| author | wenzelm | 
| Tue, 26 Sep 2023 12:30:08 +0200 | |
| changeset 78715 | 9506b852ebdf | 
| parent 78714 | eb2255d241da | 
| child 78716 | 97dfba4405e3 | 
| 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  | 
| 71883 | 18  | 
val stack_limit: unit -> int option  | 
| 60764 | 19  | 
  type params = {name: string, stack_limit: int option, interrupts: bool}
 | 
| 
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
 | 
20  | 
val attributes: params -> Thread.Thread.threadAttribute list  | 
| 78648 | 21  | 
val fork: params -> (unit -> unit) -> T  | 
22  | 
val is_active: T -> bool  | 
|
23  | 
val join: T -> unit  | 
|
| 78681 | 24  | 
val interrupt: exn  | 
25  | 
val interrupt_exn: 'a Exn.result  | 
|
26  | 
val interrupt_self: unit -> 'a  | 
|
27  | 
val interrupt_other: T -> unit  | 
|
| 
78715
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
28  | 
structure Exn: EXN  | 
| 78714 | 29  | 
val expose_interrupt_result: unit -> unit Exn.result  | 
| 78713 | 30  | 
val expose_interrupt: unit -> unit (*exception Interrupt*)  | 
| 78701 | 31  | 
val try_catch: (unit -> 'a) -> (exn -> 'a) -> 'a  | 
32  | 
val try_finally: (unit -> 'a) -> (unit -> unit) -> 'a  | 
|
| 78688 | 33  | 
val try: (unit -> 'a) -> 'a option  | 
34  | 
val can: (unit -> 'a) -> bool  | 
|
| 28241 | 35  | 
end;  | 
36  | 
||
| 71692 | 37  | 
structure Isabelle_Thread: ISABELLE_THREAD =  | 
| 28241 | 38  | 
struct  | 
39  | 
||
| 78648 | 40  | 
(* abstract type *)  | 
41  | 
||
| 
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
 | 
42  | 
abstype T = T of {thread: Thread.Thread.thread, name: string, id: int}
 | 
| 78648 | 43  | 
with  | 
44  | 
val make = T;  | 
|
45  | 
fun dest (T args) = args;  | 
|
46  | 
end;  | 
|
47  | 
||
48  | 
val get_thread = #thread o dest;  | 
|
49  | 
val get_name = #name o dest;  | 
|
50  | 
val get_id = #id o dest;  | 
|
51  | 
||
| 
78711
 
3a3a70d4d422
clarified order of modules: early access to interrupt management of Isabelle_Threads;
 
wenzelm 
parents: 
78704 
diff
changeset
 | 
52  | 
fun equal (t1, t2) = Thread.Thread.equal (get_thread t1, get_thread t2);  | 
| 78648 | 53  | 
|
54  | 
fun print t =  | 
|
55  | 
(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
 | 
56  | 
"-" ^ Int.toString (get_id t);  | 
| 78648 | 57  | 
|
58  | 
||
| 60764 | 59  | 
(* self *)  | 
60  | 
||
| 78648 | 61  | 
val make_id = Counter.make ();  | 
| 60764 | 62  | 
|
63  | 
local  | 
|
| 78648 | 64  | 
val self_var = Thread_Data.var () : T Thread_Data.var;  | 
| 60764 | 65  | 
in  | 
66  | 
||
| 78678 | 67  | 
fun init_self args =  | 
68  | 
let val t = make args in Thread_Data.put self_var (SOME t); t end;  | 
|
| 60830 | 69  | 
|
| 78648 | 70  | 
fun self () =  | 
71  | 
(case Thread_Data.get self_var of  | 
|
72  | 
SOME t => t  | 
|
| 78678 | 73  | 
  | NONE => init_self {thread = Thread.Thread.self (), name = "", id = make_id ()});
 | 
| 78648 | 74  | 
|
75  | 
fun is_self t = equal (t, self ());  | 
|
| 60764 | 76  | 
|
77  | 
end;  | 
|
78  | 
||
79  | 
||
80  | 
(* fork *)  | 
|
81  | 
||
| 
78711
 
3a3a70d4d422
clarified order of modules: early access to interrupt management of Isabelle_Threads;
 
wenzelm 
parents: 
78704 
diff
changeset
 | 
82  | 
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
 | 
83  | 
|
| 71883 | 84  | 
fun stack_limit () =  | 
85  | 
let  | 
|
| 
78711
 
3a3a70d4d422
clarified order of modules: early access to interrupt management of Isabelle_Threads;
 
wenzelm 
parents: 
78704 
diff
changeset
 | 
86  | 
val limit = Real.floor (! threads_stack_limit * 1024.0 * 1024.0 * 1024.0);  | 
| 
 
3a3a70d4d422
clarified order of modules: early access to interrupt management of Isabelle_Threads;
 
wenzelm 
parents: 
78704 
diff
changeset
 | 
87  | 
in if limit <= 0 then NONE else SOME limit end;  | 
| 71883 | 88  | 
|
| 60764 | 89  | 
type params = {name: string, stack_limit: int option, interrupts: bool};
 | 
90  | 
||
91  | 
fun attributes ({stack_limit, interrupts, ...}: 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
 | 
92  | 
Thread.Thread.MaximumMLStack stack_limit ::  | 
| 
64557
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
62923 
diff
changeset
 | 
93  | 
Thread_Attributes.convert_attributes  | 
| 
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
62923 
diff
changeset
 | 
94  | 
(if interrupts then Thread_Attributes.public_interrupts 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
 | 
95  | 
|
| 60764 | 96  | 
fun fork (params: params) body =  | 
| 78648 | 97  | 
let  | 
| 78677 | 98  | 
val self = Single_Assignment.var "self";  | 
99  | 
fun main () =  | 
|
100  | 
let  | 
|
| 78678 | 101  | 
        val t = init_self {thread = Thread.Thread.self (), name = #name params, id = make_id ()};
 | 
| 78677 | 102  | 
val _ = Single_Assignment.assign self t;  | 
103  | 
in body () end;  | 
|
104  | 
val _ = Thread.Thread.fork (main, attributes params);  | 
|
105  | 
in Single_Assignment.await self end;  | 
|
| 28241 | 106  | 
|
| 60764 | 107  | 
|
108  | 
(* join *)  | 
|
109  | 
||
| 
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
 | 
110  | 
val is_active = Thread.Thread.isActive o get_thread;  | 
| 78648 | 111  | 
|
112  | 
fun join t =  | 
|
113  | 
while is_active t  | 
|
| 52583 | 114  | 
do OS.Process.sleep (seconds 0.1);  | 
115  | 
||
| 60764 | 116  | 
|
| 78681 | 117  | 
(* interrupts *)  | 
| 60764 | 118  | 
|
| 78681 | 119  | 
val interrupt = Thread.Thread.Interrupt;  | 
120  | 
val interrupt_exn = Exn.Exn interrupt;  | 
|
121  | 
||
122  | 
fun interrupt_self () = raise interrupt;  | 
|
123  | 
||
124  | 
fun interrupt_other 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
 | 
125  | 
Thread.Thread.interrupt (get_thread t) handle Thread.Thread _ => ();  | 
| 28550 | 126  | 
|
| 
78715
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
127  | 
structure Exn: EXN =  | 
| 
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
128  | 
struct  | 
| 
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
129  | 
open Exn;  | 
| 
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
130  | 
val capture = capture0;  | 
| 
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
131  | 
end;  | 
| 
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
132  | 
|
| 78714 | 133  | 
fun expose_interrupt_result () =  | 
| 78713 | 134  | 
let  | 
135  | 
val orig_atts = Thread_Attributes.safe_interrupts (Thread_Attributes.get_attributes ());  | 
|
136  | 
val _ = Thread_Attributes.set_attributes Thread_Attributes.test_interrupts;  | 
|
137  | 
val test = Exn.capture Thread.Thread.testInterrupt ();  | 
|
138  | 
val _ = Thread_Attributes.set_attributes orig_atts;  | 
|
| 78714 | 139  | 
in test end;  | 
140  | 
||
141  | 
val expose_interrupt = Exn.release o expose_interrupt_result;  | 
|
| 78713 | 142  | 
|
| 78701 | 143  | 
fun try_catch e f =  | 
| 
78704
 
b54aee45cabe
more robust: catch/finally part is uninterruptible;
 
wenzelm 
parents: 
78701 
diff
changeset
 | 
144  | 
Thread_Attributes.uninterruptible (fn restore_attributes => fn () =>  | 
| 
 
b54aee45cabe
more robust: catch/finally part is uninterruptible;
 
wenzelm 
parents: 
78701 
diff
changeset
 | 
145  | 
restore_attributes e ()  | 
| 
 
b54aee45cabe
more robust: catch/finally part is uninterruptible;
 
wenzelm 
parents: 
78701 
diff
changeset
 | 
146  | 
handle exn => if Exn.is_interrupt exn then Exn.reraise exn else f exn) ();  | 
| 78701 | 147  | 
|
148  | 
fun try_finally e f =  | 
|
| 
78704
 
b54aee45cabe
more robust: catch/finally part is uninterruptible;
 
wenzelm 
parents: 
78701 
diff
changeset
 | 
149  | 
Thread_Attributes.uninterruptible (fn restore_attributes => fn () =>  | 
| 
 
b54aee45cabe
more robust: catch/finally part is uninterruptible;
 
wenzelm 
parents: 
78701 
diff
changeset
 | 
150  | 
Exn.release (Exn.capture (restore_attributes e) () before f ())) ();  | 
| 78701 | 151  | 
|
| 78688 | 152  | 
fun try e = Basics.try e ();  | 
153  | 
fun can e = Basics.can e ();  | 
|
154  | 
||
| 28241 | 155  | 
end;  | 
| 
78715
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
156  | 
|
| 
 
9506b852ebdf
clarified signature: distinction of unmanaged vs. managed interrupts (not implemented yet);
 
wenzelm 
parents: 
78714 
diff
changeset
 | 
157  | 
structure Exn = Isabelle_Thread.Exn;  |