| author | ballarin | 
| Thu, 13 May 2010 14:47:15 +0200 | |
| changeset 36906 | 9eff24f4e5db | 
| parent 36607 | e5f7235f39c5 | 
| child 36909 | 7d5587f6d5f7 | 
| permissions | -rw-r--r-- | 
| 36375 
2482446a604c
move the minimizer to the Sledgehammer directory
 blanchet parents: 
36370diff
changeset | 1 | (* Title: HOL/Tools/Sledgehammer/sledgehammer_fact_minimizer.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 | |
| 36375 
2482446a604c
move the minimizer to the Sledgehammer directory
 blanchet parents: 
36370diff
changeset | 8 | signature SLEDGEHAMMER_FACT_MINIMIZER = | 
| 32525 | 9 | sig | 
| 35969 | 10 | type params = ATP_Manager.params | 
| 35867 | 11 | type prover_result = ATP_Manager.prover_result | 
| 12 | ||
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 13 | val minimize_theorems : | 
| 36481 
af99c98121d6
make Sledgehammer more friendly if no subgoal is left
 blanchet parents: 
36402diff
changeset | 14 | params -> int -> int -> Proof.state -> (string * thm list) list | 
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 15 | -> (string * thm list) list option * string | 
| 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 16 | end; | 
| 32525 | 17 | |
| 36375 
2482446a604c
move the minimizer to the Sledgehammer directory
 blanchet parents: 
36370diff
changeset | 18 | structure Sledgehammer_Fact_Minimizer : SLEDGEHAMMER_FACT_MINIMIZER = | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 19 | struct | 
| 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 20 | |
| 36142 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 21 | open Sledgehammer_Util | 
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 22 | open Sledgehammer_Fact_Preprocessor | 
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
35969diff
changeset | 23 | open Sledgehammer_Proof_Reconstruct | 
| 35867 | 24 | open ATP_Manager | 
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 25 | |
| 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 26 | (* Linear minimization algorithm *) | 
| 33492 
4168294a9f96
Command atp_minimize uses the naive linear algorithm now
 nipkow parents: 
33316diff
changeset | 27 | |
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 28 | fun linear_minimize test s = | 
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 29 | let | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 30 | fun aux [] p = p | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 31 | | aux (x :: xs) (needed, result) = | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 32 | case test (xs @ needed) of | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 33 | SOME result => aux xs (needed, result) | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 34 | | NONE => aux xs (x :: needed, result) | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 35 | in aux s end | 
| 31236 | 36 | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: diff
changeset | 37 | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 38 | (* wrapper for calling external prover *) | 
| 31236 | 39 | |
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 40 | fun string_for_failure Unprovable = "Unprovable." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 41 | | string_for_failure TimedOut = "Timed out." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 42 | | string_for_failure OutOfResources = "Failed." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 43 | | string_for_failure OldSpass = "Error." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 44 | | string_for_failure MalformedOutput = "Error." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 45 | | string_for_failure UnknownError = "Failed." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 46 | fun string_for_outcome NONE = "Success." | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 47 | | string_for_outcome (SOME failure) = string_for_failure failure | 
| 31236 | 48 | |
| 35969 | 49 | fun sledgehammer_test_theorems (params as {full_types, ...} : params) prover
 | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 50 | timeout subgoal state filtered_clauses name_thms_pairs = | 
| 31236 | 51 | let | 
| 36142 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 52 | val num_theorems = length name_thms_pairs | 
| 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 53 |     val _ = priority ("Testing " ^ string_of_int num_theorems ^
 | 
| 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 54 | " theorem" ^ plural_s num_theorems ^ "...") | 
| 32525 | 55 | val name_thm_pairs = maps (fn (n, ths) => map (pair n) ths) name_thms_pairs | 
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 56 | val axclauses = cnf_rules_pairs (Proof.theory_of state) name_thm_pairs | 
| 36263 
56bf63d70120
use "Proof.goal" in Sledgehammer's minimizer (just like everywhere else in Sledgehammer), not "Proof.raw_goal"
 blanchet parents: 
36232diff
changeset | 57 |     val {context = ctxt, facts, goal} = Proof.goal state
 | 
| 32941 
72d48e333b77
eliminated extraneous wrapping of public records;
 wenzelm parents: 
32937diff
changeset | 58 | val problem = | 
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
35969diff
changeset | 59 |      {subgoal = subgoal, goal = (ctxt, (facts, goal)),
 | 
| 35969 | 60 |       relevance_override = {add = [], del = [], only = false},
 | 
| 36232 
4d425766a47f
don't redo an axiom selection in the first round of Sledgehammer "minimize"!;
 blanchet parents: 
36231diff
changeset | 61 | axiom_clauses = SOME axclauses, | 
| 
4d425766a47f
don't redo an axiom selection in the first round of Sledgehammer "minimize"!;
 blanchet parents: 
36231diff
changeset | 62 | filtered_clauses = SOME (the_default axclauses filtered_clauses)} | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 63 | in | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 64 | prover params (K "") timeout problem | 
| 36607 
e5f7235f39c5
made sml/nj happy about Sledgehammer and Nitpick (cf. 6f11c9b1fb3e, 3c2438efe224)
 krauss parents: 
36488diff
changeset | 65 | |> tap (fn result : prover_result => | 
| 
e5f7235f39c5
made sml/nj happy about Sledgehammer and Nitpick (cf. 6f11c9b1fb3e, 3c2438efe224)
 krauss parents: 
36488diff
changeset | 66 | priority (string_for_outcome (#outcome result))) | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 67 | end | 
| 31236 | 68 | |
| 69 | (* minimalization of thms *) | |
| 70 | ||
| 36378 | 71 | fun minimize_theorems (params as {debug, atps, minimize_timeout, isar_proof,
 | 
| 36488 
32c92af68ec9
remove Sledgehammer's "sorts" option to annotate variables with sorts in proof;
 blanchet parents: 
36481diff
changeset | 72 | shrink_factor, ...}) | 
| 36481 
af99c98121d6
make Sledgehammer more friendly if no subgoal is left
 blanchet parents: 
36402diff
changeset | 73 | i n state name_thms_pairs = | 
| 31236 | 74 | let | 
| 36378 | 75 | val thy = Proof.theory_of state | 
| 76 | val prover = case atps of | |
| 77 | [atp_name] => get_prover thy atp_name | |
| 78 | | _ => error "Expected a single ATP." | |
| 35969 | 79 | val msecs = Time.toMilliseconds minimize_timeout | 
| 31236 | 80 | val _ = | 
| 36378 | 81 |       priority ("Sledgehammer minimizer: ATP " ^ quote (the_single atps) ^
 | 
| 36224 | 82 | " with a time limit of " ^ string_of_int msecs ^ " ms.") | 
| 35969 | 83 | val test_thms_fun = | 
| 36063 
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
 blanchet parents: 
35969diff
changeset | 84 | sledgehammer_test_theorems params prover minimize_timeout i state | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31409diff
changeset | 85 | fun test_thms filtered thms = | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 86 | case test_thms_fun filtered thms of | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 87 |         (result as {outcome = NONE, ...}) => SOME result
 | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 88 | | _ => NONE | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 89 | |
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 90 |     val {context = ctxt, facts, goal} = Proof.goal state;
 | 
| 31236 | 91 | in | 
| 92 | (* try prove first to check result and get used theorems *) | |
| 31409 
d8537ba165b5
split preparing clauses and writing problemfile;
 immler@in.tum.de parents: 
31236diff
changeset | 93 | (case test_thms_fun NONE name_thms_pairs of | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36393diff
changeset | 94 |       result as {outcome = NONE, pool, internal_thm_names, conjecture_shape,
 | 
| 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36393diff
changeset | 95 | filtered_clauses, ...} => | 
| 31236 | 96 | let | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 97 | val used = internal_thm_names |> Vector.foldr (op ::) [] | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 98 | |> sort_distinct string_ord | 
| 31236 | 99 | val to_use = | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 100 | if length used < length name_thms_pairs then | 
| 36393 
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
 blanchet parents: 
36382diff
changeset | 101 | filter (fn (name1, _) => exists (curry (op =) name1) used) | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 102 | name_thms_pairs | 
| 33305 | 103 | else name_thms_pairs | 
| 36231 
bede2d49ba3b
get rid of "conjecture_pos", which is no longer necessary now that it's Metis's job, not Sledgehammer's, to report inconsistent contexts
 blanchet parents: 
36224diff
changeset | 104 |           val (min_thms, {proof, internal_thm_names, ...}) =
 | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 105 | linear_minimize (test_thms (SOME filtered_clauses)) to_use | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 106 | ([], result) | 
| 36481 
af99c98121d6
make Sledgehammer more friendly if no subgoal is left
 blanchet parents: 
36402diff
changeset | 107 | val m = length min_thms | 
| 32947 
3c19b98a35cd
ATP_Manager.get_prover: canonical argument order;
 wenzelm parents: 
32942diff
changeset | 108 | val _ = priority (cat_lines | 
| 36481 
af99c98121d6
make Sledgehammer more friendly if no subgoal is left
 blanchet parents: 
36402diff
changeset | 109 | ["Minimized: " ^ string_of_int m ^ " theorem" ^ plural_s m] ^ ".") | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 110 | in | 
| 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 111 | (SOME min_thms, | 
| 36402 
1b20805974c7
introduced direct proof reconstruction code, eliminating the need for the "neg_clausify" method;
 blanchet parents: 
36393diff
changeset | 112 | proof_text isar_proof | 
| 36488 
32c92af68ec9
remove Sledgehammer's "sorts" option to annotate variables with sorts in proof;
 blanchet parents: 
36481diff
changeset | 113 | (pool, debug, shrink_factor, ctxt, conjecture_shape) | 
| 36287 
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
 blanchet parents: 
36283diff
changeset | 114 | (K "", proof, internal_thm_names, goal, i) |> fst) | 
| 36223 
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
 blanchet parents: 
36143diff
changeset | 115 | end | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 116 |     | {outcome = SOME TimedOut, ...} =>
 | 
| 36142 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 117 | (NONE, "Timeout: You can increase the time limit using the \"timeout\" \ | 
| 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 118 | \option (e.g., \"timeout = " ^ | 
| 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 119 | string_of_int (10 + msecs div 1000) ^ " s\").") | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 120 |     | {outcome = SOME UnknownError, ...} =>
 | 
| 36142 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 121 | (* Failure sometimes mean timeout, unfortunately. *) | 
| 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 122 | (NONE, "Failure: No proof was found with the current time limit. You \ | 
| 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 123 | \can increase the time limit using the \"timeout\" \ | 
| 
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
 blanchet parents: 
36063diff
changeset | 124 | \option (e.g., \"timeout = " ^ | 
| 36370 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 125 | string_of_int (10 + msecs div 1000) ^ " s\").") | 
| 
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
 blanchet parents: 
36369diff
changeset | 126 |     | {message, ...} => (NONE, "ATP error: " ^ message))
 | 
| 36382 | 127 | handle Sledgehammer_HOL_Clause.TRIVIAL => (SOME [], metis_line i n []) | 
| 128 | | ERROR msg => (NONE, "Error: " ^ msg) | |
| 31236 | 129 | end | 
| 130 | ||
| 35866 
513074557e06
move the Sledgehammer Isar commands together into one file;
 blanchet parents: 
35865diff
changeset | 131 | end; |