| author | wenzelm | 
| Tue, 17 Oct 2023 11:47:13 +0200 | |
| changeset 78785 | c8202ed06a5c | 
| parent 78765 | f971ccccfd8e | 
| child 78786 | 85efa3d01b16 | 
| permissions | -rw-r--r-- | 
| 71692 | 1  | 
(* Title: Pure/Concurrent/isabelle_thread.ML  | 
| 28241 | 2  | 
Author: Makarius  | 
3  | 
||
| 71692 | 4  | 
Isabelle-specific thread management.  | 
| 28241 | 5  | 
*)  | 
6  | 
||
| 71692 | 7  | 
signature ISABELLE_THREAD =  | 
| 28241 | 8  | 
sig  | 
| 78648 | 9  | 
type T  | 
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
78648 
diff
changeset
 | 
10  | 
val get_thread: T -> Thread.Thread.thread  | 
| 78648 | 11  | 
val get_name: T -> string  | 
12  | 
val get_id: T -> int  | 
|
13  | 
val equal: T * T -> bool  | 
|
14  | 
val print: T -> string  | 
|
15  | 
val self: unit -> T  | 
|
16  | 
val is_self: T -> bool  | 
|
| 
78711
 
3a3a70d4d422
clarified order of modules: early access to interrupt management of Isabelle_Threads;
 
wenzelm 
parents: 
78704 
diff
changeset
 | 
17  | 
val threads_stack_limit: real Unsynchronized.ref  | 
| 78753 | 18  | 
val default_stack_limit: unit -> int option  | 
19  | 
type params  | 
|
20  | 
val params: string -> params  | 
|
21  | 
val stack_limit: int -> params -> params  | 
|
22  | 
val interrupts: params -> params  | 
|
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
78648 
diff
changeset
 | 
23  | 
val attributes: params -> Thread.Thread.threadAttribute list  | 
| 78648 | 24  | 
val fork: params -> (unit -> unit) -> T  | 
25  | 
val is_active: T -> bool  | 
|
26  | 
val join: T -> unit  | 
|
| 78681 | 27  | 
val interrupt: exn  | 
28  | 
val interrupt_exn: 'a Exn.result  | 
|
29  | 
val interrupt_self: unit -> 'a  | 
|
30  | 
val interrupt_other: T -> unit  | 
|
| 
78715
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
31  | 
structure Exn: EXN  | 
| 78714 | 32  | 
val expose_interrupt_result: unit -> unit Exn.result  | 
| 
78740
 
45ff003d337c
discontinue obsolete "Interrupt" constructor (NB: catch-all pattern produces ML compiler error);
 
wenzelm 
parents: 
78720 
diff
changeset
 | 
33  | 
val expose_interrupt: unit -> unit (*exception Exn.is_interrupt*)  | 
| 78701 | 34  | 
val try_catch: (unit -> 'a) -> (exn -> 'a) -> 'a  | 
35  | 
val try_finally: (unit -> 'a) -> (unit -> unit) -> 'a  | 
|
| 78688 | 36  | 
val try: (unit -> 'a) -> 'a option  | 
37  | 
val can: (unit -> 'a) -> bool  | 
|
| 28241 | 38  | 
end;  | 
39  | 
||
| 71692 | 40  | 
structure Isabelle_Thread: ISABELLE_THREAD =  | 
| 28241 | 41  | 
struct  | 
42  | 
||
| 78648 | 43  | 
(* abstract type *)  | 
44  | 
||
| 
78764
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
45  | 
abstype T = T of {thread: Thread.Thread.thread, name: string, id: int, break: bool Synchronized.var}
 | 
| 78648 | 46  | 
with  | 
47  | 
val make = T;  | 
|
48  | 
fun dest (T args) = args;  | 
|
49  | 
end;  | 
|
50  | 
||
51  | 
val get_thread = #thread o dest;  | 
|
52  | 
val get_name = #name o dest;  | 
|
53  | 
val get_id = #id o dest;  | 
|
54  | 
||
| 
78711
 
3a3a70d4d422
clarified order of modules: early access to interrupt management of Isabelle_Threads;
 
wenzelm 
parents: 
78704 
diff
changeset
 | 
55  | 
fun equal (t1, t2) = Thread.Thread.equal (get_thread t1, get_thread t2);  | 
| 78648 | 56  | 
|
57  | 
fun print t =  | 
|
58  | 
(case get_name t of "" => "ML" | a => "Isabelle." ^ a) ^  | 
|
| 
78711
 
3a3a70d4d422
clarified order of modules: early access to interrupt management of Isabelle_Threads;
 
wenzelm 
parents: 
78704 
diff
changeset
 | 
59  | 
"-" ^ Int.toString (get_id t);  | 
| 78648 | 60  | 
|
61  | 
||
| 60764 | 62  | 
(* self *)  | 
63  | 
||
| 78648 | 64  | 
val make_id = Counter.make ();  | 
| 60764 | 65  | 
|
66  | 
local  | 
|
| 78648 | 67  | 
val self_var = Thread_Data.var () : T Thread_Data.var;  | 
| 60764 | 68  | 
in  | 
69  | 
||
| 78763 | 70  | 
fun init_self name =  | 
71  | 
let  | 
|
| 
78764
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
72  | 
val break = Synchronized.var "Isabelle_Thread.break" false;  | 
| 
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
73  | 
    val t = make {thread = Thread.Thread.self (), name = name, id = make_id (), break = break};
 | 
| 78763 | 74  | 
in Thread_Data.put self_var (SOME t); t end;  | 
| 60830 | 75  | 
|
| 78648 | 76  | 
fun self () =  | 
77  | 
(case Thread_Data.get self_var of  | 
|
78  | 
SOME t => t  | 
|
| 78763 | 79  | 
| NONE => init_self "");  | 
| 78648 | 80  | 
|
81  | 
fun is_self t = equal (t, self ());  | 
|
| 60764 | 82  | 
|
83  | 
end;  | 
|
84  | 
||
85  | 
||
86  | 
(* fork *)  | 
|
87  | 
||
| 
78711
 
3a3a70d4d422
clarified order of modules: early access to interrupt management of Isabelle_Threads;
 
wenzelm 
parents: 
78704 
diff
changeset
 | 
88  | 
val threads_stack_limit = Unsynchronized.ref 0.25;  | 
| 
 
3a3a70d4d422
clarified order of modules: early access to interrupt management of Isabelle_Threads;
 
wenzelm 
parents: 
78704 
diff
changeset
 | 
89  | 
|
| 78753 | 90  | 
fun default_stack_limit () =  | 
| 78762 | 91  | 
let val limit = Real.floor (! threads_stack_limit * 1024.0 * 1024.0 * 1024.0)  | 
| 
78711
 
3a3a70d4d422
clarified order of modules: early access to interrupt management of Isabelle_Threads;
 
wenzelm 
parents: 
78704 
diff
changeset
 | 
92  | 
in if limit <= 0 then NONE else SOME limit end;  | 
| 71883 | 93  | 
|
| 78753 | 94  | 
abstype params = Params of {name: string, stack_limit: int option, interrupts: bool}
 | 
95  | 
with  | 
|
96  | 
||
97  | 
fun make_params (name, stack_limit, interrupts) =  | 
|
98  | 
  Params {name = name, stack_limit = stack_limit, interrupts = interrupts};
 | 
|
99  | 
||
100  | 
fun params name = make_params (name, default_stack_limit (), false);  | 
|
101  | 
fun stack_limit limit (Params {name, interrupts, ...}) = make_params (name, SOME limit, interrupts);
 | 
|
102  | 
fun interrupts (Params {name, stack_limit, ...}) = make_params (name, stack_limit, true);
 | 
|
| 60764 | 103  | 
|
| 78753 | 104  | 
fun params_name (Params {name, ...}) = name;
 | 
105  | 
fun params_stack_limit (Params {stack_limit, ...}) = stack_limit;
 | 
|
106  | 
fun params_interrupts (Params {interrupts, ...}) = interrupts;
 | 
|
107  | 
||
108  | 
end;  | 
|
109  | 
||
110  | 
fun attributes params =  | 
|
111  | 
Thread.Thread.MaximumMLStack (params_stack_limit params) ::  | 
|
| 
64557
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
62923 
diff
changeset
 | 
112  | 
Thread_Attributes.convert_attributes  | 
| 78753 | 113  | 
(if params_interrupts params  | 
114  | 
then Thread_Attributes.public_interrupts  | 
|
115  | 
else Thread_Attributes.no_interrupts);  | 
|
| 
59468
 
fe6651760643
explicit threads_stack_limit (for recent Poly/ML SVN versions), which leads to soft interrupt instead of exhaustion of virtual memory, which is particularly relevant for the bigger address space of x86_64;
 
wenzelm 
parents: 
59055 
diff
changeset
 | 
116  | 
|
| 78753 | 117  | 
fun fork params body =  | 
| 78648 | 118  | 
let  | 
| 78677 | 119  | 
val self = Single_Assignment.var "self";  | 
120  | 
fun main () =  | 
|
121  | 
let  | 
|
| 78763 | 122  | 
val t = init_self (params_name params);  | 
| 78677 | 123  | 
val _ = Single_Assignment.assign self t;  | 
124  | 
in body () end;  | 
|
125  | 
val _ = Thread.Thread.fork (main, attributes params);  | 
|
126  | 
in Single_Assignment.await self end;  | 
|
| 28241 | 127  | 
|
| 60764 | 128  | 
|
129  | 
(* join *)  | 
|
130  | 
||
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
78648 
diff
changeset
 | 
131  | 
val is_active = Thread.Thread.isActive o get_thread;  | 
| 78648 | 132  | 
|
133  | 
fun join t =  | 
|
134  | 
while is_active t  | 
|
| 52583 | 135  | 
do OS.Process.sleep (seconds 0.1);  | 
136  | 
||
| 60764 | 137  | 
|
| 78681 | 138  | 
(* interrupts *)  | 
| 60764 | 139  | 
|
| 
78764
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
140  | 
val interrupt = Exn.Interrupt_Break;  | 
| 78681 | 141  | 
val interrupt_exn = Exn.Exn interrupt;  | 
142  | 
||
143  | 
fun interrupt_self () = raise interrupt;  | 
|
144  | 
||
145  | 
fun interrupt_other t =  | 
|
| 
78764
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
146  | 
Synchronized.change (#break (dest t))  | 
| 
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
147  | 
(fn b => (Thread.Thread.interrupt (get_thread t); true) handle Thread.Thread _ => b);  | 
| 
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
148  | 
|
| 78785 | 149  | 
fun reset_interrupt () =  | 
150  | 
Synchronized.change_result (#break (dest (self ()))) (fn b => (b, false));  | 
|
151  | 
||
152  | 
fun make_interrupt break = if break then Exn.Interrupt_Break else Exn.Interrupt_Breakdown;  | 
|
153  | 
||
| 
78764
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
154  | 
fun check_interrupt exn =  | 
| 78785 | 155  | 
if Exn.is_interrupt_raw exn then make_interrupt (reset_interrupt ()) else exn;  | 
| 
78764
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
156  | 
|
| 
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
157  | 
fun check_interrupt_exn result =  | 
| 
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
158  | 
Exn.map_exn check_interrupt result;  | 
| 28550 | 159  | 
|
| 
78715
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
160  | 
structure Exn: EXN =  | 
| 
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
161  | 
struct  | 
| 
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
162  | 
open Exn;  | 
| 
78765
 
f971ccccfd8e
proper Exn.capture with check_interrupt (amending a3dcae9a2ebe);
 
wenzelm 
parents: 
78764 
diff
changeset
 | 
163  | 
fun capture f x = Res (f x) handle e => Exn (check_interrupt e);  | 
| 78757 | 164  | 
fun capture_body e = capture e ();  | 
| 
78715
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
165  | 
end;  | 
| 
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
166  | 
|
| 78714 | 167  | 
fun expose_interrupt_result () =  | 
| 78713 | 168  | 
let  | 
169  | 
val orig_atts = Thread_Attributes.safe_interrupts (Thread_Attributes.get_attributes ());  | 
|
| 78755 | 170  | 
fun main () =  | 
171  | 
(Thread_Attributes.set_attributes Thread_Attributes.test_interrupts;  | 
|
172  | 
Thread.Thread.testInterrupt ());  | 
|
| 
78764
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
173  | 
val test = check_interrupt_exn (Exn.capture_body main);  | 
| 78713 | 174  | 
val _ = Thread_Attributes.set_attributes orig_atts;  | 
| 78714 | 175  | 
in test end;  | 
176  | 
||
177  | 
val expose_interrupt = Exn.release o expose_interrupt_result;  | 
|
| 78713 | 178  | 
|
| 78701 | 179  | 
fun try_catch e f =  | 
| 78720 | 180  | 
Thread_Attributes.uninterruptible_body (fn run =>  | 
| 
78764
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
181  | 
run e () handle exn =>  | 
| 
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
182  | 
if Exn.is_interrupt exn then Exn.reraise (check_interrupt exn) else f exn);  | 
| 78701 | 183  | 
|
184  | 
fun try_finally e f =  | 
|
| 78720 | 185  | 
Thread_Attributes.uninterruptible_body (fn run =>  | 
| 
78764
 
a3dcae9a2ebe
distinguish proper interrupts from Poly/ML RTS breakdown;
 
wenzelm 
parents: 
78763 
diff
changeset
 | 
186  | 
Exn.release (check_interrupt_exn (Exn.capture_body (run e)) before f ()));  | 
| 78701 | 187  | 
|
| 78688 | 188  | 
fun try e = Basics.try e ();  | 
189  | 
fun can e = Basics.can e ();  | 
|
190  | 
||
| 28241 | 191  | 
end;  | 
| 
78715
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
192  | 
|
| 
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
193  | 
structure Exn = Isabelle_Thread.Exn;  |