| author | haftmann | 
| Tue, 02 Jun 2009 15:53:07 +0200 | |
| changeset 31378 | d1cbf6393964 | 
| parent 31037 | ac8669134e7a | 
| child 31368 | 763f4b0fd579 | 
| 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 | Author: Fabian Immler, TU Muenchen | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 3 | |
| 28571 | 4 | ATP threads are registered here. | 
| 5 | Threads with the same birth-time are seen as one group. | |
| 6 | All threads of a group are killed when one thread of it has been successful, | |
| 7 | or after a certain time, | |
| 8 | 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 | 9 | *) | 
| 
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 | signature ATP_MANAGER = | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 12 | sig | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 13 | val get_atps: unit -> string | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 14 | val set_atps: string -> unit | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 15 | val get_max_atps: unit -> int | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 16 | val set_max_atps: int -> unit | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 17 | val get_timeout: unit -> int | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 18 | val set_timeout: int -> unit | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 19 | val kill: unit -> unit | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 20 | val info: unit -> unit | 
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 21 | val messages: int option -> unit | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 22 | type prover = int -> (thm * (string * int)) list option -> string -> int -> | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 23 | Proof.context * (thm list * thm) -> bool * string * string * string vector | 
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 24 | val add_prover: string -> prover -> theory -> theory | 
| 28484 | 25 | val print_provers: theory -> unit | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 26 | val get_prover: string -> theory -> prover option | 
| 28571 | 27 | val sledgehammer: string list -> Proof.state -> unit | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 28 | end; | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 29 | |
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 30 | structure AtpManager: ATP_MANAGER = | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 31 | struct | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 32 | |
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 33 | (** preferences **) | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 34 | |
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 35 | val message_store_limit = 20; | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 36 | val message_display_limit = 5; | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 37 | |
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 38 | local | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 39 | |
| 29598 | 40 | val atps = ref "e remote_vampire"; | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 41 | val max_atps = ref 5; (* ~1 means infinite number of atps *) | 
| 29784 | 42 | val timeout = ref 60; | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 43 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 44 | in | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 45 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 46 | fun get_atps () = CRITICAL (fn () => ! atps); | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 47 | fun set_atps str = CRITICAL (fn () => atps := str); | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 48 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 49 | fun get_max_atps () = CRITICAL (fn () => ! max_atps); | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 50 | fun set_max_atps number = CRITICAL (fn () => max_atps := number); | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 51 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 52 | fun get_timeout () = CRITICAL (fn () => ! timeout); | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 53 | fun set_timeout time = CRITICAL (fn () => timeout := time); | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 54 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 55 | val _ = | 
| 30982 | 56 | ProofGeneralPgip.add_preference Preferences.category_proof | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 57 | (Preferences.string_pref atps | 
| 28589 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 58 | "ATP: provers" "Default automatic provers (separated by whitespace)"); | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 59 | |
| 30982 | 60 | val _ = | 
| 61 | ProofGeneralPgip.add_preference Preferences.category_proof | |
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 62 | (Preferences.int_pref max_atps | 
| 28589 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 63 | "ATP: maximum number" "How many provers may run in parallel"); | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 64 | |
| 30982 | 65 | val _ = | 
| 66 | ProofGeneralPgip.add_preference Preferences.category_proof | |
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 67 | (Preferences.int_pref timeout | 
| 28589 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 68 | "ATP: timeout" "ATPs will be interrupted after this time (in seconds)"); | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 69 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 70 | end; | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 71 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 72 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 73 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 74 | (** thread management **) | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 75 | |
| 28582 | 76 | (* data structures over threads *) | 
| 77 | ||
| 78 | structure ThreadHeap = HeapFun | |
| 79 | ( | |
| 80 | type elem = Time.time * Thread.thread; | |
| 81 | fun ord ((a, _), (b, _)) = Time.compare (a, b); | |
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 82 | ); | 
| 28582 | 83 | |
| 84 | val lookup_thread = AList.lookup Thread.equal; | |
| 85 | val delete_thread = AList.delete Thread.equal; | |
| 86 | val update_thread = AList.update Thread.equal; | |
| 87 | ||
| 88 | ||
| 89 | (* state of thread manager *) | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 90 | |
| 28582 | 91 | datatype T = State of | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 92 |  {managing_thread: Thread.thread option,
 | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 93 | timeout_heap: ThreadHeap.T, | 
| 28582 | 94 | oldest_heap: ThreadHeap.T, | 
| 95 | active: (Thread.thread * (Time.time * Time.time * string)) list, | |
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 96 | cancelling: (Thread.thread * (Time.time * Time.time * string)) list, | 
| 29620 | 97 | messages: string list, | 
| 98 | store: string list}; | |
| 28582 | 99 | |
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 100 | fun make_state managing_thread timeout_heap oldest_heap active cancelling messages store = | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 101 |   State {managing_thread = managing_thread, timeout_heap = timeout_heap, oldest_heap = oldest_heap,
 | 
| 29620 | 102 | active = active, cancelling = cancelling, messages = messages, store = store}; | 
| 28582 | 103 | |
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 104 | val state = Synchronized.var "atp_manager" | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 105 | (make_state NONE ThreadHeap.empty ThreadHeap.empty [] [] [] []); | 
| 28582 | 106 | |
| 29150 
8af5ee47f30c
unregister: refrain from cancelling self, which sometimes caused sledghammer to fail just before printing the result;
 wenzelm parents: 
29116diff
changeset | 107 | (* unregister thread *) | 
| 28582 | 108 | |
| 29620 | 109 | fun unregister (success, message) thread = Synchronized.change state | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 110 |   (fn state as State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
 | 
| 29150 
8af5ee47f30c
unregister: refrain from cancelling self, which sometimes caused sledghammer to fail just before printing the result;
 wenzelm parents: 
29116diff
changeset | 111 | (case lookup_thread active thread of | 
| 
8af5ee47f30c
unregister: refrain from cancelling self, which sometimes caused sledghammer to fail just before printing the result;
 wenzelm parents: 
29116diff
changeset | 112 | SOME (birthtime, _, description) => | 
| 
8af5ee47f30c
unregister: refrain from cancelling self, which sometimes caused sledghammer to fail just before printing the result;
 wenzelm parents: 
29116diff
changeset | 113 | let | 
| 
8af5ee47f30c
unregister: refrain from cancelling self, which sometimes caused sledghammer to fail just before printing the result;
 wenzelm parents: 
29116diff
changeset | 114 | val (group, active') = | 
| 
8af5ee47f30c
unregister: refrain from cancelling self, which sometimes caused sledghammer to fail just before printing the result;
 wenzelm parents: 
29116diff
changeset | 115 | if success then List.partition (fn (_, (tb, _, _)) => tb = birthtime) active | 
| 
8af5ee47f30c
unregister: refrain from cancelling self, which sometimes caused sledghammer to fail just before printing the result;
 wenzelm parents: 
29116diff
changeset | 116 | else List.partition (fn (th, _) => Thread.equal (th, thread)) active | 
| 28582 | 117 | |
| 29150 
8af5ee47f30c
unregister: refrain from cancelling self, which sometimes caused sledghammer to fail just before printing the result;
 wenzelm parents: 
29116diff
changeset | 118 | val now = Time.now () | 
| 
8af5ee47f30c
unregister: refrain from cancelling self, which sometimes caused sledghammer to fail just before printing the result;
 wenzelm parents: 
29116diff
changeset | 119 | val cancelling' = | 
| 29620 | 120 | fold (fn (th, (tb, _, desc)) => update_thread (th, (tb, now, desc))) group cancelling | 
| 28582 | 121 | |
| 29620 | 122 | val message' = description ^ "\n" ^ message ^ | 
| 29596 | 123 | (if length group <= 1 then "" | 
| 124 | else "\nInterrupted " ^ string_of_int (length group - 1) ^ " other group members") | |
| 29620 | 125 | val store' = message' :: | 
| 126 | (if length store <= message_store_limit then store | |
| 127 | else #1 (chop message_store_limit store)) | |
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 128 | in make_state | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 129 | managing_thread timeout_heap oldest_heap active' cancelling' (message' :: messages) store' | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 130 | end | 
| 30800 | 131 | | NONE => state)); | 
| 28582 | 132 | |
| 133 | ||
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 134 | (* kill excessive atp threads *) | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 135 | |
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 136 | fun excessive_atps active = | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 137 | let val max = get_max_atps () | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 138 | in length active > max andalso max > ~1 end; | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 139 | |
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 140 | local | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 141 | |
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 142 | fun kill_oldest () = | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 143 | let exception Unchanged in | 
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 144 | Synchronized.change_result state | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 145 |       (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
 | 
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 146 | if ThreadHeap.is_empty oldest_heap orelse not (excessive_atps active) | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 147 | then raise Unchanged | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 148 | else | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 149 | let val ((_, oldest_thread), oldest_heap') = ThreadHeap.min_elem oldest_heap | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 150 | in (oldest_thread, | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 151 | make_state managing_thread timeout_heap oldest_heap' active cancelling messages store) end) | 
| 29620 | 152 | |> unregister (false, "Interrupted (maximum number of ATPs exceeded)") | 
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 153 | handle Unchanged => () | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 154 | end; | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 155 | |
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 156 | in | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 157 | |
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 158 | fun kill_excessive () = | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 159 |   let val State {active, ...} = Synchronized.value state
 | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 160 | in if excessive_atps active then (kill_oldest (); kill_excessive ()) else () end; | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 161 | |
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 162 | end; | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 163 | |
| 29620 | 164 | fun print_new_messages () = | 
| 165 | let val to_print = Synchronized.change_result state | |
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 166 |     (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
 | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 167 | (messages, make_state managing_thread timeout_heap oldest_heap active cancelling [] store)) | 
| 30800 | 168 | in | 
| 169 | if null to_print then () | |
| 170 |     else priority ("Sledgehammer: " ^ space_implode "\n\n" to_print)
 | |
| 171 | end; | |
| 29620 | 172 | |
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 173 | |
| 30800 | 174 | (* start a watching thread -- only one may exist *) | 
| 28582 | 175 | |
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 176 | fun check_thread_manager () = Synchronized.change state | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 177 |   (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
 | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 178 | if (case managing_thread of SOME thread => Thread.isActive thread | NONE => false) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 179 | then make_state managing_thread timeout_heap oldest_heap active cancelling messages store | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 180 | else let val managing_thread = SOME (SimpleThread.fork false (fn () => | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 181 | let | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 182 | val min_wait_time = Time.fromMilliseconds 300 | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 183 | val max_wait_time = Time.fromSeconds 10 | 
| 28582 | 184 | |
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 185 | (* wait for next thread to cancel, or maximum*) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 186 |         fun time_limit (State {timeout_heap, ...}) =
 | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 187 | (case try ThreadHeap.min timeout_heap of | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 188 | NONE => SOME (Time.+ (Time.now (), max_wait_time)) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 189 | | SOME (time, _) => SOME time) | 
| 28582 | 190 | |
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 191 | (* action: find threads whose timeout is reached, and interrupt cancelling threads *) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 192 |         fun action (State {managing_thread, timeout_heap, oldest_heap, active, cancelling,
 | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 193 | messages, store}) = | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 194 | let val (timeout_threads, timeout_heap') = | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 195 | ThreadHeap.upto (Time.now (), Thread.self ()) timeout_heap | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 196 | in | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 197 | if null timeout_threads andalso null cancelling andalso not (excessive_atps active) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 198 | then NONE | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 199 | else | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 200 | let | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 201 | val _ = List.app (SimpleThread.interrupt o #1) cancelling | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 202 | val cancelling' = filter (Thread.isActive o #1) cancelling | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 203 | val state' = make_state | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 204 | managing_thread timeout_heap' oldest_heap active cancelling' messages store | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 205 | in SOME (map #2 timeout_threads, state') end | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 206 | end | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 207 | in | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 208 | while Synchronized.change_result state | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 209 | (fn st as | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 210 |             State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
 | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 211 | if (null active) andalso (null cancelling) andalso (null messages) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 212 | then (false, make_state NONE timeout_heap oldest_heap active cancelling messages store) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 213 | else (true, st)) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 214 | do | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 215 | (Synchronized.timed_access state time_limit action | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 216 | |> these | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 217 | |> List.app (unregister (false, "Interrupted (reached timeout)")); | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 218 | kill_excessive (); | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 219 | print_new_messages (); | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 220 | (*give threads time to respond to interrupt*) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 221 | OS.Process.sleep min_wait_time) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 222 | end)) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 223 | in make_state managing_thread timeout_heap oldest_heap active cancelling messages store end); | 
| 28582 | 224 | |
| 225 | ||
| 226 | (* thread is registered here by sledgehammer *) | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 227 | |
| 28582 | 228 | fun register birthtime deadtime (thread, desc) = | 
| 30798 | 229 | (Synchronized.change state | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 230 |     (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
 | 
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 231 | let | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 232 | val timeout_heap' = ThreadHeap.insert (deadtime, thread) timeout_heap | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 233 | val oldest_heap' = ThreadHeap.insert (birthtime, thread) oldest_heap | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 234 | val active' = update_thread (thread, (birthtime, deadtime, desc)) active | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 235 | in make_state managing_thread timeout_heap' oldest_heap' active' cancelling messages store end); | 
| 30798 | 236 | check_thread_manager ()); | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 237 | |
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 238 | |
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 239 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 240 | (** user commands **) | 
| 28582 | 241 | |
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 242 | (* kill: move all threads to cancelling *) | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 243 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 244 | fun kill () = Synchronized.change state | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 245 |   (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
 | 
| 28582 | 246 | let val formerly_active = map (fn (th, (tb, _, desc)) => (th, (tb, Time.now (), desc))) active | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 247 | in make_state | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 248 | managing_thread timeout_heap oldest_heap [] (formerly_active @ cancelling) messages store | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 249 | end); | 
| 28582 | 250 | |
| 251 | ||
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 252 | (* ATP info *) | 
| 28582 | 253 | |
| 254 | fun info () = | |
| 255 | let | |
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 256 |     val State {active, cancelling, ...} = Synchronized.value state
 | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 257 | |
| 28571 | 258 | fun running_info (_, (birth_time, dead_time, desc)) = "Running: " | 
| 28589 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 259 | ^ (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 | 260 | ^ " s -- " | 
| 28589 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 261 | ^ (string_of_int o Time.toSeconds) (Time.- (dead_time, Time.now ())) | 
| 28571 | 262 | ^ " s to live:\n" ^ desc | 
| 263 | 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: 
28582diff
changeset | 264 | ^ (string_of_int o Time.toSeconds) (Time.- (Time.now (), dead_time)) | 
| 28571 | 265 | ^ " s:\n" ^ desc | 
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 266 | |
| 28589 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 267 | val running = | 
| 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 268 | if null active then "No ATPs running." | 
| 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 269 |       else space_implode "\n\n" ("Running ATPs:" :: map running_info active)
 | 
| 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 270 | val interrupting = | 
| 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 271 | if null cancelling then "" | 
| 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 272 | else space_implode "\n\n" | 
| 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 273 |         ("Trying to interrupt the following ATPs:" :: map cancelling_info cancelling)
 | 
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 274 | |
| 28582 | 275 | in writeln (running ^ "\n" ^ interrupting) end; | 
| 276 | ||
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 277 | fun messages opt_limit = | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 278 | let | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 279 | val limit = the_default message_display_limit opt_limit; | 
| 29620 | 280 |     val State {store = msgs, ...} = Synchronized.value state
 | 
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 281 | val header = "Recent ATP messages" ^ | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 282 |       (if length msgs <= limit then ":" else " (" ^ string_of_int limit ^ " displayed):");
 | 
| 29116 | 283 | in writeln (space_implode "\n\n" (header :: #1 (chop limit msgs))) end; | 
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 284 | |
| 28582 | 285 | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 286 | |
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 287 | (** The Sledgehammer **) | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 288 | |
| 28582 | 289 | (* named provers *) | 
| 28484 | 290 | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 291 | type prover = int -> (thm * (string * int)) list option -> string -> int -> | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 292 | Proof.context * (thm list * thm) -> bool * string * string * string vector | 
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 293 | |
| 28582 | 294 | 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 | 295 | |
| 28582 | 296 | structure Provers = TheoryDataFun | 
| 297 | ( | |
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 298 | type T = (prover * stamp) Symtab.table | 
| 28582 | 299 | val empty = Symtab.empty | 
| 300 | val copy = I | |
| 301 | val extend = I | |
| 302 | fun merge _ tabs : T = Symtab.merge (eq_snd op =) tabs | |
| 28589 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 303 | handle Symtab.DUP dup => err_dup_prover dup | 
| 28582 | 304 | ); | 
| 28484 | 305 | |
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 306 | fun add_prover name prover thy = | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 307 | Provers.map (Symtab.update_new (name, (prover, stamp ()))) thy | 
| 28582 | 308 | handle Symtab.DUP dup => err_dup_prover dup; | 
| 309 | ||
| 310 | fun print_provers thy = Pretty.writeln | |
| 311 |   (Pretty.strs ("external provers:" :: sort_strings (Symtab.keys (Provers.get thy))));
 | |
| 28571 | 312 | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 313 | fun get_prover name thy = case Symtab.lookup (Provers.get thy) name of | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 314 | NONE => NONE | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 315 | | SOME (prover, _) => SOME prover; | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 316 | |
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 317 | (* start prover thread *) | 
| 28484 | 318 | |
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 319 | fun start_prover name birthtime deadtime i proof_state = | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 320 | (case get_prover name (Proof.theory_of proof_state) of | 
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 321 |     NONE => warning ("Unknown external prover: " ^ quote name)
 | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 322 | | SOME prover => | 
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 323 | let | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 324 | val (ctxt, (_, goal)) = Proof.get_goal proof_state | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 325 | val desc = | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 326 | "external prover " ^ quote name ^ " for subgoal " ^ string_of_int i ^ ":\n" ^ | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 327 | Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)) | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 328 | val _ = SimpleThread.fork true (fn () => | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 329 | let | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 330 | val _ = register birthtime deadtime (Thread.self (), desc) | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 331 | val result = | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 332 | let val (success, message, _, _) = | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 333 | prover (get_timeout ()) NONE name i (Proof.get_goal proof_state) | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30982diff
changeset | 334 | in (success, message) end | 
| 28835 | 335 | handle ResHolClause.TOO_TRIVIAL | 
| 336 | => (true, "Empty clause: Try this command: " ^ Markup.markup Markup.sendback "apply metis") | |
| 337 | | ERROR msg | |
| 338 | => (false, "Error: " ^ msg) | |
| 29620 | 339 | val _ = unregister result (Thread.self ()) | 
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 340 | in () end handle Interrupt => ()) | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 341 | in () end); | 
| 28582 | 342 | |
| 343 | ||
| 344 | (* sledghammer for first subgoal *) | |
| 345 | ||
| 346 | fun sledgehammer names proof_state = | |
| 347 | let | |
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 348 | val provers = | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 349 | if null names then String.tokens (Symbol.is_ascii_blank o String.str) (get_atps ()) | 
| 28582 | 350 | else names | 
| 351 | val birthtime = Time.now () | |
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 352 | val deadtime = Time.+ (birthtime, Time.fromSeconds (get_timeout ())) | 
| 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 353 | in List.app (fn name => start_prover name birthtime deadtime 1 proof_state) provers end; | 
| 28582 | 354 | |
| 355 | ||
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 356 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 357 | (** Isar command syntax **) | 
| 28582 | 358 | |
| 359 | local structure K = OuterKeyword and P = OuterParse in | |
| 360 | ||
| 361 | val _ = | |
| 362 | 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: 
28582diff
changeset | 363 | (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill)); | 
| 28582 | 364 | |
| 365 | val _ = | |
| 366 | OuterSyntax.improper_command "atp_info" "print information about managed provers" K.diag | |
| 367 | (Scan.succeed (Toplevel.no_timing o Toplevel.imperative info)); | |
| 368 | ||
| 369 | val _ = | |
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 370 | OuterSyntax.improper_command "atp_messages" "print recent messages issued by managed provers" K.diag | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 371 |     (Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")") >>
 | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 372 | (fn limit => Toplevel.no_timing o Toplevel.imperative (fn () => messages limit))); | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 373 | |
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 374 | val _ = | 
| 28582 | 375 | OuterSyntax.improper_command "print_atps" "print external provers" K.diag | 
| 376 | (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o | |
| 377 | Toplevel.keep (print_provers o Toplevel.theory_of))); | |
| 378 | ||
| 379 | val _ = | |
| 380 | OuterSyntax.command "sledgehammer" "call all automatic theorem provers" K.diag | |
| 381 | (Scan.repeat P.xname >> (fn names => Toplevel.no_timing o Toplevel.unknown_proof o | |
| 30800 | 382 | Toplevel.keep (sledgehammer names o Toplevel.proof_of))); | 
| 28582 | 383 | |
| 384 | end; | |
| 385 | ||
| 386 | end; | |
| 30537 | 387 |