| author | wenzelm | 
| Sun, 03 May 2015 00:01:10 +0200 | |
| changeset 60215 | 5fb4990dfc73 | 
| parent 59149 | 0070053570c4 | 
| child 62826 | eb94e570c1a4 | 
| permissions | -rw-r--r-- | 
| 40393 | 1  | 
(* Title: Pure/General/timing.ML  | 
| 31686 | 2  | 
Author: Makarius  | 
3  | 
||
| 40336 | 4  | 
Basic support for time measurement.  | 
| 31686 | 5  | 
*)  | 
6  | 
||
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
7  | 
signature BASIC_TIMING =  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
8  | 
sig  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
9  | 
val cond_timeit: bool -> string -> (unit -> 'a) -> 'a  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
10  | 
val timeit: (unit -> 'a) -> 'a  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
11  | 
  val timeap: ('a -> 'b) -> 'a -> 'b
 | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
12  | 
  val timeap_msg: string -> ('a -> 'b) -> 'a -> 'b
 | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
13  | 
end  | 
| 
40300
 
d4487353b3a0
added convenience operation seconds: real -> time;
 
wenzelm 
parents: 
32935 
diff
changeset
 | 
14  | 
|
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
15  | 
signature TIMING =  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
16  | 
sig  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
17  | 
include BASIC_TIMING  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
18  | 
  type timing = {elapsed: Time.time, cpu: Time.time, gc: Time.time}
 | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
19  | 
type start  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
20  | 
val start: unit -> start  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
21  | 
val result: start -> timing  | 
| 
42013
 
1694cc477045
pure Timing.timing, without any exception handling;
 
wenzelm 
parents: 
42012 
diff
changeset
 | 
22  | 
  val timing: ('a -> 'b) -> 'a -> timing * 'b
 | 
| 51228 | 23  | 
val is_relevant_time: Time.time -> bool  | 
| 44440 | 24  | 
val is_relevant: timing -> bool  | 
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
25  | 
val message: timing -> string  | 
| 
51662
 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 
wenzelm 
parents: 
51606 
diff
changeset
 | 
26  | 
val protocol_message: Properties.T -> timing -> unit  | 
| 
 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 
wenzelm 
parents: 
51606 
diff
changeset
 | 
27  | 
  val protocol: Properties.T -> ('a -> 'b) -> 'a -> 'b
 | 
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
28  | 
end  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
29  | 
|
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
30  | 
structure Timing: TIMING =  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
31  | 
struct  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
32  | 
|
| 
51217
 
65ab2c4f4c32
support for prescient timing information within command transactions;
 
wenzelm 
parents: 
50781 
diff
changeset
 | 
33  | 
(* type timing *)  | 
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
34  | 
|
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
35  | 
type timing = {elapsed: Time.time, cpu: Time.time, gc: Time.time};
 | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
36  | 
|
| 
51217
 
65ab2c4f4c32
support for prescient timing information within command transactions;
 
wenzelm 
parents: 
50781 
diff
changeset
 | 
37  | 
|
| 
 
65ab2c4f4c32
support for prescient timing information within command transactions;
 
wenzelm 
parents: 
50781 
diff
changeset
 | 
38  | 
(* timer control *)  | 
| 
 
65ab2c4f4c32
support for prescient timing information within command transactions;
 
wenzelm 
parents: 
50781 
diff
changeset
 | 
39  | 
|
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
40  | 
abstype start = Start of  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
41  | 
Timer.real_timer * Time.time * Timer.cpu_timer *  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
42  | 
    {gc: {sys: Time.time, usr: Time.time}, nongc: {sys: Time.time, usr: Time.time}}
 | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
43  | 
with  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
44  | 
|
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
45  | 
fun start () =  | 
| 31686 | 46  | 
let  | 
47  | 
val real_timer = Timer.startRealTimer ();  | 
|
48  | 
val real_time = Timer.checkRealTimer real_timer;  | 
|
49  | 
val cpu_timer = Timer.startCPUTimer ();  | 
|
50  | 
val cpu_times = Timer.checkCPUTimes cpu_timer;  | 
|
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
51  | 
in Start (real_timer, real_time, cpu_timer, cpu_times) end;  | 
| 31686 | 52  | 
|
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
53  | 
fun result (Start (real_timer, real_time, cpu_timer, cpu_times)) =  | 
| 31686 | 54  | 
let  | 
55  | 
val real_time2 = Timer.checkRealTimer real_timer;  | 
|
56  | 
    val {nongc = {sys, usr}, gc = {sys = gc_sys, usr = gc_usr}} = cpu_times;
 | 
|
57  | 
    val {nongc = {sys = sys2, usr = usr2}, gc = {sys = gc_sys2, usr = gc_usr2}} =
 | 
|
58  | 
Timer.checkCPUTimes cpu_timer;  | 
|
59  | 
||
60  | 
open Time;  | 
|
61  | 
val elapsed = real_time2 - real_time;  | 
|
62  | 
val gc = gc_usr2 - gc_usr + gc_sys2 - gc_sys;  | 
|
63  | 
val cpu = usr2 - usr + sys2 - sys + gc;  | 
|
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
64  | 
  in {elapsed = elapsed, cpu = cpu, gc = gc} end;
 | 
| 
32935
 
2218b970325a
show direct GC time (which is included in CPU time);
 
wenzelm 
parents: 
31757 
diff
changeset
 | 
65  | 
|
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
66  | 
end;  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
67  | 
|
| 
42013
 
1694cc477045
pure Timing.timing, without any exception handling;
 
wenzelm 
parents: 
42012 
diff
changeset
 | 
68  | 
fun timing f x =  | 
| 
 
1694cc477045
pure Timing.timing, without any exception handling;
 
wenzelm 
parents: 
42012 
diff
changeset
 | 
69  | 
let  | 
| 
 
1694cc477045
pure Timing.timing, without any exception handling;
 
wenzelm 
parents: 
42012 
diff
changeset
 | 
70  | 
val start = start ();  | 
| 
 
1694cc477045
pure Timing.timing, without any exception handling;
 
wenzelm 
parents: 
42012 
diff
changeset
 | 
71  | 
val y = f x;  | 
| 
 
1694cc477045
pure Timing.timing, without any exception handling;
 
wenzelm 
parents: 
42012 
diff
changeset
 | 
72  | 
in (result start, y) end;  | 
| 
 
1694cc477045
pure Timing.timing, without any exception handling;
 
wenzelm 
parents: 
42012 
diff
changeset
 | 
73  | 
|
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
74  | 
|
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
75  | 
(* timing messages *)  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
76  | 
|
| 42819 | 77  | 
val min_time = Time.fromMilliseconds 1;  | 
78  | 
||
| 51228 | 79  | 
fun is_relevant_time time = Time.>= (time, min_time);  | 
80  | 
||
| 42819 | 81  | 
fun is_relevant {elapsed, cpu, gc} =
 | 
| 51228 | 82  | 
is_relevant_time elapsed orelse  | 
83  | 
is_relevant_time cpu orelse  | 
|
84  | 
is_relevant_time gc;  | 
|
| 42819 | 85  | 
|
| 44440 | 86  | 
fun message {elapsed, cpu, gc} =
 | 
87  | 
Time.toString elapsed ^ "s elapsed time, " ^  | 
|
88  | 
Time.toString cpu ^ "s cpu time, " ^  | 
|
89  | 
Time.toString gc ^ "s GC time" handle Time.Time => "";  | 
|
90  | 
||
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
91  | 
fun cond_timeit enabled msg e =  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
92  | 
if enabled then  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
93  | 
let  | 
| 
51606
 
2843cc095a57
additional timing status for implicitly forked terminal proofs -- proper accounting for interactive Timing dockable etc.;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
94  | 
val (t, result) = timing (Exn.interruptible_capture e) ();  | 
| 42819 | 95  | 
val _ =  | 
| 
51606
 
2843cc095a57
additional timing status for implicitly forked terminal proofs -- proper accounting for interactive Timing dockable etc.;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
96  | 
if is_relevant t then  | 
| 
 
2843cc095a57
additional timing status for implicitly forked terminal proofs -- proper accounting for interactive Timing dockable etc.;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
97  | 
let val end_msg = message t  | 
| 42819 | 98  | 
in warning (if msg = "" then end_msg else msg ^ "\n" ^ end_msg) end  | 
99  | 
else ();  | 
|
| 
42013
 
1694cc477045
pure Timing.timing, without any exception handling;
 
wenzelm 
parents: 
42012 
diff
changeset
 | 
100  | 
in Exn.release result end  | 
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
101  | 
else e ();  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
102  | 
|
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
103  | 
fun timeit e = cond_timeit true "" e;  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
104  | 
fun timeap f x = timeit (fn () => f x);  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
105  | 
fun timeap_msg msg f x = cond_timeit true msg (fn () => f x);  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
106  | 
|
| 
51662
 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 
wenzelm 
parents: 
51606 
diff
changeset
 | 
107  | 
fun protocol_message props t =  | 
| 
59149
 
0070053570c4
suppress irrelevant timing messages (the majority);
 
wenzelm 
parents: 
56333 
diff
changeset
 | 
108  | 
if is_relevant t then  | 
| 
 
0070053570c4
suppress irrelevant timing messages (the majority);
 
wenzelm 
parents: 
56333 
diff
changeset
 | 
109  | 
Output.try_protocol_message (props @ Markup.timing_properties t) []  | 
| 
 
0070053570c4
suppress irrelevant timing messages (the majority);
 
wenzelm 
parents: 
56333 
diff
changeset
 | 
110  | 
else ();  | 
| 
51662
 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 
wenzelm 
parents: 
51606 
diff
changeset
 | 
111  | 
|
| 
 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 
wenzelm 
parents: 
51606 
diff
changeset
 | 
112  | 
fun protocol props f x =  | 
| 
51606
 
2843cc095a57
additional timing status for implicitly forked terminal proofs -- proper accounting for interactive Timing dockable etc.;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
113  | 
let  | 
| 
 
2843cc095a57
additional timing status for implicitly forked terminal proofs -- proper accounting for interactive Timing dockable etc.;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
114  | 
val (t, result) = timing (Exn.interruptible_capture f) x;  | 
| 
51662
 
3391a493f39a
just one timing protocol function, with 3 implementations: TTY/PG, PIDE/document, build;
 
wenzelm 
parents: 
51606 
diff
changeset
 | 
115  | 
val _ = protocol_message props t;  | 
| 
51606
 
2843cc095a57
additional timing status for implicitly forked terminal proofs -- proper accounting for interactive Timing dockable etc.;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
116  | 
in Exn.release result end;  | 
| 
 
2843cc095a57
additional timing status for implicitly forked terminal proofs -- proper accounting for interactive Timing dockable etc.;
 
wenzelm 
parents: 
51228 
diff
changeset
 | 
117  | 
|
| 
42012
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
118  | 
end;  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
119  | 
|
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
120  | 
structure Basic_Timing: BASIC_TIMING = Timing;  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
121  | 
open Basic_Timing;  | 
| 
 
2c3fe3cbebae
structure Timing: covers former start_timing/end_timing and Output.timeit etc;
 
wenzelm 
parents: 
40393 
diff
changeset
 | 
122  |