| author | huffman | 
| Tue, 23 Aug 2011 14:11:02 -0700 | |
| changeset 44457 | d366fa5551ef | 
| parent 43630 | e42ccb132305 | 
| child 44463 | c471a2b48fa1 | 
| permissions | -rw-r--r-- | 
| 38988 | 1 | (* Title: HOL/Tools/Sledgehammer/sledgehammer_minimize.ML | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 2 | Author: Philipp Meyer, TU Muenchen | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 4 | |
| 40977 | 5 | Minimization of fact list for Metis using external provers. | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 6 | *) | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 7 | |
| 38988 | 8 | signature SLEDGEHAMMER_MINIMIZE = | 
| 32525 | 9 | sig | 
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43064diff
changeset | 10 | type locality = ATP_Translate.locality | 
| 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43064diff
changeset | 11 | type play = ATP_Reconstruct.play | 
| 41087 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 blanchet parents: 
40983diff
changeset | 12 | type params = Sledgehammer_Provers.params | 
| 35867 | 13 | |
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42579diff
changeset | 14 | val binary_min_facts : int Config.T | 
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 15 | val minimize_facts : | 
| 43064 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 blanchet parents: 
43058diff
changeset | 16 | string -> params -> bool -> int -> int -> Proof.state | 
| 41091 
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
 blanchet parents: 
41090diff
changeset | 17 | -> ((string * locality) * thm list) list | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 18 | -> ((string * locality) * thm list) list option | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 19 | * ((unit -> play) * (play -> string) * string) | 
| 38996 | 20 | val run_minimize : | 
| 21 | params -> int -> (Facts.ref * Attrib.src list) list -> Proof.state -> unit | |
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 22 | end; | 
| 32525 | 23 | |
| 38988 | 24 | structure Sledgehammer_Minimize : SLEDGEHAMMER_MINIMIZE = | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 25 | struct | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 26 | |
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43064diff
changeset | 27 | open ATP_Util | 
| 39496 
a52a4e4399c1
got caught once again by SML's pattern maching (ctor vs. var)
 blanchet parents: 
39491diff
changeset | 28 | open ATP_Proof | 
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43064diff
changeset | 29 | open ATP_Translate | 
| 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43064diff
changeset | 30 | open ATP_Reconstruct | 
| 36142 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 31 | open Sledgehammer_Util | 
| 38988 | 32 | open Sledgehammer_Filter | 
| 41087 
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
 blanchet parents: 
40983diff
changeset | 33 | open Sledgehammer_Provers | 
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 34 | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 35 | (* wrapper for calling external prover *) | 
| 31236 | 36 | |
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 37 | fun n_facts names = | 
| 38698 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 blanchet parents: 
38696diff
changeset | 38 | let val n = length names in | 
| 40061 
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
 blanchet parents: 
40060diff
changeset | 39 | string_of_int n ^ " fact" ^ plural_s n ^ | 
| 38092 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38084diff
changeset | 40 | (if n > 0 then | 
| 38698 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 blanchet parents: 
38696diff
changeset | 41 | ": " ^ (names |> map fst | 
| 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 blanchet parents: 
38696diff
changeset | 42 | |> sort_distinct string_ord |> space_implode " ") | 
| 38092 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38084diff
changeset | 43 | else | 
| 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38084diff
changeset | 44 | "") | 
| 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38084diff
changeset | 45 | end | 
| 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38084diff
changeset | 46 | |
| 41091 
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
 blanchet parents: 
41090diff
changeset | 47 | fun print silent f = if silent then () else Output.urgent_message (f ()) | 
| 
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
 blanchet parents: 
41090diff
changeset | 48 | |
| 42724 
4d6bcf846759
added "max_mono_instances" option to Sledgehammer and renamed old "monomorphize_limit" option
 blanchet parents: 
42646diff
changeset | 49 | fun test_facts ({debug, verbose, overlord, provers, max_mono_iters,
 | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43577diff
changeset | 50 | max_new_mono_instances, type_enc, isar_proof, | 
| 43015 
21b6baec55b1
renamed "metis_timeout" to "preplay_timeout" and continued implementation
 blanchet parents: 
43011diff
changeset | 51 | isar_shrink_factor, preplay_timeout, ...} : params) | 
| 43064 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 blanchet parents: 
43058diff
changeset | 52 | silent (prover : prover) timeout i n state facts = | 
| 31236 | 53 | let | 
| 41277 
1369c27c6966
reduce the minimizer slack and add verbose information
 blanchet parents: 
41267diff
changeset | 54 | val _ = | 
| 
1369c27c6966
reduce the minimizer slack and add verbose information
 blanchet parents: 
41267diff
changeset | 55 | print silent (fn () => | 
| 
1369c27c6966
reduce the minimizer slack and add verbose information
 blanchet parents: 
41267diff
changeset | 56 | "Testing " ^ n_facts (map fst facts) ^ | 
| 
1369c27c6966
reduce the minimizer slack and add verbose information
 blanchet parents: 
41267diff
changeset | 57 | (if verbose then " (timeout: " ^ string_from_time timeout ^ ")" | 
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43064diff
changeset | 58 | else "") ^ "...") | 
| 41742 
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
 blanchet parents: 
41741diff
changeset | 59 |     val {goal, ...} = Proof.goal state
 | 
| 38100 
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
 blanchet parents: 
38094diff
changeset | 60 | val params = | 
| 42060 
889d767ce5f4
make Minimizer honor "verbose" and "debug" options better
 blanchet parents: 
41824diff
changeset | 61 |       {debug = debug, verbose = verbose, overlord = overlord, blocking = true,
 | 
| 43626 
a867ebb12209
renamed "type_sys" to "type_enc", which is more accurate
 blanchet parents: 
43577diff
changeset | 62 | provers = provers, type_enc = type_enc, sound = true, | 
| 41138 
eb80538166b6
implemented partially-typed "tags" type encoding
 blanchet parents: 
41134diff
changeset | 63 | relevance_thresholds = (1.01, 1.01), max_relevant = NONE, | 
| 42740 
31334a7b109d
renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
 blanchet parents: 
42724diff
changeset | 64 | max_mono_iters = max_mono_iters, | 
| 
31334a7b109d
renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
 blanchet parents: 
42724diff
changeset | 65 | max_new_mono_instances = max_new_mono_instances, isar_proof = isar_proof, | 
| 
31334a7b109d
renamed "max_mono_instances" to "max_new_mono_instances" and changed its semantics accordingly
 blanchet parents: 
42724diff
changeset | 66 | isar_shrink_factor = isar_shrink_factor, slicing = false, | 
| 43015 
21b6baec55b1
renamed "metis_timeout" to "preplay_timeout" and continued implementation
 blanchet parents: 
43011diff
changeset | 67 | timeout = timeout, preplay_timeout = preplay_timeout, expect = ""} | 
| 40204 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40200diff
changeset | 68 | val facts = | 
| 41090 | 69 | facts |> maps (fn (n, ths) => ths |> map (Untranslated_Fact o pair n)) | 
| 40065 | 70 | val problem = | 
| 71 |       {state = state, goal = goal, subgoal = i, subgoal_count = n,
 | |
| 41741 | 72 | facts = facts, smt_filter = NONE} | 
| 43050 
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
 blanchet parents: 
43043diff
changeset | 73 |     val result as {outcome, used_facts, run_time_in_msecs, ...} =
 | 
| 43051 | 74 | prover params (K (K "")) problem | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 75 | in | 
| 43259 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 76 | print silent | 
| 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 77 | (fn () => | 
| 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 78 | case outcome of | 
| 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 79 | SOME failure => string_for_failure failure | 
| 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 80 | | NONE => | 
| 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 81 | "Found proof" ^ | 
| 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 82 | (if length used_facts = length facts then "" | 
| 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 83 | else " with " ^ n_facts used_facts) ^ | 
| 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 84 | (case run_time_in_msecs of | 
| 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 85 | SOME ms => | 
| 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 86 |                     " (" ^ string_from_time (Time.fromMilliseconds ms) ^ ")"
 | 
| 
30c141dc22d6
killed "explicit_apply" option in Sledgehammer -- the "smart" default is about as lightweight as "false" and just as complete as "true"
 blanchet parents: 
43166diff
changeset | 87 | | NONE => "") ^ "."); | 
| 38092 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38084diff
changeset | 88 | result | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 89 | end | 
| 31236 | 90 | |
| 40204 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40200diff
changeset | 91 | (* minimalization of facts *) | 
| 31236 | 92 | |
| 40977 | 93 | (* The sublinear algorithm works well in almost all situations, except when the | 
| 94 | external prover cannot return the list of used facts and hence returns all | |
| 41267 
958fee9ec275
lower threshold where the binary algorithm kick in and use the same value for automatic minimization
 blanchet parents: 
41265diff
changeset | 95 | facts as used. In that case, the binary algorithm is much more appropriate. | 
| 
958fee9ec275
lower threshold where the binary algorithm kick in and use the same value for automatic minimization
 blanchet parents: 
41265diff
changeset | 96 | We can usually detect the situation by looking at the number of used facts | 
| 
958fee9ec275
lower threshold where the binary algorithm kick in and use the same value for automatic minimization
 blanchet parents: 
41265diff
changeset | 97 | reported by the prover. *) | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42579diff
changeset | 98 | val binary_min_facts = | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42579diff
changeset | 99 |   Attrib.setup_config_int @{binding sledgehammer_minimize_binary_min_facts}
 | 
| 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42579diff
changeset | 100 | (K 20) | 
| 40977 | 101 | |
| 38015 | 102 | fun sublinear_minimize _ [] p = p | 
| 103 | | sublinear_minimize test (x :: xs) (seen, result) = | |
| 104 | case test (xs @ seen) of | |
| 40204 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40200diff
changeset | 105 |       result as {outcome = NONE, used_facts, ...} : prover_result =>
 | 
| 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40200diff
changeset | 106 | sublinear_minimize test (filter_used_facts used_facts xs) | 
| 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40200diff
changeset | 107 | (filter_used_facts used_facts seen, result) | 
| 38015 | 108 | | _ => sublinear_minimize test xs (x :: seen, result) | 
| 109 | ||
| 40977 | 110 | fun binary_minimize test xs = | 
| 111 | let | |
| 43306 | 112 | fun pred xs = #outcome (test xs : prover_result) = NONE | 
| 40977 | 113 | fun split [] p = p | 
| 114 | | split [h] (l, r) = (h :: l, r) | |
| 115 | | split (h1 :: h2 :: t) (l, r) = split t (h1 :: l, h2 :: r) | |
| 41743 | 116 | fun min _ _ [] = raise Empty | 
| 117 | | min _ _ [s0] = [s0] | |
| 118 | | min depth sup xs = | |
| 119 | let | |
| 120 | (* | |
| 121 |           val _ = warning (replicate_string depth " " ^ "{" ^ ("  " ^
 | |
| 122 | n_facts (map fst sup) ^ " and " ^ | |
| 123 | n_facts (map fst xs))) | |
| 124 | *) | |
| 125 | val (l0, r0) = split xs ([], []) | |
| 126 | in | |
| 43306 | 127 | if pred (sup @ l0) then | 
| 41743 | 128 | min (depth + 1) sup l0 | 
| 43306 | 129 | else if pred (sup @ r0) then | 
| 41743 | 130 | min (depth + 1) sup r0 | 
| 40977 | 131 | else | 
| 132 | let | |
| 41743 | 133 | val l = min (depth + 1) (sup @ r0) l0 | 
| 134 | val r = min (depth + 1) (sup @ l) r0 | |
| 40977 | 135 | in l @ r end | 
| 136 | end | |
| 41743 | 137 | (* | 
| 138 | |> tap (fn _ => warning (replicate_string depth " " ^ "}")) | |
| 139 | *) | |
| 40977 | 140 | val xs = | 
| 41743 | 141 | case min 0 [] xs of | 
| 43306 | 142 | [x] => if pred [] then [] else [x] | 
| 40977 | 143 | | xs => xs | 
| 144 | in (xs, test xs) end | |
| 145 | ||
| 146 | (* Give the external prover some slack. The ATP gets further slack because the | |
| 147 | Sledgehammer preprocessing time is included in the estimate below but isn't | |
| 148 | part of the timeout. *) | |
| 41277 
1369c27c6966
reduce the minimizer slack and add verbose information
 blanchet parents: 
41267diff
changeset | 149 | val slack_msecs = 200 | 
| 38092 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38084diff
changeset | 150 | |
| 43064 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 blanchet parents: 
43058diff
changeset | 151 | fun minimize_facts prover_name (params as {timeout, ...}) silent i n state
 | 
| 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 blanchet parents: 
43058diff
changeset | 152 | facts = | 
| 31236 | 153 | let | 
| 40941 
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
 blanchet parents: 
40553diff
changeset | 154 | val ctxt = Proof.context_of state | 
| 43021 | 155 | val prover = get_prover ctxt Minimize prover_name | 
| 38590 
bd443b426d56
get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
 blanchet parents: 
38589diff
changeset | 156 | val msecs = Time.toMilliseconds timeout | 
| 43058 
5f8bac7a2945
minimize automatically even if Metis failed, if the external prover was really fast
 blanchet parents: 
43052diff
changeset | 157 | val _ = print silent (fn () => "Sledgehammer minimizer: " ^ | 
| 40977 | 158 | quote prover_name ^ ".") | 
| 43064 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 blanchet parents: 
43058diff
changeset | 159 | fun do_test timeout = test_facts params silent prover timeout i n state | 
| 38092 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38084diff
changeset | 160 | val timer = Timer.startRealTimer () | 
| 31236 | 161 | in | 
| 40204 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40200diff
changeset | 162 | (case do_test timeout facts of | 
| 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40200diff
changeset | 163 |        result as {outcome = NONE, used_facts, ...} =>
 | 
| 38015 | 164 | let | 
| 38092 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38084diff
changeset | 165 | val time = Timer.checkRealTimer timer | 
| 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38084diff
changeset | 166 | val new_timeout = | 
| 41277 
1369c27c6966
reduce the minimizer slack and add verbose information
 blanchet parents: 
41267diff
changeset | 167 | Int.min (msecs, Time.toMilliseconds time + slack_msecs) | 
| 38092 
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
 blanchet parents: 
38084diff
changeset | 168 | |> Time.fromMilliseconds | 
| 40977 | 169 | val facts = filter_used_facts used_facts facts | 
| 43577 | 170 |          val (min_facts, {preplay, message, message_tail, ...}) =
 | 
| 42646 
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
 blanchet parents: 
42579diff
changeset | 171 | if length facts >= Config.get ctxt binary_min_facts then | 
| 40977 | 172 | binary_minimize (do_test new_timeout) facts | 
| 173 | else | |
| 174 | sublinear_minimize (do_test new_timeout) facts ([], result) | |
| 41091 
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
 blanchet parents: 
41090diff
changeset | 175 | val _ = print silent (fn () => cat_lines | 
| 43630 | 176 | ["Minimized to " ^ n_facts (map fst min_facts)] ^ | 
| 43577 | 177 | (case length (filter (curry (op =) Chained o snd o fst) min_facts) of | 
| 38698 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 blanchet parents: 
38696diff
changeset | 178 | 0 => "" | 
| 43577 | 179 | | n => "\n(including " ^ string_of_int n ^ " chained)") ^ ".") | 
| 180 | in (SOME min_facts, (preplay, message, message_tail)) end | |
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 181 |      | {outcome = SOME TimedOut, preplay, ...} =>
 | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 182 | (NONE, | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 183 | (preplay, | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 184 | fn _ => "Timeout: You can increase the time limit using the \ | 
| 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 185 | \\"timeout\" option (e.g., \"timeout = " ^ | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 186 | string_of_int (10 + msecs div 1000) ^ "\").", | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 187 | "")) | 
| 43052 
8d6a4978cc65
automatically minimize with Metis when this can be done within a few seconds
 blanchet parents: 
43051diff
changeset | 188 |      | {preplay, message, ...} =>
 | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 189 | (NONE, (preplay, prefix "Prover error: " o message, ""))) | 
| 43166 | 190 | handle ERROR msg => | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 191 | (NONE, (K (Failed_to_Play Metis), fn _ => "Error: " ^ msg, "")) | 
| 31236 | 192 | end | 
| 193 | ||
| 41265 
a393d6d8e198
let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
 blanchet parents: 
41259diff
changeset | 194 | fun run_minimize (params as {provers, ...}) i refs state =
 | 
| 38045 | 195 | let | 
| 196 | val ctxt = Proof.context_of state | |
| 38696 
4c6b65d6a135
quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
 blanchet parents: 
38617diff
changeset | 197 | val reserved = reserved_isar_keyword_table () | 
| 43043 
1406f6fc5dc3
normalize indices in chained facts to make sure that backtick facts (which often result in different names) are recognized + changed definition of urgent messages
 blanchet parents: 
43033diff
changeset | 198 | val chained_ths = normalize_chained_theorems (#facts (Proof.goal state)) | 
| 40204 
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
 blanchet parents: 
40200diff
changeset | 199 | val facts = | 
| 41091 
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
 blanchet parents: 
41090diff
changeset | 200 | refs | 
| 
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
 blanchet parents: 
41090diff
changeset | 201 | |> maps (map (apsnd single) o fact_from_ref ctxt reserved chained_ths) | 
| 38045 | 202 | in | 
| 203 | case subgoal_count state of | |
| 40132 
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
 wenzelm parents: 
40114diff
changeset | 204 | 0 => Output.urgent_message "No subgoal!" | 
| 41265 
a393d6d8e198
let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
 blanchet parents: 
41259diff
changeset | 205 | | n => case provers of | 
| 
a393d6d8e198
let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
 blanchet parents: 
41259diff
changeset | 206 | [] => error "No prover is set." | 
| 
a393d6d8e198
let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
 blanchet parents: 
41259diff
changeset | 207 | | prover :: _ => | 
| 
a393d6d8e198
let each prover minimizes its own stuff (rather than taking the first prover of the list to minimize every prover's stuff)
 blanchet parents: 
41259diff
changeset | 208 | (kill_provers (); | 
| 43064 
b6e61d22fa61
made "explicit_apply" smarter -- no need to force explicit applications in minimizer on all constants, better do it more fine granularly
 blanchet parents: 
43058diff
changeset | 209 | minimize_facts prover params false i n state facts | 
| 43261 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 210 | |> (fn (_, (preplay, message, message_tail)) => | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 211 | message (preplay ()) ^ message_tail | 
| 
a4aeb26a6362
make sure that the message tail (timing + TPTP important message) is preserved upon automatic minimization
 blanchet parents: 
43259diff
changeset | 212 | |> Output.urgent_message)) | 
| 38045 | 213 | end | 
| 214 | ||
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 215 | end; |