| author | wenzelm | 
| Tue, 05 Mar 2024 16:06:06 +0100 | |
| changeset 79777 | db9c6be8e236 | 
| parent 79613 | 7a432595fb66 | 
| child 82720 | 956ecf2c07a0 | 
| permissions | -rw-r--r-- | 
| 
62508
 
d0b68218ea55
discontinued RAW session: bootstrap directly from isabelle_process RAW_ML_SYSTEM;
 
wenzelm 
parents: 
62501 
diff
changeset
 | 
1  | 
(* Title: Pure/Concurrent/multithreading.ML  | 
| 
24690
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 79601 | 4  | 
Multithreading in Poly/ML, see also  | 
5  | 
- $ML_SOURCES/basis/Thread.sml: numPhysicalProcessors  | 
|
6  | 
- $ML_SOURCES/libpolyml/processes.cpp: PolyThreadNumPhysicalProcessors  | 
|
| 
24690
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
signature MULTITHREADING =  | 
| 
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
sig  | 
| 
79604
 
0e8ac7db1f4d
clarified num_processors: follow Poly/ML (with its inaccuracies);
 
wenzelm 
parents: 
79601 
diff
changeset
 | 
11  | 
val num_processors: unit -> int  | 
| 62925 | 12  | 
val max_threads: unit -> int  | 
| 62359 | 13  | 
val max_threads_update: int -> unit  | 
| 68025 | 14  | 
val parallel_proofs: int ref  | 
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
68130 
diff
changeset
 | 
15  | 
val sync_wait: Time.time option -> Thread.ConditionVar.conditionVar -> Thread.Mutex.mutex ->  | 
| 
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
68130 
diff
changeset
 | 
16  | 
bool Exn.result  | 
| 
39616
 
8052101883c3
renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
 
wenzelm 
parents: 
32777 
diff
changeset
 | 
17  | 
val trace: int ref  | 
| 
32295
 
400cc493d466
renamed Multithreading.regular_interrupts to Multithreading.public_interrupts;
 
wenzelm 
parents: 
32286 
diff
changeset
 | 
18  | 
val tracing: int -> (unit -> string) -> unit  | 
| 
 
400cc493d466
renamed Multithreading.regular_interrupts to Multithreading.public_interrupts;
 
wenzelm 
parents: 
32286 
diff
changeset
 | 
19  | 
val tracing_time: bool -> Time.time -> (unit -> string) -> unit  | 
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
68130 
diff
changeset
 | 
20  | 
val synchronized: string -> Thread.Mutex.mutex -> (unit -> 'a) -> 'a  | 
| 
24690
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
end;  | 
| 
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
|
| 
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
structure Multithreading: MULTITHREADING =  | 
| 
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
struct  | 
| 
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
|
| 
79604
 
0e8ac7db1f4d
clarified num_processors: follow Poly/ML (with its inaccuracies);
 
wenzelm 
parents: 
79601 
diff
changeset
 | 
26  | 
(* physical processors *)  | 
| 62359 | 27  | 
|
| 62501 | 28  | 
fun num_processors () =  | 
| 79613 | 29  | 
Int.max  | 
30  | 
(case Thread.Thread.numPhysicalProcessors () of  | 
|
31  | 
SOME n => n  | 
|
32  | 
| NONE => Thread.Thread.numProcessors (), 1);  | 
|
| 62501 | 33  | 
|
| 
79604
 
0e8ac7db1f4d
clarified num_processors: follow Poly/ML (with its inaccuracies);
 
wenzelm 
parents: 
79601 
diff
changeset
 | 
34  | 
|
| 
 
0e8ac7db1f4d
clarified num_processors: follow Poly/ML (with its inaccuracies);
 
wenzelm 
parents: 
79601 
diff
changeset
 | 
35  | 
(* max_threads *)  | 
| 
 
0e8ac7db1f4d
clarified num_processors: follow Poly/ML (with its inaccuracies);
 
wenzelm 
parents: 
79601 
diff
changeset
 | 
36  | 
|
| 
 
0e8ac7db1f4d
clarified num_processors: follow Poly/ML (with its inaccuracies);
 
wenzelm 
parents: 
79601 
diff
changeset
 | 
37  | 
local  | 
| 
 
0e8ac7db1f4d
clarified num_processors: follow Poly/ML (with its inaccuracies);
 
wenzelm 
parents: 
79601 
diff
changeset
 | 
38  | 
|
| 62359 | 39  | 
fun max_threads_result m =  | 
| 
62926
 
9ff9bcbc2341
virtual Pure is single-threaded to avoid confusion with multiple thread farms etc.;
 
wenzelm 
parents: 
62925 
diff
changeset
 | 
40  | 
if Thread_Data.is_virtual then 1  | 
| 
 
9ff9bcbc2341
virtual Pure is single-threaded to avoid confusion with multiple thread farms etc.;
 
wenzelm 
parents: 
62925 
diff
changeset
 | 
41  | 
else if m > 0 then m  | 
| 79613 | 42  | 
else Int.min (num_processors (), 8);  | 
| 62359 | 43  | 
|
| 62925 | 44  | 
val max_threads_state = ref 1;  | 
45  | 
||
46  | 
in  | 
|
| 
32286
 
1fb5db48002d
added Multithreading.sync_wait, which turns enabled interrupts to sync ones, to ensure that wait will reaquire its lock when interrupted;
 
wenzelm 
parents: 
32186 
diff
changeset
 | 
47  | 
|
| 62925 | 48  | 
fun max_threads () = ! max_threads_state;  | 
49  | 
fun max_threads_update m = max_threads_state := max_threads_result m;  | 
|
| 67659 | 50  | 
|
| 62925 | 51  | 
end;  | 
| 
28187
 
4062882c7df3
proper values of no_interrupts, regular_interrupts;
 
wenzelm 
parents: 
28169 
diff
changeset
 | 
52  | 
|
| 62359 | 53  | 
|
| 68025 | 54  | 
(* parallel_proofs *)  | 
55  | 
||
56  | 
val parallel_proofs = ref 1;  | 
|
57  | 
||
58  | 
||
| 62359 | 59  | 
(* synchronous wait *)  | 
| 
41713
 
a21084741b37
added Multithreading.interrupted (cf. java.lang.Thread.interrupted);
 
wenzelm 
parents: 
39616 
diff
changeset
 | 
60  | 
|
| 64276 | 61  | 
fun sync_wait time cond lock =  | 
| 62923 | 62  | 
Thread_Attributes.with_attributes  | 
| 
64557
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
64276 
diff
changeset
 | 
63  | 
(Thread_Attributes.sync_interrupts (Thread_Attributes.get_attributes ()))  | 
| 62359 | 64  | 
(fn _ =>  | 
65  | 
(case time of  | 
|
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
68130 
diff
changeset
 | 
66  | 
SOME t => Exn.Res (Thread.ConditionVar.waitUntil (cond, lock, t))  | 
| 
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
68130 
diff
changeset
 | 
67  | 
| NONE => (Thread.ConditionVar.wait (cond, lock); Exn.Res true))  | 
| 62359 | 68  | 
handle exn => Exn.Exn exn);  | 
| 
32295
 
400cc493d466
renamed Multithreading.regular_interrupts to Multithreading.public_interrupts;
 
wenzelm 
parents: 
32286 
diff
changeset
 | 
69  | 
|
| 
 
400cc493d466
renamed Multithreading.regular_interrupts to Multithreading.public_interrupts;
 
wenzelm 
parents: 
32286 
diff
changeset
 | 
70  | 
|
| 
 
400cc493d466
renamed Multithreading.regular_interrupts to Multithreading.public_interrupts;
 
wenzelm 
parents: 
32286 
diff
changeset
 | 
71  | 
(* tracing *)  | 
| 
30612
 
cb6421b6a18f
future_job: do not inherit attributes, but enforce restricted interrupts -- attempt to prevent interrupt race conditions;
 
wenzelm 
parents: 
29564 
diff
changeset
 | 
72  | 
|
| 62359 | 73  | 
val trace = ref 0;  | 
74  | 
||
75  | 
fun tracing level msg =  | 
|
| 64562 | 76  | 
if ! trace < level then ()  | 
| 78720 | 77  | 
else Thread_Attributes.uninterruptible_body (fn _ =>  | 
| 62359 | 78  | 
    (TextIO.output (TextIO.stdErr, (">>> " ^ msg () ^ "\n")); TextIO.flushOut TextIO.stdErr)
 | 
| 78720 | 79  | 
handle _ (*sic*) => ());  | 
| 62359 | 80  | 
|
81  | 
fun tracing_time detailed time =  | 
|
82  | 
tracing  | 
|
83  | 
(if not detailed then 5  | 
|
| 62826 | 84  | 
else if time >= seconds 1.0 then 1  | 
85  | 
else if time >= seconds 0.1 then 2  | 
|
86  | 
else if time >= seconds 0.01 then 3  | 
|
87  | 
else if time >= seconds 0.001 then 4 else 5);  | 
|
| 62359 | 88  | 
|
| 25735 | 89  | 
|
| 
59180
 
c0fa3b3bdabd
discontinued central critical sections: NAMED_CRITICAL / CRITICAL;
 
wenzelm 
parents: 
59054 
diff
changeset
 | 
90  | 
(* synchronized evaluation *)  | 
| 
24690
 
c661dae1070a
renamed ML-Systems/multithreading_dummy.ML to ML-Systems/multithreading.ML;
 
wenzelm 
parents:  
diff
changeset
 | 
91  | 
|
| 62359 | 92  | 
fun synchronized name lock e =  | 
| 78720 | 93  | 
Exn.release (Thread_Attributes.uninterruptible_body (fn run =>  | 
| 64563 | 94  | 
if ! trace > 0 then  | 
95  | 
let  | 
|
96  | 
val immediate =  | 
|
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
68130 
diff
changeset
 | 
97  | 
if Thread.Mutex.trylock lock then true  | 
| 64563 | 98  | 
else  | 
99  | 
let  | 
|
100  | 
val _ = tracing 5 (fn () => name ^ ": locking ...");  | 
|
101  | 
val timer = Timer.startRealTimer ();  | 
|
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
68130 
diff
changeset
 | 
102  | 
val _ = Thread.Mutex.lock lock;  | 
| 64563 | 103  | 
val time = Timer.checkRealTimer timer;  | 
104  | 
val _ = tracing_time true time (fn () => name ^ ": locked after " ^ Time.toString time);  | 
|
105  | 
in false end;  | 
|
| 78716 | 106  | 
val result = Exn.capture0 (run e) ();  | 
| 64563 | 107  | 
val _ = if immediate then () else tracing 5 (fn () => name ^ ": unlocking ...");  | 
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
68130 
diff
changeset
 | 
108  | 
val _ = Thread.Mutex.unlock lock;  | 
| 64563 | 109  | 
in result end  | 
110  | 
else  | 
|
111  | 
let  | 
|
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
68130 
diff
changeset
 | 
112  | 
val _ = Thread.Mutex.lock lock;  | 
| 78716 | 113  | 
val result = Exn.capture0 (run e) ();  | 
| 
78650
 
47d0c333d155
clarified signature: retain original Poly/ML names Thread.Thread, Thread.Mutex, Thread.ConditionVar and de-emphasize them for Isabelle/ML;
 
wenzelm 
parents: 
68130 
diff
changeset
 | 
114  | 
val _ = Thread.Mutex.unlock lock;  | 
| 78720 | 115  | 
in result end));  | 
| 
59054
 
61b723761dff
load simple_thread.ML later, such that it benefits from redefined print_exception_trace;
 
wenzelm 
parents: 
54717 
diff
changeset
 | 
116  | 
|
| 
28123
 
53cd972d7db9
provide dummy thread structures, including proper Thread.getLocal/setLocal;
 
wenzelm 
parents: 
26082 
diff
changeset
 | 
117  | 
end;  |