author | blanchet |
Tue, 03 May 2011 00:10:22 +0200 | |
changeset 42646 | 4781fcd53572 |
parent 42579 | 2552c09b1a72 |
child 42724 | 4d6bcf846759 |
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:
36369
diff
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 |
38988 | 10 |
type locality = Sledgehammer_Filter.locality |
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
40983
diff
changeset
|
11 |
type params = Sledgehammer_Provers.params |
35867 | 12 |
|
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42579
diff
changeset
|
13 |
val binary_min_facts : int Config.T |
41255
a80024d7b71b
added debugging option to find out how good the relevance filter was at identifying relevant facts
blanchet
parents:
41242
diff
changeset
|
14 |
val filter_used_facts : ''a list -> (''a * 'b) list -> (''a * 'b) list |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
15 |
val minimize_facts : |
41742
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
16 |
string -> params -> bool option -> bool -> int -> int -> Proof.state |
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
17 |
-> ((string * locality) * thm list) list |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38745
diff
changeset
|
18 |
-> ((string * locality) * thm list) list option * string |
38996 | 19 |
val run_minimize : |
20 |
params -> int -> (Facts.ref * Attrib.src list) list -> Proof.state -> unit |
|
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
21 |
end; |
32525 | 22 |
|
38988 | 23 |
structure Sledgehammer_Minimize : SLEDGEHAMMER_MINIMIZE = |
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff
changeset
|
24 |
struct |
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff
changeset
|
25 |
|
39496
a52a4e4399c1
got caught once again by SML's pattern maching (ctor vs. var)
blanchet
parents:
39491
diff
changeset
|
26 |
open ATP_Proof |
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36063
diff
changeset
|
27 |
open Sledgehammer_Util |
38988 | 28 |
open Sledgehammer_Filter |
41087
d7b5fd465198
split "Sledgehammer" module into two parts, to resolve forthcoming dependency problems
blanchet
parents:
40983
diff
changeset
|
29 |
open Sledgehammer_Provers |
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
30 |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
31 |
(* wrapper for calling external prover *) |
31236 | 32 |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
33 |
fun n_facts names = |
38698
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
34 |
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:
40060
diff
changeset
|
35 |
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:
38084
diff
changeset
|
36 |
(if n > 0 then |
38698
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
37 |
": " ^ (names |> map fst |
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
38 |
|> 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:
38084
diff
changeset
|
39 |
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:
38084
diff
changeset
|
40 |
"") |
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:
38084
diff
changeset
|
41 |
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:
38084
diff
changeset
|
42 |
|
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
43 |
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:
41090
diff
changeset
|
44 |
|
42180
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
42060
diff
changeset
|
45 |
fun test_facts ({debug, verbose, overlord, provers, monomorphize_limit, |
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
42060
diff
changeset
|
46 |
type_sys, isar_proof, isar_shrink_factor, ...} : params) |
41742
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
47 |
explicit_apply_opt silent (prover : prover) timeout i n state facts = |
31236 | 48 |
let |
41742
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
49 |
val thy = Proof.theory_of state |
41277
1369c27c6966
reduce the minimizer slack and add verbose information
blanchet
parents:
41267
diff
changeset
|
50 |
val _ = |
1369c27c6966
reduce the minimizer slack and add verbose information
blanchet
parents:
41267
diff
changeset
|
51 |
print silent (fn () => |
1369c27c6966
reduce the minimizer slack and add verbose information
blanchet
parents:
41267
diff
changeset
|
52 |
"Testing " ^ n_facts (map fst facts) ^ |
1369c27c6966
reduce the minimizer slack and add verbose information
blanchet
parents:
41267
diff
changeset
|
53 |
(if verbose then " (timeout: " ^ string_from_time timeout ^ ")" |
1369c27c6966
reduce the minimizer slack and add verbose information
blanchet
parents:
41267
diff
changeset
|
54 |
else "") ^ "...") |
41742
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
55 |
val {goal, ...} = Proof.goal state |
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
56 |
val explicit_apply = |
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
57 |
case explicit_apply_opt of |
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
58 |
SOME explicit_apply => explicit_apply |
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
59 |
| NONE => |
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
60 |
let val (_, hyp_ts, concl_t) = strip_subgoal goal i in |
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
61 |
not (forall (Meson.is_fol_term thy) |
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
62 |
(concl_t :: hyp_ts @ maps (map prop_of o snd) facts)) |
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
63 |
end |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
64 |
val params = |
42060
889d767ce5f4
make Minimizer honor "verbose" and "debug" options better
blanchet
parents:
41824
diff
changeset
|
65 |
{debug = debug, verbose = verbose, overlord = overlord, blocking = true, |
41138
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41134
diff
changeset
|
66 |
provers = provers, type_sys = type_sys, explicit_apply = explicit_apply, |
eb80538166b6
implemented partially-typed "tags" type encoding
blanchet
parents:
41134
diff
changeset
|
67 |
relevance_thresholds = (1.01, 1.01), max_relevant = NONE, |
42579
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42444
diff
changeset
|
68 |
monomorphize_limit = monomorphize_limit, isar_proof = isar_proof, |
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42444
diff
changeset
|
69 |
isar_shrink_factor = isar_shrink_factor, slicing = false, |
2552c09b1a72
implement the new ATP type system in Sledgehammer
blanchet
parents:
42444
diff
changeset
|
70 |
timeout = 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:
40200
diff
changeset
|
71 |
val facts = |
41090 | 72 |
facts |> maps (fn (n, ths) => ths |> map (Untranslated_Fact o pair n)) |
40065 | 73 |
val problem = |
74 |
{state = state, goal = goal, subgoal = i, subgoal_count = n, |
|
41741 | 75 |
facts = facts, smt_filter = NONE} |
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40200
diff
changeset
|
76 |
val result as {outcome, used_facts, ...} = prover params (K "") problem |
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36143
diff
changeset
|
77 |
in |
41277
1369c27c6966
reduce the minimizer slack and add verbose information
blanchet
parents:
41267
diff
changeset
|
78 |
print silent (fn () => |
1369c27c6966
reduce the minimizer slack and add verbose information
blanchet
parents:
41267
diff
changeset
|
79 |
case outcome of |
41745 | 80 |
SOME failure => string_for_failure failure |
81 |
| NONE => if length used_facts = length facts then "Found proof." |
|
82 |
else "Found proof with " ^ n_facts used_facts ^ "."); |
|
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:
38084
diff
changeset
|
83 |
result |
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36143
diff
changeset
|
84 |
end |
31236 | 85 |
|
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40200
diff
changeset
|
86 |
(* minimalization of facts *) |
31236 | 87 |
|
40977 | 88 |
(* The sublinear algorithm works well in almost all situations, except when the |
89 |
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:
41265
diff
changeset
|
90 |
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:
41265
diff
changeset
|
91 |
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:
41265
diff
changeset
|
92 |
reported by the prover. *) |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42579
diff
changeset
|
93 |
val binary_min_facts = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42579
diff
changeset
|
94 |
Attrib.setup_config_int @{binding sledgehammer_minimize_binary_min_facts} |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42579
diff
changeset
|
95 |
(K 20) |
40977 | 96 |
|
40204
da97d75e20e6
standardize on "fact" terminology (vs. "axiom" or "theorem") in Sledgehammer -- but keep "Axiom" in the lower-level "ATP_Problem" module
blanchet
parents:
40200
diff
changeset
|
97 |
fun filter_used_facts used = filter (member (op =) used o fst) |
38015 | 98 |
|
99 |
fun sublinear_minimize _ [] p = p |
|
100 |
| sublinear_minimize test (x :: xs) (seen, result) = |
|
101 |
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:
40200
diff
changeset
|
102 |
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:
40200
diff
changeset
|
103 |
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:
40200
diff
changeset
|
104 |
(filter_used_facts used_facts seen, result) |
38015 | 105 |
| _ => sublinear_minimize test xs (x :: seen, result) |
106 |
||
40977 | 107 |
fun binary_minimize test xs = |
108 |
let |
|
109 |
fun p xs = #outcome (test xs : prover_result) = NONE |
|
110 |
fun split [] p = p |
|
111 |
| split [h] (l, r) = (h :: l, r) |
|
112 |
| split (h1 :: h2 :: t) (l, r) = split t (h1 :: l, h2 :: r) |
|
41743 | 113 |
fun min _ _ [] = raise Empty |
114 |
| min _ _ [s0] = [s0] |
|
115 |
| min depth sup xs = |
|
116 |
let |
|
117 |
(* |
|
118 |
val _ = warning (replicate_string depth " " ^ "{" ^ (" " ^ |
|
119 |
n_facts (map fst sup) ^ " and " ^ |
|
120 |
n_facts (map fst xs))) |
|
121 |
*) |
|
122 |
val (l0, r0) = split xs ([], []) |
|
123 |
in |
|
40977 | 124 |
if p (sup @ l0) then |
41743 | 125 |
min (depth + 1) sup l0 |
40977 | 126 |
else if p (sup @ r0) then |
41743 | 127 |
min (depth + 1) sup r0 |
40977 | 128 |
else |
129 |
let |
|
41743 | 130 |
val l = min (depth + 1) (sup @ r0) l0 |
131 |
val r = min (depth + 1) (sup @ l) r0 |
|
40977 | 132 |
in l @ r end |
133 |
end |
|
41743 | 134 |
(* |
135 |
|> tap (fn _ => warning (replicate_string depth " " ^ "}")) |
|
136 |
*) |
|
40977 | 137 |
val xs = |
41743 | 138 |
case min 0 [] xs of |
40977 | 139 |
[x] => if p [] then [] else [x] |
140 |
| xs => xs |
|
141 |
in (xs, test xs) end |
|
142 |
||
143 |
(* Give the external prover some slack. The ATP gets further slack because the |
|
144 |
Sledgehammer preprocessing time is included in the estimate below but isn't |
|
145 |
part of the timeout. *) |
|
41277
1369c27c6966
reduce the minimizer slack and add verbose information
blanchet
parents:
41267
diff
changeset
|
146 |
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:
38084
diff
changeset
|
147 |
|
41742
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
148 |
fun minimize_facts prover_name (params as {timeout, ...}) explicit_apply_opt |
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
149 |
silent i n state facts = |
31236 | 150 |
let |
40941
a3e6f8634a11
replace "smt" prover with specific SMT solvers, e.g. "z3" -- whatever the SMT module gives us
blanchet
parents:
40553
diff
changeset
|
151 |
val ctxt = Proof.context_of state |
42444
8e5438dc70bb
cleanup: get rid of "may_slice" arguments without changing semantics
blanchet
parents:
42443
diff
changeset
|
152 |
val prover = get_prover ctxt false prover_name |
38590
bd443b426d56
get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
blanchet
parents:
38589
diff
changeset
|
153 |
val msecs = Time.toMilliseconds timeout |
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
154 |
val _ = print silent (fn () => "Sledgehammer minimize: " ^ |
40977 | 155 |
quote prover_name ^ ".") |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
156 |
fun do_test timeout = |
41742
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
157 |
test_facts params explicit_apply_opt 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:
38084
diff
changeset
|
158 |
val timer = Timer.startRealTimer () |
31236 | 159 |
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:
40200
diff
changeset
|
160 |
(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:
40200
diff
changeset
|
161 |
result as {outcome = NONE, used_facts, ...} => |
38015 | 162 |
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:
38084
diff
changeset
|
163 |
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:
38084
diff
changeset
|
164 |
val new_timeout = |
41277
1369c27c6966
reduce the minimizer slack and add verbose information
blanchet
parents:
41267
diff
changeset
|
165 |
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:
38084
diff
changeset
|
166 |
|> Time.fromMilliseconds |
40977 | 167 |
val facts = filter_used_facts used_facts facts |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
168 |
val (min_thms, {message, ...}) = |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42579
diff
changeset
|
169 |
if length facts >= Config.get ctxt binary_min_facts then |
40977 | 170 |
binary_minimize (do_test new_timeout) facts |
171 |
else |
|
172 |
sublinear_minimize (do_test new_timeout) facts ([], result) |
|
38094 | 173 |
val n = length min_thms |
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
174 |
val _ = print silent (fn () => cat_lines |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
175 |
["Minimized: " ^ string_of_int n ^ " fact" ^ plural_s n] ^ |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38745
diff
changeset
|
176 |
(case length (filter (curry (op =) Chained o snd o fst) min_thms) of |
38698
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
177 |
0 => "" |
41491 | 178 |
| n => " (including " ^ string_of_int n ^ " chained)") ^ ".") |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
179 |
in (SOME min_thms, message) end |
38015 | 180 |
| {outcome = SOME TimedOut, ...} => |
181 |
(NONE, "Timeout: You can increase the time limit using the \"timeout\" \ |
|
182 |
\option (e.g., \"timeout = " ^ |
|
40341
03156257040f
standardize on seconds for Nitpick and Sledgehammer timeouts
blanchet
parents:
40205
diff
changeset
|
183 |
string_of_int (10 + msecs div 1000) ^ "\").") |
40977 | 184 |
| {message, ...} => (NONE, "Prover error: " ^ message)) |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37926
diff
changeset
|
185 |
handle ERROR msg => (NONE, "Error: " ^ msg) |
31236 | 186 |
end |
187 |
||
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:
41259
diff
changeset
|
188 |
fun run_minimize (params as {provers, ...}) i refs state = |
38045 | 189 |
let |
190 |
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:
38617
diff
changeset
|
191 |
val reserved = reserved_isar_keyword_table () |
38045 | 192 |
val chained_ths = #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:
40200
diff
changeset
|
193 |
val facts = |
41091
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
194 |
refs |
0afdf5cde874
implicitly call the minimizer for SMT solvers that don't return an unsat core
blanchet
parents:
41090
diff
changeset
|
195 |
|> maps (map (apsnd single) o fact_from_ref ctxt reserved chained_ths) |
38045 | 196 |
in |
197 |
case subgoal_count state of |
|
40132
7ee65dbffa31
renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents:
40114
diff
changeset
|
198 |
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:
41259
diff
changeset
|
199 |
| 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:
41259
diff
changeset
|
200 |
[] => 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:
41259
diff
changeset
|
201 |
| 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:
41259
diff
changeset
|
202 |
(kill_provers (); |
41742
11e862c68b40
automatically minimize Z3-as-an-ATP proofs (cf. CVC3 and Yices)
blanchet
parents:
41741
diff
changeset
|
203 |
minimize_facts prover params NONE false i n state facts |
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:
41259
diff
changeset
|
204 |
|> #2 |> Output.urgent_message) |
38045 | 205 |
end |
206 |
||
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
207 |
end; |