author | wenzelm |
Tue, 14 Oct 2008 15:16:11 +0200 | |
changeset 28586 | d238b83ba3fc |
parent 28582 | c269a3045fdf |
child 28589 | 581b2ab9827a |
permissions | -rw-r--r-- |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
1 |
(* Title: HOL/Tools/atp_manager.ML |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
3 |
Author: Fabian Immler, TU Muenchen |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
4 |
|
28571 | 5 |
ATP threads are registered here. |
6 |
Threads with the same birth-time are seen as one group. |
|
7 |
All threads of a group are killed when one thread of it has been successful, |
|
8 |
or after a certain time, |
|
9 |
or when the maximum number of threads exceeds; then the oldest thread is killed. |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
10 |
*) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
11 |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
12 |
signature ATP_MANAGER = |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
13 |
sig |
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
14 |
val get_atps: unit -> string |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
15 |
val set_atps: string -> unit |
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
16 |
val get_max_atps: unit -> int |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
17 |
val set_max_atps: int -> unit |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
18 |
val get_timeout: unit -> int |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
19 |
val set_timeout: int -> unit |
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
20 |
val kill: unit -> unit |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
21 |
val info: unit -> unit |
28571 | 22 |
val atp_thread: (unit -> 'a option) -> ('a -> string) -> Thread.thread |
23 |
val add_prover: string -> (int -> Proof.state -> Thread.thread) -> theory -> theory |
|
28484 | 24 |
val print_provers: theory -> unit |
28571 | 25 |
val sledgehammer: string list -> Proof.state -> unit |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
26 |
end; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
27 |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
28 |
structure AtpManager: ATP_MANAGER = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
29 |
struct |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
30 |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
31 |
(** preferences **) |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
32 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
33 |
local |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
34 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
35 |
val atps = ref "e"; |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
36 |
val max_atps = ref 5; (* ~1 means infinite number of atps *) |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
37 |
val timeout = ref 60; |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
38 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
39 |
in |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
40 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
41 |
fun get_atps () = CRITICAL (fn () => ! atps); |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
42 |
fun set_atps str = CRITICAL (fn () => atps := str); |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
43 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
44 |
fun get_max_atps () = CRITICAL (fn () => ! max_atps); |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
45 |
fun set_max_atps number = CRITICAL (fn () => max_atps := number); |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
46 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
47 |
fun get_timeout () = CRITICAL (fn () => ! timeout); |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
48 |
fun set_timeout time = CRITICAL (fn () => timeout := time); |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
49 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
50 |
val _ = |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
51 |
ProofGeneralPgip.add_preference "Proof" |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
52 |
(Preferences.string_pref atps |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
53 |
"ATP: provers" "Default automatic provers (separated by whitespace)") |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
54 |
handle ERROR _ => warning "Preference already exists"; |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
55 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
56 |
val _ = ProofGeneralPgip.add_preference "Proof" |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
57 |
(Preferences.int_pref max_atps |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
58 |
"ATP: maximum number" "How many provers may run in parallel") |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
59 |
handle ERROR _ => warning "Preference already exists"; |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
60 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
61 |
val _ = ProofGeneralPgip.add_preference "Proof" |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
62 |
(Preferences.int_pref timeout |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
63 |
"ATP: timeout" "ATPs will be interrupted after this time (in seconds)") |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
64 |
handle ERROR _ => warning "Preference already exists"; |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
65 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
66 |
end; |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
67 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
68 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
69 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
70 |
(** thread management **) |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
71 |
|
28582 | 72 |
(* data structures over threads *) |
73 |
||
74 |
structure ThreadHeap = HeapFun |
|
75 |
( |
|
76 |
type elem = Time.time * Thread.thread; |
|
77 |
fun ord ((a, _), (b, _)) = Time.compare (a, b); |
|
78 |
) |
|
79 |
||
80 |
val lookup_thread = AList.lookup Thread.equal; |
|
81 |
val delete_thread = AList.delete Thread.equal; |
|
82 |
val update_thread = AList.update Thread.equal; |
|
83 |
||
84 |
||
85 |
(* state of thread manager *) |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
86 |
|
28582 | 87 |
datatype T = State of |
88 |
{timeout_heap: ThreadHeap.T, |
|
89 |
oldest_heap: ThreadHeap.T, |
|
90 |
active: (Thread.thread * (Time.time * Time.time * string)) list, |
|
91 |
cancelling: (Thread.thread * (Time.time * Time.time * string)) list}; |
|
92 |
||
93 |
fun make_state timeout_heap oldest_heap active cancelling = |
|
94 |
State {timeout_heap = timeout_heap, oldest_heap = oldest_heap, |
|
95 |
active = active, cancelling = cancelling}; |
|
96 |
||
97 |
val state = Synchronized.var "atp_manager" (make_state ThreadHeap.empty ThreadHeap.empty [] []); |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
98 |
|
28582 | 99 |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
100 |
(* the managing thread *) |
28582 | 101 |
|
102 |
(*watches over running threads and interrupts them if required*) |
|
103 |
val managing_thread = ref (NONE: Thread.thread option); |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
104 |
|
28582 | 105 |
|
106 |
(* unregister thread from thread manager -- move to cancelling *) |
|
107 |
||
108 |
fun unregister success message thread = Synchronized.change_result state |
|
109 |
(fn State {timeout_heap, oldest_heap, active, cancelling} => |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
110 |
let |
28582 | 111 |
val info = lookup_thread active thread |
112 |
||
113 |
(* get birthtime of unregistering thread if successful - for group-killing*) |
|
114 |
val birthtime = case info of NONE => Time.zeroTime |
|
115 |
| SOME (tb, _, _) => if success then tb else Time.zeroTime |
|
116 |
||
117 |
(* move unregistering thread to cancelling *) |
|
118 |
val active' = delete_thread thread active |
|
119 |
val cancelling' = case info of NONE => cancelling |
|
120 |
| SOME (tb, _, desc) => update_thread (thread, (tb, Time.now (), desc)) cancelling |
|
121 |
||
122 |
(* move all threads of the same group to cancelling *) |
|
123 |
val group_threads = active |> map_filter (fn (th, (tb, _, desc)) => |
|
124 |
if tb = birthtime then SOME (th, (tb, Time.now (), desc)) else NONE) |
|
125 |
val active'' = filter_out (fn (_, (tb, _, _)) => tb = birthtime) active' |
|
126 |
val cancelling'' = append group_threads cancelling' |
|
127 |
||
128 |
(* message for user *) |
|
129 |
val message' = case info of NONE => "" |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
130 |
| SOME (_, _, desc) => "Sledgehammer: " ^ desc ^ "\n" ^ message ^ |
28582 | 131 |
(if null group_threads then "" |
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
132 |
else "\nInterrupted " ^ string_of_int (length group_threads - 1) ^ " other group members") |
28582 | 133 |
in (message', make_state timeout_heap oldest_heap active'' cancelling'') end); |
134 |
||
135 |
||
136 |
(* start a watching thread which runs forever -- only one may exist *) |
|
137 |
||
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
138 |
fun check_thread_manager () = CRITICAL (fn () => |
28582 | 139 |
if (case ! managing_thread of SOME thread => Thread.isActive thread | NONE => false) |
140 |
then () else managing_thread := SOME (SimpleThread.fork false (fn () => |
|
28571 | 141 |
let |
28582 | 142 |
val min_wait_time = Time.fromMilliseconds 300 |
143 |
val max_wait_time = Time.fromSeconds 10 |
|
144 |
||
145 |
(* wait for next thread to cancel, or maximum*) |
|
146 |
fun time_limit (State {timeout_heap, ...}) = |
|
147 |
(case try ThreadHeap.min timeout_heap of |
|
148 |
NONE => SOME (Time.+ (Time.now (), max_wait_time)) |
|
149 |
| SOME (time, _) => SOME time) |
|
150 |
||
151 |
(* action: cancel find threads whose timeout is reached, and interrupt cancelling threads *) |
|
152 |
fun action (State {timeout_heap, oldest_heap, active, cancelling}) = |
|
153 |
let val (timeout_threads, timeout_heap') = |
|
154 |
ThreadHeap.upto (Time.now (), Thread.self ()) timeout_heap |
|
155 |
in |
|
156 |
if null timeout_threads andalso null cancelling then NONE |
|
157 |
else |
|
158 |
let |
|
159 |
val _ = List.app (SimpleThread.interrupt o #1) cancelling |
|
160 |
val cancelling' = filter (Thread.isActive o #1) cancelling |
|
161 |
val state' = make_state timeout_heap' oldest_heap active cancelling' |
|
162 |
in SOME (map #2 timeout_threads, state') end |
|
163 |
end |
|
164 |
in |
|
165 |
while true do |
|
166 |
((* cancel threads found by action *) |
|
167 |
Synchronized.timed_access state time_limit action |
|
168 |
|> these |
|
169 |
|> List.app (priority o unregister false "Interrupted (reached timeout)"); |
|
28571 | 170 |
(* give threads time to respond to interrupt *) |
171 |
OS.Process.sleep min_wait_time) |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
172 |
end))); |
28582 | 173 |
|
174 |
||
175 |
(* thread is registered here by sledgehammer *) |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
176 |
|
28582 | 177 |
fun register birthtime deadtime (thread, desc) = |
178 |
(check_thread_manager (); |
|
179 |
Synchronized.change state (fn State {timeout_heap, oldest_heap, active, cancelling} => |
|
180 |
let |
|
181 |
val timeout_heap' = ThreadHeap.insert (deadtime, thread) timeout_heap |
|
182 |
val oldest_heap' = ThreadHeap.insert (birthtime, thread) oldest_heap |
|
183 |
val active' = update_thread (thread, (birthtime, deadtime, desc)) active |
|
184 |
in make_state timeout_heap' oldest_heap' active' cancelling end)); |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
185 |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
186 |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
187 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
188 |
(** user commands **) |
28582 | 189 |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
190 |
(* kill: move all threads to cancelling *) |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
191 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
192 |
fun kill () = Synchronized.change state |
28582 | 193 |
(fn State {timeout_heap, oldest_heap, active, cancelling} => |
194 |
let val formerly_active = map (fn (th, (tb, _, desc)) => (th, (tb, Time.now (), desc))) active |
|
195 |
in make_state timeout_heap oldest_heap [] (formerly_active @ cancelling) end); |
|
196 |
||
197 |
||
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
198 |
(* info: information on running threads *) |
28582 | 199 |
|
200 |
fun info () = |
|
201 |
let |
|
202 |
val State {timeout_heap, oldest_heap, active, cancelling} = Synchronized.value state |
|
28571 | 203 |
fun running_info (_, (birth_time, dead_time, desc)) = "Running: " |
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
204 |
^ ((string_of_int o Time.toSeconds) (Time.- (Time.now (), birth_time))) |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
205 |
^ " s -- " |
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
206 |
^ ((string_of_int o Time.toSeconds) (Time.- (dead_time, Time.now ()))) |
28571 | 207 |
^ " s to live:\n" ^ desc |
208 |
fun cancelling_info (_, (_, dead_time, desc)) = "Trying to interrupt thread since " |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
209 |
^ (string_of_int o Time.toSeconds) (Time.- (Time.now (), dead_time)) |
28571 | 210 |
^ " s:\n" ^ desc |
211 |
val running = if null active then "No ATPs running." |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
212 |
else space_implode "\n\n" ("--- RUNNING ATPs ---" :: |
28571 | 213 |
(map (fn entry => running_info entry) active)) |
214 |
val interrupting = if null cancelling then "" |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
215 |
else space_implode "\n\n" ("--- TRYING TO INTERRUPT FOLLOWING ATPs ---" :: |
28571 | 216 |
(map (fn entry => cancelling_info entry) cancelling)) |
28582 | 217 |
in writeln (running ^ "\n" ^ interrupting) end; |
218 |
||
219 |
||
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
220 |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
221 |
(** The Sledgehammer **) |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
222 |
|
28582 | 223 |
(* named provers *) |
28484 | 224 |
|
28582 | 225 |
fun err_dup_prover name = error ("Duplicate prover: " ^ quote name); |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
226 |
|
28582 | 227 |
structure Provers = TheoryDataFun |
228 |
( |
|
229 |
type T = ((int -> Proof.state -> Thread.thread) * stamp) Symtab.table |
|
230 |
val empty = Symtab.empty |
|
231 |
val copy = I |
|
232 |
val extend = I |
|
233 |
fun merge _ tabs : T = Symtab.merge (eq_snd op =) tabs |
|
234 |
handle Symtab.DUP dup => err_dup_prover dup; |
|
235 |
); |
|
28484 | 236 |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
237 |
fun add_prover name prover_fn thy = |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
238 |
Provers.map (Symtab.update_new (name, (prover_fn, stamp ()))) thy |
28582 | 239 |
handle Symtab.DUP dup => err_dup_prover dup; |
240 |
||
241 |
fun print_provers thy = Pretty.writeln |
|
242 |
(Pretty.strs ("external provers:" :: sort_strings (Symtab.keys (Provers.get thy)))); |
|
28571 | 243 |
|
28582 | 244 |
fun prover_desc state subgoal name = |
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
245 |
let val (ctxt, (_, goal)) = Proof.get_goal state in |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
246 |
"external prover " ^ quote name ^ " for subgoal " ^ string_of_int subgoal ^ ":\n" ^ |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
247 |
Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal subgoal)) |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
248 |
end; |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
249 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
250 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
251 |
(* thread wrapping an atp-call *) |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
252 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
253 |
fun atp_thread call_prover produce_answer = |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
254 |
SimpleThread.fork true (fn () => |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
255 |
let |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
256 |
val result = call_prover () |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
257 |
val message = case result of NONE => "Failed." |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
258 |
| SOME result => "Try this command: " ^ produce_answer result |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
259 |
in priority (unregister (is_some result) message (Thread.self ())) |
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
260 |
end handle Interrupt => ()); |
28582 | 261 |
|
262 |
fun run_prover state subgoal name = |
|
263 |
(case Symtab.lookup (Provers.get (Proof.theory_of state)) name of |
|
264 |
NONE => (warning ("Unknown external prover: " ^ quote name); NONE) |
|
265 |
| SOME (prover_fn, _) => SOME (prover_fn subgoal state, prover_desc state subgoal name)); |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
266 |
|
28484 | 267 |
|
28582 | 268 |
(* kill excessive atp threads *) |
28484 | 269 |
|
28582 | 270 |
local |
28484 | 271 |
|
28582 | 272 |
fun excessive_atps active = |
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
273 |
let val max = get_max_atps () |
28582 | 274 |
in length active > max andalso max > ~1 end; |
28484 | 275 |
|
28582 | 276 |
fun kill_oldest () = |
277 |
let exception Unchanged in |
|
278 |
Synchronized.change_result state (fn State {timeout_heap, oldest_heap, active, cancelling} => |
|
279 |
if ThreadHeap.is_empty oldest_heap orelse not (excessive_atps active) |
|
280 |
then raise Unchanged |
|
281 |
else |
|
282 |
let val ((_, oldest_thread), oldest_heap') = ThreadHeap.min_elem oldest_heap |
|
283 |
in (oldest_thread, make_state timeout_heap oldest_heap' active cancelling) end) |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
284 |
|> (priority o unregister false "Interrupted (maximum number of ATPs exceeded)") |
28582 | 285 |
handle Unchanged => () |
28484 | 286 |
end; |
287 |
||
28582 | 288 |
in |
289 |
||
290 |
fun kill_excessive () = |
|
291 |
let val State {active, ...} = Synchronized.value state |
|
292 |
in if excessive_atps active then (kill_oldest (); kill_excessive ()) else () end; |
|
293 |
||
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
294 |
end; |
28582 | 295 |
|
296 |
||
297 |
(* sledghammer for first subgoal *) |
|
298 |
||
299 |
fun sledgehammer names proof_state = |
|
300 |
let |
|
301 |
val proverids = |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
302 |
if null names then String.tokens (Symbol.is_ascii_blank o String.str) (get_atps ()) |
28582 | 303 |
else names |
304 |
val threads_names = map_filter (run_prover proof_state 1) proverids |
|
305 |
val birthtime = Time.now () |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
306 |
val deadtime = Time.+ (Time.now (), Time.fromSeconds (get_timeout ())) |
28582 | 307 |
val _ = List.app (register birthtime deadtime) threads_names |
308 |
val _ = kill_excessive () |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
309 |
in () end; |
28582 | 310 |
|
311 |
||
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
312 |
|
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
313 |
(** Isar command syntax **) |
28582 | 314 |
|
315 |
local structure K = OuterKeyword and P = OuterParse in |
|
316 |
||
317 |
val _ = |
|
318 |
OuterSyntax.improper_command "atp_kill" "kill all managed provers" K.diag |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
319 |
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill)); |
28582 | 320 |
|
321 |
val _ = |
|
322 |
OuterSyntax.improper_command "atp_info" "print information about managed provers" K.diag |
|
323 |
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative info)); |
|
324 |
||
325 |
val _ = |
|
326 |
OuterSyntax.improper_command "print_atps" "print external provers" K.diag |
|
327 |
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o |
|
328 |
Toplevel.keep (print_provers o Toplevel.theory_of))); |
|
329 |
||
330 |
val _ = |
|
331 |
OuterSyntax.command "sledgehammer" "call all automatic theorem provers" K.diag |
|
332 |
(Scan.repeat P.xname >> (fn names => Toplevel.no_timing o Toplevel.unknown_proof o |
|
333 |
Toplevel.keep ((sledgehammer names) o Toplevel.proof_of))); |
|
334 |
||
335 |
end; |
|
336 |
||
337 |
end; |