Skip to content

Commit 407f476

Browse files
committed
Introduce --stats
1 parent 7eac1a1 commit 407f476

File tree

6 files changed

+220
-9
lines changed

6 files changed

+220
-9
lines changed

src/basic/FStarC.Options.fst

Lines changed: 31 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -97,19 +97,25 @@ let as_comma_string_list = function
9797
* No stack should ever be empty! Any of these failwiths should never be
9898
* triggered externally. IOW, the API should protect this invariant.
9999
*
100-
* We also keep a snapshot of the Debug module's state.
100+
* We also keep a snapshot of the stateful modules that are modified by
101+
* changing options. Currently: Debug, Ext (extension options) and Stats.
101102
*)
102-
let history1 = Debug.saved_state & Ext.ext_state & optionstate
103+
let history1 =
104+
Debug.saved_state &
105+
Ext.ext_state &
106+
bool & (* value of Stats.enabled *)
107+
optionstate
103108

104109
let fstar_options : ref optionstate = mk_ref (PSMap.empty ())
105110

106111
let snapshot_all () : history1 =
107-
(Debug.snapshot (), Ext.save (), !fstar_options)
112+
(Debug.snapshot (), Ext.save (), !Stats.enabled, !fstar_options)
108113

109114
let restore_all (h : history1) : unit =
110-
let dbg, ext, opts = h in
115+
let dbg, ext, stats, opts = h in
111116
Debug.restore dbg;
112117
Ext.restore ext;
118+
Stats.enabled := stats;
113119
fstar_options := opts
114120

115121

@@ -120,17 +126,15 @@ let peek () = !fstar_options
120126

121127
let internal_push () =
122128
let lev1::rest = !history in
123-
let newhd = (Debug.snapshot (), Ext.save (), !fstar_options) in
129+
let newhd = snapshot_all () in
124130
history := (newhd :: lev1) :: rest
125131

126132
let internal_pop () =
127133
let lev1::rest = !history in
128134
match lev1 with
129135
| [] -> false
130-
| (dbg, ext, opts)::lev1' ->
131-
Debug.restore dbg;
132-
Ext.restore ext;
133-
fstar_options := opts;
136+
| snap :: lev1' ->
137+
restore_all snap;
134138
history := lev1' :: rest;
135139
true
136140

@@ -299,6 +303,7 @@ let defaults =
299303
("smtencoding.valid_intro" , Bool true);
300304
("smtencoding.valid_elim" , Bool false);
301305
("split_queries" , String "on_failure");
306+
("stats" , Bool false);
302307
("tactics_failhard" , Bool false);
303308
("tactics_info" , Bool false);
304309
("tactic_raw_binders" , Bool false);
@@ -558,6 +563,7 @@ let get_smtencoding_l_arith_repr() = lookup_opt "smtencoding.l_arith_repr"
558563
let get_smtencoding_valid_intro () = lookup_opt "smtencoding.valid_intro" as_bool
559564
let get_smtencoding_valid_elim () = lookup_opt "smtencoding.valid_elim" as_bool
560565
let get_split_queries () = lookup_opt "split_queries" as_string
566+
let get_stats () = lookup_opt "stats" as_bool
561567
let get_tactic_raw_binders () = lookup_opt "tactic_raw_binders" as_bool
562568
let get_tactics_failhard () = lookup_opt "tactics_failhard" as_bool
563569
let get_tactics_info () = lookup_opt "tactics_info" as_bool
@@ -1414,6 +1420,20 @@ let rec specs_with_types warn_unsafe : list (char & string & opt_type & Pprint.d
14141420
text "Use 'yes' to always split.";
14151421
]);
14161422
1423+
( noshort,
1424+
"stats",
1425+
PostProcessed ((fun b ->
1426+
(match b with
1427+
| Bool true ->
1428+
Stats.enabled := true;
1429+
Stats.ever_enabled := true
1430+
| Bool false ->
1431+
Stats.enabled := false
1432+
| _ -> ());
1433+
b),
1434+
BoolStr),
1435+
text "Print some statistics on the time spent in each phase of the compiler");
1436+
14171437
( noshort,
14181438
"tactic_raw_binders",
14191439
Const (Bool true),
@@ -1777,6 +1797,7 @@ let settable = function
17771797
| "smtencoding.valid_intro"
17781798
| "smtencoding.valid_elim"
17791799
| "split_queries"
1800+
| "stats"
17801801
| "tactic_raw_binders"
17811802
| "tactics_failhard"
17821803
| "tactics_info"
@@ -2164,6 +2185,7 @@ let parse_split_queries (s:string) : option split_queries_t =
21642185
| _ -> None
21652186
21662187
let split_queries () = get_split_queries () |> parse_split_queries |> Util.must
2188+
let stats () = get_stats ()
21672189
let tactic_raw_binders () = get_tactic_raw_binders ()
21682190
let tactics_failhard () = get_tactics_failhard ()
21692191
let tactics_info () = get_tactics_info ()

src/basic/FStarC.Options.fsti

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -232,6 +232,7 @@ val smtencoding_l_arith_native : unit -> bool
232232
val smtencoding_valid_intro : unit -> bool
233233
val smtencoding_valid_elim : unit -> bool
234234
val split_queries : unit -> split_queries_t
235+
val stats : unit -> bool
235236
val tactic_raw_binders : unit -> bool
236237
val tactics_failhard : unit -> bool
237238
val tactics_info : unit -> bool

src/basic/FStarC.Profiling.fst

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -67,6 +67,7 @@ let create_or_lookup_counter cid =
6767

6868
(* Time an operation, if the the profiler is enabled *)
6969
let profile (f: unit -> 'a) (module_name:option string) (cid:string) : 'a =
70+
// Stats.record cid fun () ->
7071
if Options.profile_enabled module_name cid
7172
then let c = create_or_lookup_counter cid in
7273
if !c.running //if the counter is already running

src/basic/FStarC.Stats.fst

Lines changed: 142 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,142 @@
1+
(*
2+
Copyright 2008-2025 Nikhil Swamy and Microsoft Research
3+
4+
Licensed under the Apache License, Version 2.0 (the "License");
5+
you may not use this file except in compliance with the License.
6+
You may obtain a copy of the License at
7+
8+
http://www.apache.org/licenses/LICENSE-2.0
9+
10+
Unless required by applicable law or agreed to in writing, software
11+
distributed under the License is distributed on an "AS IS" BASIS,
12+
WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied.
13+
See the License for the specific language governing permissions and
14+
limitations under the License.
15+
*)
16+
module FStarC.Stats
17+
18+
open FStarC.Effect
19+
open FStarC.Class.Monoid
20+
21+
let enabled = alloc false
22+
let ever_enabled = alloc false
23+
24+
type stat = {
25+
ns_tree : int;
26+
ns_exn : int;
27+
ns_sub : int;
28+
ncalls : int;
29+
}
30+
31+
instance _ : monoid stat = {
32+
mzero = {
33+
ns_tree = 0;
34+
ns_exn = 0;
35+
ns_sub = 0;
36+
ncalls = 0;
37+
};
38+
mplus = (fun s1 s2 ->
39+
{
40+
ns_tree = s1.ns_tree + s2.ns_tree;
41+
ns_exn = s1.ns_exn + s2.ns_exn;
42+
ns_sub = s1.ns_sub + s2.ns_sub;
43+
ncalls = s1.ncalls + s2.ncalls;
44+
});
45+
}
46+
47+
(* the ref bool marks whether a given key is currently
48+
being recorded. This is so we avoid double counting
49+
the time taken by reentrant calls. *)
50+
let st : SMap.t (ref bool & stat) = SMap.create 10
51+
52+
(* Current stats we are logging. This is used to distinguish
53+
"tree" time (all the time taken by some call)
54+
vs "point" time (time taken by some call, subtracting
55+
the time in subcalls, if any). *)
56+
let stack : ref (list string) = mk_ref []
57+
58+
let r_running (k : string) : ref bool =
59+
match SMap.try_find st k with
60+
| None ->
61+
let r = alloc false in
62+
SMap.add st k (r, mzero);
63+
r
64+
| Some (r, _) ->
65+
r
66+
67+
let add (k : string) (s1 : stat) : unit =
68+
let (r, s0) =
69+
match SMap.try_find st k with
70+
| None -> (alloc false, mzero)
71+
| Some rs -> rs
72+
in
73+
SMap.add st k (r, mplus s0 s1)
74+
75+
let do_record
76+
(key : string)
77+
(f : unit -> 'a)
78+
: 'a
79+
=
80+
stack := key :: !stack;
81+
let running = r_running key in
82+
let was_running = !running in
83+
running := true;
84+
let t0 = Timing.now_ns () in
85+
let resexn =
86+
try Inr (f ())
87+
with | e -> Inl e
88+
in
89+
running := was_running;
90+
let t1 = Timing.now_ns () in
91+
let ns = Timing.diff_ns t0 t1 in
92+
stack := List.tl !stack;
93+
if not was_running then (
94+
add key { mzero with ns_tree = ns };
95+
(* Take time out of the parent, if any. *)
96+
begin match !stack with
97+
| [] -> ()
98+
| k_par::_ -> add k_par { mzero with ns_sub = ns }
99+
end
100+
);
101+
add key { mzero with ncalls = 1 };
102+
match resexn with
103+
| Inr r ->
104+
r
105+
| Inl e ->
106+
add key { mzero with ns_exn = ns };
107+
raise e
108+
109+
let record key f =
110+
if !enabled then
111+
do_record key f
112+
else
113+
f ()
114+
115+
let lpad (len:int) (s:string) : string =
116+
let l = String.length s in
117+
if l >= len then s else String.make (len - l) ' ' ^ s
118+
119+
let max x y =
120+
if x > y then x else y
121+
122+
let print_all () : string =
123+
let keys = SMap.keys st in
124+
let points = List.map (fun k -> k, snd <| Some?.v <| SMap.try_find st k) keys in
125+
(* Sort by (point) time. *)
126+
let points =
127+
points |>
128+
Class.Ord.sort_by (fun (_, s1) (_, s2) ->
129+
(s2.ns_tree - s2.ns_sub) `Class.Ord.cmp` (s1.ns_tree - s1.ns_sub))
130+
in
131+
let longest_key = List.fold_left (fun acc (k, _) -> max acc (String.length k)) 20 points in
132+
let pr1 (p : (string & stat)) : string =
133+
let k, st = p in
134+
Util.format5 " %s %s %s ms %s ms %s ms"
135+
(lpad longest_key k)
136+
(lpad 8 (string_of_int st.ncalls))
137+
(lpad 6 (string_of_int (st.ns_tree / 1000000)))
138+
(lpad 6 (string_of_int ((st.ns_tree - st.ns_sub) / 1000000)))
139+
(lpad 6 (string_of_int (st.ns_exn / 1000000)))
140+
in
141+
Util.format5 " %s %s %s %s %s" (lpad longest_key "key") (lpad 8 "calls") (lpad 9 "tree") (lpad 9 "point") (lpad 9 "exn") ^ "\n" ^
142+
(points |> List.map pr1 |> String.concat "\n")

src/basic/FStarC.Stats.fsti

Lines changed: 37 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,37 @@
1+
(*
2+
Copyright 2008-2025 Nikhil Swamy and Microsoft Research
3+
4+
Licensed under the Apache License, Version 2.0 (the "License");
5+
you may not use this file except in compliance with the License.
6+
You may obtain a copy of the License at
7+
8+
http://www.apache.org/licenses/LICENSE-2.0
9+
10+
Unless required by applicable law or agreed to in writing, software
11+
distributed under the License is distributed on an "AS IS" BASIS,
12+
WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied.
13+
See the License for the specific language governing permissions and
14+
limitations under the License.
15+
*)
16+
module FStarC.Stats
17+
18+
open FStarC.Effect
19+
20+
(* We only record stats when this is set to true. This is
21+
set by the Options module. *)
22+
val enabled : ref bool
23+
24+
(* This is set to true by the Options module whenever
25+
--stats true is passed, and never set to false. We use it
26+
to decide whether to print the stats at the end of
27+
the run. *)
28+
val ever_enabled : ref bool
29+
30+
(* Count the time taken by `f ()` under a given stats key. *)
31+
val record
32+
(key : string)
33+
(f : unit -> 'a)
34+
: 'a
35+
36+
(* Generates a message with a table for all stat keys. *)
37+
val print_all () : string

src/fstar/FStarC.Main.fst

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -29,6 +29,10 @@ open FStarC.Class.Show
2929
module E = FStarC.Errors
3030
module UF = FStarC.Syntax.Unionfind
3131

32+
let print_stats () =
33+
if !Stats.ever_enabled then
34+
print1_error "Stats:\n%s\n" (Stats.print_all ())
35+
3236
(* These modules only mentioned to put them in the dep graph
3337
and hence compile and link them in. They do not export anything,
3438
instead they register primitive steps in the normalizer during
@@ -65,6 +69,7 @@ let report_errors fmods =
6569
let nerrs = FStarC.Errors.get_err_count() in
6670
if nerrs > 0 then begin
6771
finished_message fmods nerrs;
72+
print_stats ();
6873
exit 1
6974
end
7075

@@ -125,6 +130,7 @@ let set_error_trap () =
125130
Errors.diag Range.dummyRange [
126131
text "GOT SIGINT! Exiting";
127132
];
133+
print_stats ();
128134
exit 1
129135
in
130136
set_sigint_handler (sigint_handler_f h')
@@ -422,8 +428,10 @@ let main () =
422428
then Util.print2_error "TOTAL TIME %s ms: %s\n"
423429
(FStarC.Util.string_of_int time)
424430
(String.concat " " (FStarC.Getopt.cmdline()));
431+
print_stats();
425432
exit 0
426433
with
427434
| e ->
428435
handle_error e;
436+
print_stats();
429437
exit 1

0 commit comments

Comments
 (0)