| author | blanchet | 
| Wed, 08 Sep 2010 15:57:50 +0200 | |
| changeset 39222 | decf607a5a67 | 
| parent 39005 | 42fcb25de082 | 
| child 39262 | bdfcf2434601 | 
| 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 | |
| 35867 | 5 | Minimization of theorem list for Metis using automatic theorem 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 | 
| 38021 
e024504943d1
rename "ATP_Manager" ML module to "Sledgehammer";
 blanchet parents: 
38015diff
changeset | 11 | type params = Sledgehammer.params | 
| 35867 | 12 | |
| 38050 | 13 | val minimize_theorems : | 
| 38752 
6628adcae4a7
consider "locality" when assigning weights to facts
 blanchet parents: 
38745diff
changeset | 14 | params -> int -> int -> Proof.state -> ((string * locality) * thm list) list | 
| 
6628adcae4a7
consider "locality" when assigning weights to facts
 blanchet parents: 
38745diff
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: 
35865diff
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 | |
| 38586 | 23 | open ATP_Systems | 
| 36142 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 24 | open Sledgehammer_Util | 
| 38988 | 25 | open Sledgehammer_Filter | 
| 39004 
f1b465f889b5
translate the axioms to FOF once and for all ATPs
 blanchet parents: 
38998diff
changeset | 26 | open Sledgehammer_Translate | 
| 38988 | 27 | open Sledgehammer_Reconstruct | 
| 38021 
e024504943d1
rename "ATP_Manager" ML module to "Sledgehammer";
 blanchet parents: 
38015diff
changeset | 28 | open Sledgehammer | 
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 29 | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 30 | (* wrapper for calling external prover *) | 
| 31236 | 31 | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 32 | fun string_for_failure Unprovable = "Unprovable." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 33 | | string_for_failure TimedOut = "Timed out." | 
| 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 | 34 | | string_for_failure _ = "Unknown error." | 
| 31236 | 35 | |
| 38698 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 blanchet parents: 
38696diff
changeset | 36 | fun n_theorems names = | 
| 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 blanchet parents: 
38696diff
changeset | 37 | let val n = length names 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: 
38084diff
changeset | 38 | string_of_int n ^ " theorem" ^ plural_s n ^ | 
| 
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 | 39 | (if n > 0 then | 
| 38698 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 blanchet parents: 
38696diff
changeset | 40 | ": " ^ (names |> map fst | 
| 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 blanchet parents: 
38696diff
changeset | 41 | |> 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 | 42 | 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 | 43 | "") | 
| 
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 | 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 | 45 | |
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 46 | fun test_theorems ({debug, verbose, overlord, atps, full_types, isar_proof,
 | 
| 38741 | 47 | isar_shrink_factor, ...} : params) | 
| 38100 
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
 blanchet parents: 
38094diff
changeset | 48 | (prover : prover) explicit_apply timeout subgoal state | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 49 | axioms = | 
| 31236 | 50 | let | 
| 38015 | 51 | val _ = | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 52 |       priority ("Testing " ^ n_theorems (map fst axioms) ^ "...")
 | 
| 38100 
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
 blanchet parents: 
38094diff
changeset | 53 | val params = | 
| 38982 
820b8221ed48
added "blocking" option to Sledgehammer to run in synchronous mode;
 blanchet parents: 
38752diff
changeset | 54 |       {blocking = true, debug = debug, verbose = verbose, overlord = overlord,
 | 
| 
820b8221ed48
added "blocking" option to Sledgehammer to run in synchronous mode;
 blanchet parents: 
38752diff
changeset | 55 | atps = atps, full_types = full_types, explicit_apply = explicit_apply, | 
| 38998 | 56 | relevance_thresholds = (1.01, 1.01), | 
| 57 | max_relevant = SOME 65536 (* a large number *), isar_proof = isar_proof, | |
| 58 | isar_shrink_factor = isar_shrink_factor, timeout = timeout, expect = ""} | |
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 59 | val axioms = maps (fn (n, ths) => map (pair n) ths) axioms | 
| 38998 | 60 |     val {context = ctxt, goal, ...} = Proof.goal state
 | 
| 39004 
f1b465f889b5
translate the axioms to FOF once and for all ATPs
 blanchet parents: 
38998diff
changeset | 61 | val problem = | 
| 
f1b465f889b5
translate the axioms to FOF once and for all ATPs
 blanchet parents: 
38998diff
changeset | 62 |       {ctxt = ctxt, goal = goal, subgoal = subgoal,
 | 
| 39005 | 63 | axioms = map (prepare_axiom ctxt) axioms} | 
| 38698 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 blanchet parents: 
38696diff
changeset | 64 |     val result as {outcome, used_thm_names, ...} = prover params (K "") problem
 | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 65 | 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: 
38084diff
changeset | 66 | priority (case outcome of | 
| 
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 | 67 | NONE => | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 68 | if length used_thm_names = length axioms then | 
| 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 | 69 | "Found proof." | 
| 
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 | 70 | 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 | 71 | "Found proof with " ^ n_theorems used_thm_names ^ "." | 
| 
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 | 72 | | SOME failure => string_for_failure failure); | 
| 
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 | 73 | result | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 74 | end | 
| 31236 | 75 | |
| 76 | (* minimalization of thms *) | |
| 77 | ||
| 38698 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 blanchet parents: 
38696diff
changeset | 78 | fun filter_used_facts used = filter (member (op =) used o fst) | 
| 38015 | 79 | |
| 80 | fun sublinear_minimize _ [] p = p | |
| 81 | | sublinear_minimize test (x :: xs) (seen, result) = | |
| 82 | case test (xs @ seen) of | |
| 38488 | 83 |       result as {outcome = NONE, proof, used_thm_names, ...} : prover_result =>
 | 
| 38015 | 84 | sublinear_minimize test (filter_used_facts used_thm_names xs) | 
| 85 | (filter_used_facts used_thm_names seen, result) | |
| 86 | | _ => sublinear_minimize test xs (x :: seen, result) | |
| 87 | ||
| 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 | (* 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: 
38084diff
changeset | 89 | 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: 
38084diff
changeset | 90 | timeout. *) | 
| 38590 
bd443b426d56
get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
 blanchet parents: 
38589diff
changeset | 91 | 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: 
38084diff
changeset | 92 | |
| 
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 | 93 | fun minimize_theorems {atps = [], ...} _ _ _ _ = error "No ATP is set."
 | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 94 |   | minimize_theorems (params as {debug, atps = atp :: _, full_types,
 | 
| 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 95 | isar_proof, isar_shrink_factor, timeout, ...}) | 
| 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 96 | i n state axioms = | 
| 31236 | 97 | let | 
| 36378 | 98 | val thy = Proof.theory_of 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 | 99 | val prover = get_prover_fun thy atp | 
| 38590 
bd443b426d56
get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
 blanchet parents: 
38589diff
changeset | 100 | val msecs = Time.toMilliseconds timeout | 
| 
bd443b426d56
get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
 blanchet parents: 
38589diff
changeset | 101 |     val _ = priority ("Sledgehammer minimize: ATP " ^ quote atp ^ ".")
 | 
| 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 | 102 |     val {context = ctxt, 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 | 103 | 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: 
38094diff
changeset | 104 | val explicit_apply = | 
| 
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
 blanchet parents: 
38094diff
changeset | 105 | 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: 
38741diff
changeset | 106 | (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: 
38094diff
changeset | 107 | fun do_test timeout = | 
| 
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
 blanchet parents: 
38094diff
changeset | 108 | test_theorems params prover explicit_apply timeout i 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 | 109 | val timer = Timer.startRealTimer () | 
| 31236 | 110 | in | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 111 | (case do_test timeout axioms of | 
| 38083 
c4b57f68ddb3
remove the "extra_clauses" business introduced in 19a5f1c8a844;
 blanchet parents: 
38050diff
changeset | 112 |        result as {outcome = NONE, pool, used_thm_names,
 | 
| 
c4b57f68ddb3
remove the "extra_clauses" business introduced in 19a5f1c8a844;
 blanchet parents: 
38050diff
changeset | 113 | conjecture_shape, ...} => | 
| 38015 | 114 | 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 | 115 | 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 | 116 | 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: 
38084diff
changeset | 117 | 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: 
38084diff
changeset | 118 | |> Time.fromMilliseconds | 
| 38282 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 blanchet parents: 
38100diff
changeset | 119 |          val (min_thms, {proof, axiom_names, ...}) =
 | 
| 38100 
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
 blanchet parents: 
38094diff
changeset | 120 | sublinear_minimize (do_test new_timeout) | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 121 | (filter_used_facts used_thm_names axioms) ([], result) | 
| 38094 | 122 | val n = length min_thms | 
| 38015 | 123 | val _ = priority (cat_lines | 
| 38094 | 124 | ["Minimized: " ^ string_of_int n ^ " theorem" ^ plural_s n] ^ | 
| 38752 
6628adcae4a7
consider "locality" when assigning weights to facts
 blanchet parents: 
38745diff
changeset | 125 | (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: 
38696diff
changeset | 126 | 0 => "" | 
| 
d19c3a7ce38b
clean handling of whether a fact is chained or not;
 blanchet parents: 
38696diff
changeset | 127 | | n => " (including " ^ Int.toString n ^ " chained)") ^ ".") | 
| 38015 | 128 | in | 
| 129 | (SOME min_thms, | |
| 130 | proof_text isar_proof | |
| 131 | (pool, debug, isar_shrink_factor, ctxt, conjecture_shape) | |
| 38282 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 blanchet parents: 
38100diff
changeset | 132 | (full_types, K "", proof, axiom_names, goal, i) |> fst) | 
| 38015 | 133 | end | 
| 134 |      | {outcome = SOME TimedOut, ...} =>
 | |
| 135 | (NONE, "Timeout: You can increase the time limit using the \"timeout\" \ | |
| 136 | \option (e.g., \"timeout = " ^ | |
| 137 | string_of_int (10 + msecs div 1000) ^ " s\").") | |
| 138 |      | {outcome = SOME UnknownError, ...} =>
 | |
| 139 | (* Failure sometimes mean timeout, unfortunately. *) | |
| 140 | (NONE, "Failure: No proof was found with the current time limit. You \ | |
| 141 | \can increase the time limit using the \"timeout\" \ | |
| 142 | \option (e.g., \"timeout = " ^ | |
| 143 | string_of_int (10 + msecs div 1000) ^ " s\").") | |
| 144 |      | {message, ...} => (NONE, "ATP error: " ^ message))
 | |
| 37994 
b04307085a09
make TPTP generator accept full first-order formulas
 blanchet parents: 
37926diff
changeset | 145 | handle ERROR msg => (NONE, "Error: " ^ msg) | 
| 31236 | 146 | end | 
| 147 | ||
| 38282 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 blanchet parents: 
38100diff
changeset | 148 | fun run_minimize params i refs state = | 
| 38045 | 149 | let | 
| 150 | 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 | 151 | val reserved = reserved_isar_keyword_table () | 
| 38045 | 152 | 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: 
38741diff
changeset | 153 | val axioms = | 
| 38752 
6628adcae4a7
consider "locality" when assigning weights to facts
 blanchet parents: 
38745diff
changeset | 154 | maps (map (apsnd single) | 
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 155 | o name_thm_pairs_from_ref ctxt reserved chained_ths) refs | 
| 38045 | 156 | in | 
| 157 | case subgoal_count state of | |
| 158 | 0 => priority "No subgoal!" | |
| 159 | | n => | |
| 38744 
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
 blanchet parents: 
38741diff
changeset | 160 | (kill_atps (); priority (#2 (minimize_theorems params i n state axioms))) | 
| 38045 | 161 | end | 
| 162 | ||
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 163 | end; |