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