author | nipkow |
Thu, 26 Jun 2025 17:30:33 +0200 | |
changeset 82772 | 59b937edcff8 |
parent 80910 | 406a85a25189 |
permissions | -rw-r--r-- |
59471
ca459352d8c5
more explicit indication of Async_Manager_Legacy as Proof General legacy;
wenzelm
parents:
59468
diff
changeset
|
1 |
(* Title: HOL/Tools/Sledgehammer/async_manager_legacy.ML |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
2 |
Author: Fabian Immler, TU Muenchen |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
3 |
Author: Makarius |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
4 |
Author: Jasmin Blanchette, TU Muenchen |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
5 |
|
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
6 |
Central manager for asynchronous diagnosis tool threads. |
58894 | 7 |
|
8 |
Proof General legacy! |
|
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
9 |
*) |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
10 |
|
59471
ca459352d8c5
more explicit indication of Async_Manager_Legacy as Proof General legacy;
wenzelm
parents:
59468
diff
changeset
|
11 |
signature ASYNC_MANAGER_LEGACY = |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
12 |
sig |
57556 | 13 |
val thread : string -> Time.time -> Time.time -> string * string -> (unit -> bool * string) -> |
14 |
unit |
|
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
15 |
val has_running_threads : string -> bool |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
16 |
end; |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
17 |
|
59471
ca459352d8c5
more explicit indication of Async_Manager_Legacy as Proof General legacy;
wenzelm
parents:
59468
diff
changeset
|
18 |
structure Async_Manager_Legacy : ASYNC_MANAGER_LEGACY = |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
19 |
struct |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
20 |
|
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
21 |
fun implode_message (workers, work) = |
80910 | 22 |
implode_space (Try.serial_commas "and" workers) ^ work |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
23 |
|
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
24 |
structure Thread_Heap = Heap |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
25 |
( |
78648 | 26 |
type elem = Time.time * Isabelle_Thread.T |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
27 |
fun ord ((a, _), (b, _)) = Time.compare (a, b) |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
28 |
) |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
29 |
|
78648 | 30 |
fun lookup_thread xs = AList.lookup Isabelle_Thread.equal xs |
31 |
fun delete_thread xs = AList.delete Isabelle_Thread.equal xs |
|
32 |
fun update_thread xs = AList.update Isabelle_Thread.equal xs |
|
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
33 |
|
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
34 |
type state = |
78648 | 35 |
{manager: Isabelle_Thread.T option, |
37585 | 36 |
timeout_heap: Thread_Heap.T, |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
37 |
active: |
78648 | 38 |
(Isabelle_Thread.T |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
39 |
* (string * Time.time * Time.time * (string * string))) list, |
78648 | 40 |
canceling: (Isabelle_Thread.T * (string * Time.time * (string * string))) list, |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
41 |
messages: (bool * (string * (string * string))) list} |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
42 |
|
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
43 |
fun make_state manager timeout_heap active canceling messages : state = |
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
44 |
{manager = manager, timeout_heap = timeout_heap, active = active, canceling = canceling, |
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
45 |
messages = messages} |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
46 |
|
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
47 |
val global_state = Synchronized.var "async_manager" (make_state NONE Thread_Heap.empty [] [] []) |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
48 |
|
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
49 |
fun unregister (urgent, message) thread = |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
50 |
Synchronized.change global_state |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
51 |
(fn state as {manager, timeout_heap, active, canceling, messages} => |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
52 |
(case lookup_thread active thread of |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
53 |
SOME (tool, _, _, desc as (worker, its_desc)) => |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
54 |
let |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
55 |
val active' = delete_thread thread active |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
56 |
val now = Time.now () |
37585 | 57 |
val canceling' = (thread, (tool, now, desc)) :: canceling |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
58 |
val message' = |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
59 |
(worker, its_desc ^ (if message = "" then "" else "\n" ^ message)) |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
60 |
val messages' = (urgent, (tool, message')) :: messages |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
61 |
in make_state manager timeout_heap active' canceling' messages' end |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
62 |
| NONE => state)) |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
63 |
|
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
64 |
val min_wait_time = seconds 0.3 |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
65 |
val max_wait_time = seconds 10.0 |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
66 |
|
37585 | 67 |
fun print_new_messages () = |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
68 |
Synchronized.change_result global_state |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
69 |
(fn {manager, timeout_heap, active, canceling, messages} => |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
70 |
messages |
43055 | 71 |
|> List.partition |
72 |
(fn (urgent, _) => |
|
73 |
(null active andalso null canceling) orelse urgent) |
|
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
74 |
||> make_state manager timeout_heap active canceling) |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
75 |
|> map (fn (_, (tool, (worker, work))) => ((tool, work), worker)) |
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
76 |
|> AList.group (op =) |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
77 |
|> List.app (fn ((_, ""), _) => () |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
78 |
| ((tool, work), workers) => |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
79 |
tool ^ ": " ^ |
47945 | 80 |
implode_message (workers |> sort_distinct string_ord, work) |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
81 |
|> writeln) |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
82 |
|
39006
a02cb5717677
remove time information in output, since it's confusing anyway
blanchet
parents:
38990
diff
changeset
|
83 |
fun check_thread_manager () = Synchronized.change global_state |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
84 |
(fn state as {manager, timeout_heap, active, canceling, messages} => |
78648 | 85 |
if (case manager of SOME thread => Isabelle_Thread.is_active thread | NONE => false) then state |
78753 | 86 |
else let val manager = SOME (Isabelle_Thread.fork (Isabelle_Thread.params "async_manager") (fn () => |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
87 |
let |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
88 |
fun time_limit timeout_heap = |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
89 |
(case try Thread_Heap.min timeout_heap of |
62826 | 90 |
NONE => Time.now () + max_wait_time |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
91 |
| SOME (time, _) => time) |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
92 |
|
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
93 |
(*action: find threads whose timeout is reached, and interrupt canceling threads*) |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
94 |
fun action {manager, timeout_heap, active, canceling, messages} = |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
95 |
let val (timeout_threads, timeout_heap') = |
78648 | 96 |
Thread_Heap.upto (Time.now (), Isabelle_Thread.self ()) timeout_heap |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
97 |
in |
37585 | 98 |
if null timeout_threads andalso null canceling then |
99 |
NONE |
|
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
100 |
else |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
101 |
let |
78787 | 102 |
val _ = List.app (Isabelle_Thread.interrupt_thread o #1) canceling |
78648 | 103 |
val canceling' = filter (Isabelle_Thread.is_active o #1) canceling |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
104 |
val state' = make_state manager timeout_heap' active canceling' messages |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
105 |
in SOME (map #2 timeout_threads, state') end |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
106 |
end |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
107 |
in |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
108 |
while Synchronized.change_result global_state |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
109 |
(fn state as {timeout_heap, active, canceling, messages, ...} => |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
110 |
if null active andalso null canceling andalso null messages |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
111 |
then (false, make_state NONE timeout_heap active canceling messages) |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
112 |
else (true, state)) |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
113 |
do |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
114 |
(Synchronized.timed_access global_state |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
115 |
(SOME o time_limit o #timeout_heap) action |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
116 |
|> these |
63692 | 117 |
|> List.app (unregister (false, "Timed out")); |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
118 |
print_new_messages (); |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
119 |
(* give threads some time to respond to interrupt *) |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
120 |
OS.Process.sleep min_wait_time) |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
121 |
end)) |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
122 |
in make_state manager timeout_heap active canceling messages end) |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
123 |
|
39006
a02cb5717677
remove time information in output, since it's confusing anyway
blanchet
parents:
38990
diff
changeset
|
124 |
fun register tool birth_time death_time desc thread = |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
125 |
(Synchronized.change global_state |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
126 |
(fn {manager, timeout_heap, active, canceling, messages} => |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
127 |
let |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
128 |
val timeout_heap' = Thread_Heap.insert (death_time, thread) timeout_heap |
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
129 |
val active' = update_thread (thread, (tool, birth_time, death_time, desc)) active |
61312
6d779a71086d
further reduced dependency on legacy async thread manager
blanchet
parents:
60764
diff
changeset
|
130 |
val state' = make_state manager timeout_heap' active' canceling messages |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
131 |
in state' end); |
39006
a02cb5717677
remove time information in output, since it's confusing anyway
blanchet
parents:
38990
diff
changeset
|
132 |
check_thread_manager ()) |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
133 |
|
52048
9003b293e775
tuned signature -- emphasize thread creation here;
wenzelm
parents:
48319
diff
changeset
|
134 |
fun thread tool birth_time death_time desc f = |
78754 | 135 |
(Isabelle_Thread.fork (Isabelle_Thread.params "async_worker" |> Isabelle_Thread.interrupts) |
37584 | 136 |
(fn () => |
137 |
let |
|
78648 | 138 |
val self = Isabelle_Thread.self () |
39006
a02cb5717677
remove time information in output, since it's confusing anyway
blanchet
parents:
38990
diff
changeset
|
139 |
val _ = register tool birth_time death_time desc self |
43005
c96f06bffd90
merge timeout messages from several ATPs into one message to avoid clutter
blanchet
parents:
41042
diff
changeset
|
140 |
in unregister (f ()) self end); |
37584 | 141 |
()) |
142 |
||
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
143 |
fun has_running_threads tool = |
57556 | 144 |
exists (fn (_, (tool', _, _, _)) => tool' = tool) (#active (Synchronized.value global_state)) |
48319
340187063d84
use async manager to manage MaSh learners to make sure they get killed cleanly
blanchet
parents:
47945
diff
changeset
|
145 |
|
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
diff
changeset
|
146 |
end; |