| author | haftmann | 
| Thu, 17 Jun 2010 15:59:48 +0200 | |
| changeset 37450 | 45073611170a | 
| parent 37413 | e856582fe9c4 | 
| child 37480 | d5a85d35ef62 | 
| 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 | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 3 | Author: Makarius | 
| 35969 | 4 | Author: Jasmin Blanchette, TU Muenchen | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 5 | |
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 6 | Central manager component for ATP threads. | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 7 | *) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 8 | |
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 9 | signature ATP_MANAGER = | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 10 | sig | 
| 36393 
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
 blanchet parents: 
36392diff
changeset | 11 | type name_pool = Sledgehammer_HOL_Clause.name_pool | 
| 35969 | 12 | type relevance_override = Sledgehammer_Fact_Filter.relevance_override | 
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36235diff
changeset | 13 | type minimize_command = Sledgehammer_Proof_Reconstruct.minimize_command | 
| 35969 | 14 | type params = | 
| 15 |     {debug: bool,
 | |
| 16 | verbose: bool, | |
| 36143 
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
 blanchet parents: 
36064diff
changeset | 17 | overlord: bool, | 
| 35969 | 18 | atps: string list, | 
| 19 | full_types: bool, | |
| 36235 
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
 blanchet parents: 
36231diff
changeset | 20 | explicit_apply: bool, | 
| 36058 
8256d5a185bd
added "respect_no_atp" and "convergence" options to Sledgehammer;
 blanchet parents: 
35969diff
changeset | 21 | respect_no_atp: bool, | 
| 35969 | 22 | relevance_threshold: real, | 
| 36922 | 23 | relevance_convergence: real, | 
| 36220 
f3655a3ae1ab
rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
 blanchet parents: 
36184diff
changeset | 24 | theory_relevant: bool option, | 
| 36922 | 25 | defs_relevant: bool, | 
| 35969 | 26 | isar_proof: bool, | 
| 36924 | 27 | isar_shrink_factor: int, | 
| 35969 | 28 | timeout: Time.time, | 
| 29 | minimize_timeout: Time.time} | |
| 35867 | 30 | type problem = | 
| 35969 | 31 |     {subgoal: int,
 | 
| 32 | goal: Proof.context * (thm list * thm), | |
| 33 | relevance_override: relevance_override, | |
| 34 | axiom_clauses: (thm * (string * int)) list option, | |
| 35 | filtered_clauses: (thm * (string * int)) list option} | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 36 | datatype failure = | 
| 37413 | 37 | Unprovable | IncompleteUnprovable | TimedOut | OutOfResources | OldSpass | | 
| 38 | MalformedInput | MalformedOutput | UnknownError | |
| 35867 | 39 | type prover_result = | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 40 |     {outcome: failure option,
 | 
| 35969 | 41 | message: string, | 
| 36393 
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
 blanchet parents: 
36392diff
changeset | 42 | pool: name_pool option, | 
| 35969 | 43 | relevant_thm_names: string list, | 
| 44 | atp_run_time_in_msecs: int, | |
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 45 | output: string, | 
| 35969 | 46 | proof: string, | 
| 47 | internal_thm_names: string Vector.vector, | |
| 36400 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 48 | conjecture_shape: int list list, | 
| 35969 | 49 | filtered_clauses: (thm * (string * int)) list} | 
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36235diff
changeset | 50 | type prover = | 
| 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36235diff
changeset | 51 | params -> minimize_command -> Time.time -> problem -> prover_result | 
| 35867 | 52 | |
| 35969 | 53 | val kill_atps: unit -> unit | 
| 54 | val running_atps: unit -> unit | |
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 55 | val messages: int option -> unit | 
| 35867 | 56 | val add_prover: string * prover -> theory -> theory | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 57 | val get_prover: theory -> string -> prover | 
| 35969 | 58 | val available_atps: theory -> unit | 
| 36373 
66af0a49de39
move some sledgehammer stuff out of "atp_manager.ML"
 blanchet parents: 
36372diff
changeset | 59 | val start_prover_thread: | 
| 36482 
1281be23bd23
added total goal count as argument + message when killing ATPs
 blanchet parents: 
36473diff
changeset | 60 | params -> Time.time -> Time.time -> int -> int -> relevance_override | 
| 36373 
66af0a49de39
move some sledgehammer stuff out of "atp_manager.ML"
 blanchet parents: 
36372diff
changeset | 61 | -> (string -> minimize_command) -> Proof.state -> string -> unit | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 62 | end; | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 63 | |
| 35865 | 64 | structure ATP_Manager : ATP_MANAGER = | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 65 | struct | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 66 | |
| 36229 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 67 | open Sledgehammer_Util | 
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36059diff
changeset | 68 | open Sledgehammer_Fact_Filter | 
| 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
36059diff
changeset | 69 | open Sledgehammer_Proof_Reconstruct | 
| 35969 | 70 | |
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36235diff
changeset | 71 | (** problems, results, provers, etc. **) | 
| 35969 | 72 | |
| 73 | type params = | |
| 74 |   {debug: bool,
 | |
| 75 | verbose: bool, | |
| 36143 
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
 blanchet parents: 
36064diff
changeset | 76 | overlord: bool, | 
| 35969 | 77 | atps: string list, | 
| 78 | full_types: bool, | |
| 36235 
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
 blanchet parents: 
36231diff
changeset | 79 | explicit_apply: bool, | 
| 36058 
8256d5a185bd
added "respect_no_atp" and "convergence" options to Sledgehammer;
 blanchet parents: 
35969diff
changeset | 80 | respect_no_atp: bool, | 
| 35969 | 81 | relevance_threshold: real, | 
| 36922 | 82 | relevance_convergence: real, | 
| 36220 
f3655a3ae1ab
rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
 blanchet parents: 
36184diff
changeset | 83 | theory_relevant: bool option, | 
| 36922 | 84 | defs_relevant: bool, | 
| 35969 | 85 | isar_proof: bool, | 
| 36924 | 86 | isar_shrink_factor: int, | 
| 35969 | 87 | timeout: Time.time, | 
| 88 | minimize_timeout: Time.time} | |
| 35867 | 89 | |
| 90 | type problem = | |
| 35969 | 91 |   {subgoal: int,
 | 
| 92 | goal: Proof.context * (thm list * thm), | |
| 93 | relevance_override: relevance_override, | |
| 94 | axiom_clauses: (thm * (string * int)) list option, | |
| 95 | filtered_clauses: (thm * (string * int)) list option}; | |
| 35867 | 96 | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 97 | datatype failure = | 
| 37413 | 98 | Unprovable | IncompleteUnprovable | TimedOut | OutOfResources | OldSpass | | 
| 99 | MalformedInput | MalformedOutput | UnknownError | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 100 | |
| 35867 | 101 | type prover_result = | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 102 |   {outcome: failure option,
 | 
| 35969 | 103 | message: string, | 
| 36393 
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
 blanchet parents: 
36392diff
changeset | 104 | pool: name_pool option, | 
| 35969 | 105 | relevant_thm_names: string list, | 
| 106 | atp_run_time_in_msecs: int, | |
| 36369 
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
 blanchet parents: 
36289diff
changeset | 107 | output: string, | 
| 35969 | 108 | proof: string, | 
| 109 | internal_thm_names: string Vector.vector, | |
| 36400 
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
 blanchet parents: 
36393diff
changeset | 110 | conjecture_shape: int list list, | 
| 35969 | 111 | filtered_clauses: (thm * (string * int)) list}; | 
| 35867 | 112 | |
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36235diff
changeset | 113 | type prover = | 
| 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36235diff
changeset | 114 | params -> minimize_command -> Time.time -> problem -> prover_result | 
| 35867 | 115 | |
| 116 | ||
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 117 | (** preferences **) | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 118 | |
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 119 | val message_store_limit = 20; | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 120 | val message_display_limit = 5; | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 121 | |
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 122 | |
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 123 | (** thread management **) | 
| 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 124 | |
| 28582 | 125 | (* data structures over threads *) | 
| 126 | ||
| 32939 | 127 | structure Thread_Heap = Heap | 
| 28582 | 128 | ( | 
| 129 | type elem = Time.time * Thread.thread; | |
| 130 | fun ord ((a, _), (b, _)) = Time.compare (a, b); | |
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 131 | ); | 
| 28582 | 132 | |
| 31368 | 133 | fun lookup_thread xs = AList.lookup Thread.equal xs; | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 134 | fun delete_thread xs = AList.delete Thread.equal xs; | 
| 31368 | 135 | fun update_thread xs = AList.update Thread.equal xs; | 
| 28582 | 136 | |
| 137 | ||
| 138 | (* state of thread manager *) | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 139 | |
| 32938 | 140 | type state = | 
| 141 |  {manager: Thread.thread option,
 | |
| 142 | timeout_heap: Thread_Heap.T, | |
| 28582 | 143 | active: (Thread.thread * (Time.time * Time.time * string)) list, | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 144 | cancelling: (Thread.thread * (Time.time * string)) list, | 
| 29620 | 145 | messages: string list, | 
| 146 | store: string list}; | |
| 28582 | 147 | |
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 148 | fun make_state manager timeout_heap active cancelling messages store : state = | 
| 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 149 |   {manager = manager, timeout_heap = timeout_heap, active = active,
 | 
| 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 150 | cancelling = cancelling, messages = messages, store = store}; | 
| 28582 | 151 | |
| 32938 | 152 | val global_state = Synchronized.var "atp_manager" | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 153 | (make_state NONE Thread_Heap.empty [] [] [] []); | 
| 28582 | 154 | |
| 31368 | 155 | |
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 156 | (* unregister ATP thread *) | 
| 28582 | 157 | |
| 36184 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 158 | fun unregister ({verbose, ...} : params) message thread =
 | 
| 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 159 | Synchronized.change global_state | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 160 |   (fn state as {manager, timeout_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 | 161 | (case lookup_thread active thread of | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 162 | SOME (birth_time, _, desc) => | 
| 29150 
8af5ee47f30c
unregister: refrain from cancelling self, which sometimes caused sledghammer to fail just before printing the result;
 wenzelm parents: 
29116diff
changeset | 163 | let | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 164 | val active' = delete_thread thread active; | 
| 36184 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 165 | val now = Time.now () | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 166 | val cancelling' = (thread, (now, desc)) :: cancelling; | 
| 36184 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 167 | val message' = | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 168 | desc ^ "\n" ^ message ^ | 
| 36184 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 169 | (if verbose then | 
| 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 170 | "Total time: " ^ Int.toString (Time.toMilliseconds | 
| 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 171 | (Time.- (now, birth_time))) ^ " ms.\n" | 
| 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 172 | else | 
| 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 173 | "") | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 174 | val messages' = message' :: messages; | 
| 29620 | 175 | val store' = message' :: | 
| 176 | (if length store <= message_store_limit then store | |
| 32938 | 177 | else #1 (chop message_store_limit store)); | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 178 | in make_state manager timeout_heap active' cancelling' messages' store' end | 
| 30800 | 179 | | NONE => state)); | 
| 28582 | 180 | |
| 181 | ||
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 182 | (* main manager thread -- only one may exist *) | 
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 183 | |
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 184 | val min_wait_time = Time.fromMilliseconds 300; | 
| 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 185 | val max_wait_time = Time.fromSeconds 10; | 
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 186 | |
| 36229 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 187 | (* This is a workaround for Proof General's off-by-a-few sendback display bug, | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 188 | whereby "pr" in "proof" is not highlighted. *) | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 189 | val break_into_chunks = | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 190 | map (replace_all "\n\n" "\000") #> maps (space_explode "\000") | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 191 | |
| 29620 | 192 | fun print_new_messages () = | 
| 36229 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 193 | case Synchronized.change_result global_state | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 194 |          (fn {manager, timeout_heap, active, cancelling, messages, store} =>
 | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 195 | (messages, make_state manager timeout_heap active cancelling [] | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 196 | store)) of | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 197 | [] => () | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 198 | | msgs => | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 199 | msgs |> break_into_chunks | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 200 | |> (fn msg :: msgs => "Sledgehammer: " ^ msg :: msgs) | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 201 | |> List.app priority | 
| 29620 | 202 | |
| 36184 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 203 | fun check_thread_manager params = Synchronized.change global_state | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 204 |   (fn state as {manager, timeout_heap, active, cancelling, messages, store} =>
 | 
| 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 205 | if (case manager of SOME thread => Thread.isActive thread | NONE => false) then state | 
| 33604 
d4220df6fde2
Toplevel.thread provides Isar-style exception output;
 wenzelm parents: 
33522diff
changeset | 206 | else let val manager = SOME (Toplevel.thread false (fn () => | 
| 32938 | 207 | let | 
| 208 | fun time_limit timeout_heap = | |
| 209 | (case try Thread_Heap.min timeout_heap of | |
| 210 | NONE => Time.+ (Time.now (), max_wait_time) | |
| 211 | | SOME (time, _) => time); | |
| 28582 | 212 | |
| 32938 | 213 | (*action: find threads whose timeout is reached, and interrupt cancelling threads*) | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 214 |         fun action {manager, timeout_heap, active, cancelling, messages, store} =
 | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 215 | let val (timeout_threads, timeout_heap') = | 
| 32938 | 216 | Thread_Heap.upto (Time.now (), Thread.self ()) timeout_heap; | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 217 | in | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 218 | if null timeout_threads andalso null cancelling | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 219 | then NONE | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 220 | else | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 221 | let | 
| 37216 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 wenzelm parents: 
36965diff
changeset | 222 | val _ = List.app (Simple_Thread.interrupt o #1) cancelling; | 
| 32938 | 223 | val cancelling' = filter (Thread.isActive o #1) cancelling; | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 224 | val state' = make_state manager timeout_heap' active cancelling' messages store; | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 225 | in SOME (map #2 timeout_threads, state') end | 
| 32938 | 226 | end; | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 227 | in | 
| 32938 | 228 | while Synchronized.change_result global_state | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 229 |           (fn state as {timeout_heap, active, cancelling, messages, store, ...} =>
 | 
| 32938 | 230 | if null active andalso null cancelling andalso null messages | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 231 | then (false, make_state NONE timeout_heap active cancelling messages store) | 
| 32938 | 232 | else (true, state)) | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 233 | do | 
| 32938 | 234 | (Synchronized.timed_access global_state (SOME o time_limit o #timeout_heap) action | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 235 | |> these | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 236 | |> List.app (unregister params "Timed out.\n"); | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 237 | print_new_messages (); | 
| 32938 | 238 | (*give threads some time to respond to interrupt*) | 
| 30830 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 239 | OS.Process.sleep min_wait_time) | 
| 
263064c4d0c3
included managing_thread in state of AtpManager:
 immler@in.tum.de parents: 
30800diff
changeset | 240 | end)) | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 241 | in make_state manager timeout_heap active cancelling messages store end); | 
| 28582 | 242 | |
| 243 | ||
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 244 | (* register ATP thread *) | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 245 | |
| 36184 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 246 | fun register params birth_time death_time (thread, desc) = | 
| 32938 | 247 | (Synchronized.change global_state | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 248 |     (fn {manager, timeout_heap, active, cancelling, messages, store} =>
 | 
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 249 | let | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 250 | val timeout_heap' = Thread_Heap.insert (death_time, thread) timeout_heap; | 
| 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 251 | val active' = update_thread (thread, (birth_time, death_time, desc)) active; | 
| 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 252 | val state' = make_state manager timeout_heap' active' cancelling messages store; | 
| 32938 | 253 | in state' end); | 
| 36184 
54a9c0679079
output total time taken by Sledgehammer if "verbose" is set
 blanchet parents: 
36181diff
changeset | 254 | check_thread_manager params); | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 255 | |
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 256 | |
| 28586 
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 | (** user commands **) | 
| 28582 | 259 | |
| 35969 | 260 | (* kill ATPs *) | 
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 261 | |
| 35969 | 262 | fun kill_atps () = Synchronized.change global_state | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 263 |   (fn {manager, timeout_heap, active, cancelling, messages, store} =>
 | 
| 32938 | 264 | let | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 265 | val killing = map (fn (th, (_, _, desc)) => (th, (Time.now (), desc))) active; | 
| 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 266 | val state' = make_state manager timeout_heap [] (killing @ cancelling) messages store; | 
| 36482 
1281be23bd23
added total goal count as argument + message when killing ATPs
 blanchet parents: 
36473diff
changeset | 267 | val _ = if null active then () | 
| 
1281be23bd23
added total goal count as argument + message when killing ATPs
 blanchet parents: 
36473diff
changeset | 268 | else priority "Killed active Sledgehammer threads." | 
| 32938 | 269 | in state' end); | 
| 28582 | 270 | |
| 271 | ||
| 35969 | 272 | (* running_atps *) | 
| 28582 | 273 | |
| 32938 | 274 | fun seconds time = string_of_int (Time.toSeconds time) ^ "s"; | 
| 275 | ||
| 35969 | 276 | fun running_atps () = | 
| 28582 | 277 | let | 
| 32938 | 278 |     val {active, cancelling, ...} = Synchronized.value global_state;
 | 
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 279 | |
| 32938 | 280 | val now = Time.now (); | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 281 | fun running_info (_, (birth_time, death_time, desc)) = | 
| 32938 | 282 | "Running: " ^ seconds (Time.- (now, birth_time)) ^ " -- " ^ | 
| 32996 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 283 | seconds (Time.- (death_time, now)) ^ " to live:\n" ^ desc; | 
| 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 284 | fun cancelling_info (_, (deadth_time, desc)) = | 
| 
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
 wenzelm parents: 
32995diff
changeset | 285 | "Trying to interrupt thread since " ^ seconds (Time.- (now, deadth_time)) ^ ":\n" ^ desc; | 
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 286 | |
| 28589 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 287 | val running = | 
| 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 288 | if null active then "No ATPs running." | 
| 32938 | 289 |       else space_implode "\n\n" ("Running ATPs:" :: map running_info active);
 | 
| 28589 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 290 | val interrupting = | 
| 
581b2ab9827a
adding preferences is now permissive, no error handling here;
 wenzelm parents: 
28586diff
changeset | 291 | if null cancelling then "" | 
| 32938 | 292 | else | 
| 293 | space_implode "\n\n" | |
| 294 |           ("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 | 295 | |
| 36229 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 296 | in priority (running ^ "\n" ^ interrupting) end; | 
| 28582 | 297 | |
| 29112 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 298 | fun messages opt_limit = | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 299 | let | 
| 
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
 wenzelm parents: 
28835diff
changeset | 300 | val limit = the_default message_display_limit opt_limit; | 
| 32938 | 301 |     val {store, ...} = Synchronized.value global_state;
 | 
| 302 | val header = | |
| 303 | "Recent ATP messages" ^ | |
| 304 |         (if length store <= limit then ":" else " (" ^ string_of_int limit ^ " displayed):");
 | |
| 36229 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 305 | in List.app priority (header :: break_into_chunks (#1 (chop limit store))) end | 
| 28582 | 306 | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 307 | |
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 308 | (** The Sledgehammer **) | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: diff
changeset | 309 | |
| 28582 | 310 | (* named provers *) | 
| 28484 | 311 | |
| 36229 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 312 | 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 | 313 | |
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 314 | structure Data = Theory_Data | 
| 28582 | 315 | ( | 
| 35867 | 316 | type T = (prover * stamp) Symtab.table; | 
| 32938 | 317 | val empty = Symtab.empty; | 
| 318 | val extend = I; | |
| 33522 | 319 | fun merge data : T = Symtab.merge (eq_snd op =) data | 
| 36281 
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
 blanchet parents: 
36235diff
changeset | 320 | handle Symtab.DUP name => err_dup_prover name; | 
| 28582 | 321 | ); | 
| 28484 | 322 | |
| 32864 
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
 boehmes parents: 
32824diff
changeset | 323 | fun add_prover (name, prover) thy = | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 324 | Data.map (Symtab.update_new (name, (prover, stamp ()))) thy | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 325 | handle Symtab.DUP name => err_dup_prover name; | 
| 28582 | 326 | |
| 32995 | 327 | fun get_prover thy name = | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 328 | case Symtab.lookup (Data.get thy) name of | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 329 | SOME (prover, _) => prover | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 330 |   | NONE => error ("Unknown ATP: " ^ name)
 | 
| 32995 | 331 | |
| 36229 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 332 | fun available_atps thy = | 
| 
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
 blanchet parents: 
36226diff
changeset | 333 |   priority ("Available ATPs: " ^
 | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 334 | commas (sort_strings (Symtab.keys (Data.get thy))) ^ ".") | 
| 28571 | 335 | |
| 28586 
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
 wenzelm parents: 
28582diff
changeset | 336 | |
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 337 | (* start prover thread *) | 
| 28484 | 338 | |
| 36482 
1281be23bd23
added total goal count as argument + message when killing ATPs
 blanchet parents: 
36473diff
changeset | 339 | fun start_prover_thread (params as {timeout, ...}) birth_time death_time i n
 | 
| 36373 
66af0a49de39
move some sledgehammer stuff out of "atp_manager.ML"
 blanchet parents: 
36372diff
changeset | 340 | relevance_override minimize_command proof_state name = | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 341 | let | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 342 | val prover = get_prover (Proof.theory_of proof_state) name | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 343 |     val {context = ctxt, facts, goal} = Proof.goal proof_state;
 | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 344 | val desc = | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 345 | "ATP " ^ quote name ^ " for subgoal " ^ string_of_int i ^ ":\n" ^ | 
| 36392 | 346 | Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)); | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 347 | val _ = Toplevel.thread true (fn () => | 
| 28595 
67e3945b53f1
add_prover: plain prover function, without thread;
 wenzelm parents: 
28589diff
changeset | 348 | let | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 349 | val _ = register params birth_time death_time (Thread.self (), desc) | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 350 | val problem = | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 351 |           {subgoal = i, goal = (ctxt, (facts, goal)),
 | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 352 | relevance_override = relevance_override, axiom_clauses = NONE, | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 353 | filtered_clauses = NONE} | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 354 | val message = | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 355 | #message (prover params (minimize_command name) timeout problem) | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 356 | handle Sledgehammer_HOL_Clause.TRIVIAL => metis_line i n [] | 
| 36383 | 357 | | ERROR message => "Error: " ^ message ^ "\n" | 
| 36379 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 358 | val _ = unregister params message (Thread.self ()); | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 359 | in () end) | 
| 
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
 blanchet parents: 
36373diff
changeset | 360 | in () end | 
| 28582 | 361 | |
| 362 | end; |