author | blanchet |
Fri, 22 Oct 2010 18:31:45 +0200 | |
changeset 40114 | acb75271cdce |
parent 40069 | 6f7bf79b1506 |
child 40132 | 7ee65dbffa31 |
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 |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
5 |
Minimization of fact list for Metis using ATPs. |
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 |
38021
e024504943d1
rename "ATP_Manager" ML module to "Sledgehammer";
blanchet
parents:
38015
diff
changeset
|
11 |
type params = Sledgehammer.params |
35867 | 12 |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
13 |
val minimize_facts : |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38745
diff
changeset
|
14 |
params -> int -> int -> Proof.state -> ((string * locality) * thm list) list |
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38745
diff
changeset
|
15 |
-> ((string * locality) * thm list) list option * string |
38996 | 16 |
val run_minimize : |
17 |
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
|
18 |
end; |
32525 | 19 |
|
38988 | 20 |
structure Sledgehammer_Minimize : SLEDGEHAMMER_MINIMIZE = |
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff
changeset
|
21 |
struct |
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff
changeset
|
22 |
|
39496
a52a4e4399c1
got caught once again by SML's pattern maching (ctor vs. var)
blanchet
parents:
39491
diff
changeset
|
23 |
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
|
24 |
open Sledgehammer_Util |
38988 | 25 |
open Sledgehammer_Filter |
38021
e024504943d1
rename "ATP_Manager" ML module to "Sledgehammer";
blanchet
parents:
38015
diff
changeset
|
26 |
open Sledgehammer |
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
27 |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
28 |
(* wrapper for calling external prover *) |
31236 | 29 |
|
39496
a52a4e4399c1
got caught once again by SML's pattern maching (ctor vs. var)
blanchet
parents:
39491
diff
changeset
|
30 |
fun string_for_failure Unprovable = "Unprovable." |
a52a4e4399c1
got caught once again by SML's pattern maching (ctor vs. var)
blanchet
parents:
39491
diff
changeset
|
31 |
| string_for_failure TimedOut = "Timed out." |
a52a4e4399c1
got caught once again by SML's pattern maching (ctor vs. var)
blanchet
parents:
39491
diff
changeset
|
32 |
| string_for_failure Interrupted = "Interrupted." |
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
|
33 |
| string_for_failure _ = "Unknown error." |
31236 | 34 |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
35 |
fun n_facts names = |
38698
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
36 |
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
|
37 |
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
|
38 |
(if n > 0 then |
38698
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
39 |
": " ^ (names |> map fst |
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
40 |
|> 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
|
41 |
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
|
42 |
"") |
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
|
43 |
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
|
44 |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
45 |
fun test_facts ({debug, verbose, overlord, provers, full_types, isar_proof, |
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
46 |
isar_shrink_factor, ...} : params) (prover : prover) |
40065 | 47 |
explicit_apply timeout i n state axioms = |
31236 | 48 |
let |
38015 | 49 |
val _ = |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
50 |
priority ("Testing " ^ n_facts (map fst axioms) ^ "...") |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
51 |
val params = |
38982
820b8221ed48
added "blocking" option to Sledgehammer to run in synchronous mode;
blanchet
parents:
38752
diff
changeset
|
52 |
{blocking = true, debug = debug, verbose = verbose, overlord = overlord, |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39496
diff
changeset
|
53 |
provers = provers, full_types = full_types, |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39496
diff
changeset
|
54 |
explicit_apply = explicit_apply, relevance_thresholds = (1.01, 1.01), |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
55 |
max_relevant = NONE, isar_proof = isar_proof, |
38998 | 56 |
isar_shrink_factor = isar_shrink_factor, timeout = timeout, expect = ""} |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
57 |
val axioms = |
40114 | 58 |
axioms |> maps (fn (n, ths) => ths |> map (Untranslated o pair n)) |
40069 | 59 |
val {goal, ...} = Proof.goal state |
40065 | 60 |
val problem = |
61 |
{state = state, goal = goal, subgoal = i, subgoal_count = n, |
|
62 |
axioms = axioms, only = true} |
|
63 |
val result as {outcome, used_axioms, ...} = prover params (K "") problem |
|
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36143
diff
changeset
|
64 |
in |
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
|
65 |
priority (case outcome of |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
66 |
SOME failure => string_for_failure failure |
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
67 |
| NONE => |
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
68 |
if length used_axioms = length axioms then "Found proof." |
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
69 |
else "Found proof with " ^ n_facts used_axioms ^ "."); |
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
|
70 |
result |
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36143
diff
changeset
|
71 |
end |
31236 | 72 |
|
73 |
(* minimalization of thms *) |
|
74 |
||
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
75 |
fun filter_used_axioms used = filter (member (op =) used o fst) |
38015 | 76 |
|
77 |
fun sublinear_minimize _ [] p = p |
|
78 |
| sublinear_minimize test (x :: xs) (seen, result) = |
|
79 |
case test (xs @ seen) of |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
80 |
result as {outcome = NONE, used_axioms, ...} : prover_result => |
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
81 |
sublinear_minimize test (filter_used_axioms used_axioms xs) |
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
82 |
(filter_used_axioms used_axioms seen, result) |
38015 | 83 |
| _ => sublinear_minimize test xs (x :: seen, result) |
84 |
||
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
|
85 |
(* Give the ATP some slack. The ATP gets further slack because the Sledgehammer |
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
|
86 |
preprocessing time is included in the estimate below but isn't part of the |
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
|
87 |
timeout. *) |
38590
bd443b426d56
get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
blanchet
parents:
38589
diff
changeset
|
88 |
val fudge_msecs = 1000 |
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
|
89 |
|
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
90 |
fun minimize_facts {provers = [], ...} _ _ _ _ = error "No prover is set." |
40065 | 91 |
| minimize_facts (params as {provers = prover_name :: _, timeout, ...}) i n |
92 |
state axioms = |
|
31236 | 93 |
let |
36378 | 94 |
val thy = Proof.theory_of state |
40063
d086e3699e78
bring ATPs and SMT solvers more in line with each other
blanchet
parents:
40062
diff
changeset
|
95 |
val prover = get_prover thy 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
|
96 |
val msecs = Time.toMilliseconds timeout |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
97 |
val _ = priority ("Sledgehammer minimize: " ^ quote prover_name ^ ".") |
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
98 |
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:
38094
diff
changeset
|
99 |
val (_, hyp_ts, concl_t) = strip_subgoal goal i |
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
100 |
val explicit_apply = |
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
101 |
not (forall (Meson.is_fol_term thy) |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
102 |
(concl_t :: hyp_ts @ maps (map prop_of o snd) axioms)) |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
103 |
fun do_test timeout = |
40065 | 104 |
test_facts params prover explicit_apply 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
|
105 |
val timer = Timer.startRealTimer () |
31236 | 106 |
in |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
107 |
(case do_test timeout axioms of |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
108 |
result as {outcome = NONE, used_axioms, ...} => |
38015 | 109 |
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
|
110 |
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
|
111 |
val new_timeout = |
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
|
112 |
Int.min (msecs, Time.toMilliseconds time + fudge_msecs) |
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
|
113 |
|> Time.fromMilliseconds |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
114 |
val (min_thms, {message, ...}) = |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
115 |
sublinear_minimize (do_test new_timeout) |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
116 |
(filter_used_axioms used_axioms axioms) ([], result) |
38094 | 117 |
val n = length min_thms |
38015 | 118 |
val _ = priority (cat_lines |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
119 |
["Minimized: " ^ string_of_int n ^ " fact" ^ plural_s n] ^ |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38745
diff
changeset
|
120 |
(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
|
121 |
0 => "" |
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
122 |
| n => " (including " ^ Int.toString n ^ " chained)") ^ ".") |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
123 |
in (SOME min_thms, message) end |
38015 | 124 |
| {outcome = SOME TimedOut, ...} => |
125 |
(NONE, "Timeout: You can increase the time limit using the \"timeout\" \ |
|
126 |
\option (e.g., \"timeout = " ^ |
|
127 |
string_of_int (10 + msecs div 1000) ^ " s\").") |
|
128 |
| {outcome = SOME UnknownError, ...} => |
|
129 |
(* Failure sometimes mean timeout, unfortunately. *) |
|
130 |
(NONE, "Failure: No proof was found with the current time limit. You \ |
|
131 |
\can increase the time limit using the \"timeout\" \ |
|
132 |
\option (e.g., \"timeout = " ^ |
|
133 |
string_of_int (10 + msecs div 1000) ^ " s\").") |
|
134 |
| {message, ...} => (NONE, "ATP error: " ^ message)) |
|
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37926
diff
changeset
|
135 |
handle ERROR msg => (NONE, "Error: " ^ msg) |
31236 | 136 |
end |
137 |
||
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
38100
diff
changeset
|
138 |
fun run_minimize params i refs state = |
38045 | 139 |
let |
140 |
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
|
141 |
val reserved = reserved_isar_keyword_table () |
38045 | 142 |
val chained_ths = #facts (Proof.goal state) |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
143 |
val axioms = |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38745
diff
changeset
|
144 |
maps (map (apsnd single) |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
145 |
o name_thm_pairs_from_ref ctxt reserved chained_ths) refs |
38045 | 146 |
in |
147 |
case subgoal_count state of |
|
148 |
0 => priority "No subgoal!" |
|
149 |
| n => |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39496
diff
changeset
|
150 |
(kill_provers (); |
40061
71cc5aac8b76
generalization of the Sledgehammer minimizer, to make it possible to handle SMT solvers as well
blanchet
parents:
40060
diff
changeset
|
151 |
priority (#2 (minimize_facts params i n state axioms))) |
38045 | 152 |
end |
153 |
||
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
154 |
end; |