| author | blanchet | 
| Sun, 15 Dec 2013 18:01:38 +0100 | |
| changeset 54750 | f50693bab67c | 
| parent 54746 | 6db5fbc02436 | 
| child 54751 | 9b93f9117f8b | 
| permissions | -rw-r--r-- | 
| 49883 | 1 | (* Title: HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML | 
| 2 | Author: Jasmin Blanchette, TU Muenchen | |
| 3 | Author: Steffen Juilf Smolka, TU Muenchen | |
| 4 | ||
| 49914 | 5 | Isar proof reconstruction from ATP proofs. | 
| 49883 | 6 | *) | 
| 7 | ||
| 52374 | 8 | signature SLEDGEHAMMER_RECONSTRUCT = | 
| 49883 | 9 | sig | 
| 54495 | 10 |   type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
 | 
| 53586 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
53052diff
changeset | 11 | type 'a atp_proof = 'a ATP_Proof.atp_proof | 
| 49914 | 12 | type stature = ATP_Problem_Generate.stature | 
| 54495 | 13 | type one_line_params = Sledgehammer_Reconstructor.one_line_params | 
| 49914 | 14 | |
| 15 | type isar_params = | |
| 54505 | 16 | bool * bool * string * string * Time.time option * real * bool * (term, string) atp_step list * | 
| 17 | thm | |
| 49914 | 18 | |
| 49883 | 19 | val isar_proof_text : | 
| 53052 
a0db255af8c5
sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
 wenzelm parents: 
53047diff
changeset | 20 | Proof.context -> bool option -> isar_params -> one_line_params -> string | 
| 49883 | 21 | val proof_text : | 
| 54500 | 22 | Proof.context -> bool option -> (unit -> isar_params) -> int -> one_line_params -> string | 
| 49883 | 23 | end; | 
| 24 | ||
| 54750 | 25 | structure Sledgehammer_Reconstruct (* ### : SLEDGEHAMMER_RECONSTRUCT *) = | 
| 49883 | 26 | struct | 
| 27 | ||
| 28 | open ATP_Util | |
| 49914 | 29 | open ATP_Problem | 
| 49883 | 30 | open ATP_Proof | 
| 31 | open ATP_Problem_Generate | |
| 32 | open ATP_Proof_Reconstruct | |
| 49918 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 blanchet parents: 
49917diff
changeset | 33 | open Sledgehammer_Util | 
| 52555 | 34 | open Sledgehammer_Reconstructor | 
| 50264 
a9ec48b98734
renamed sledgehammer_isar_reconstruct to sledgehammer_proof
 smolkas parents: 
50262diff
changeset | 35 | open Sledgehammer_Proof | 
| 50258 | 36 | open Sledgehammer_Annotate | 
| 52555 | 37 | open Sledgehammer_Print | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52555diff
changeset | 38 | open Sledgehammer_Preplay | 
| 51130 
76d68444cd59
renamed sledgehammer_shrink to sledgehammer_compress
 smolkas parents: 
51129diff
changeset | 39 | open Sledgehammer_Compress | 
| 52592 
8a25b17e3d79
optimize isar-proofs by trying different proof methods
 smolkas parents: 
52556diff
changeset | 40 | open Sledgehammer_Try0 | 
| 52611 
831f7479c74f
minimize dependencies (used facts) of Isar proof steps; remove unreferenced steps
 smolkas parents: 
52592diff
changeset | 41 | open Sledgehammer_Minimize_Isar | 
| 49914 | 42 | |
| 43 | structure String_Redirect = ATP_Proof_Redirect( | |
| 53586 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
53052diff
changeset | 44 | type key = atp_step_name | 
| 49914 | 45 | val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s') | 
| 46 | val string_of = fst) | |
| 47 | ||
| 49883 | 48 | open String_Redirect | 
| 49 | ||
| 54746 | 50 | fun maybe_mk_Trueprop t = t |> fastype_of t = HOLogic.boolT ? HOLogic.mk_Trueprop | 
| 51 | val maybe_dest_Trueprop = perhaps (try HOLogic.dest_Trueprop) | |
| 52 | ||
| 53 | val e_skolemize_rule = "skolemize" | |
| 54 | val vampire_skolemisation_rule = "skolemisation" | |
| 55 | (* TODO: Use "Z3_Proof.string_of_rule" once it is moved to Isabelle *) | |
| 56 | val z3_skolemize_rule = "skolemize" | |
| 57 | val z3_hypothesis_rule = "hypothesis" | |
| 58 | val z3_lemma_rule = "lemma" | |
| 59 | val z3_intro_def_rule = "intro-def" | |
| 60 | val z3_apply_def_rule = "apply-def" | |
| 61 | ||
| 62 | val is_skolemize_rule = | |
| 63 | member (op =) [e_skolemize_rule, vampire_skolemisation_rule, z3_skolemize_rule] | |
| 64 | ||
| 54501 | 65 | fun raw_label_of_num num = (num, 0) | 
| 49914 | 66 | |
| 54501 | 67 | fun label_of_clause [(num, _)] = raw_label_of_num num | 
| 68 | | label_of_clause c = (space_implode "___" (map (fst o raw_label_of_num o fst) c), 0) | |
| 50005 | 69 | |
| 54505 | 70 | fun add_fact_of_dependencies [(_, ss as _ :: _)] = apsnd (union (op =) ss) | 
| 71 | | add_fact_of_dependencies names = apfst (insert (op =) (label_of_clause names)) | |
| 49914 | 72 | |
| 73 | fun replace_one_dependency (old, new) dep = | |
| 74 | if is_same_atp_step dep old then new else [dep] | |
| 51201 | 75 | fun replace_dependencies_in_line p (name, role, t, rule, deps) = | 
| 76 | (name, role, t, rule, fold (union (op =) o replace_one_dependency p) deps []) | |
| 49914 | 77 | |
| 54746 | 78 | fun inline_z3_defs _ [] = [] | 
| 79 | | inline_z3_defs defs ((line as (name, role, t, rule, deps)) :: lines) = | |
| 80 | if rule = z3_intro_def_rule then | |
| 81 | let val def = t |> maybe_dest_Trueprop |> HOLogic.dest_eq |> swap in | |
| 82 | inline_z3_defs (insert (op =) def defs) | |
| 83 | (map (replace_dependencies_in_line (name, [])) lines) | |
| 84 | end | |
| 85 | else if rule = z3_apply_def_rule then | |
| 86 | inline_z3_defs defs (map (replace_dependencies_in_line (name, [])) lines) | |
| 87 | else | |
| 88 | (name, role, Term.subst_atomic defs t, rule, deps) :: inline_z3_defs defs lines | |
| 89 | ||
| 54750 | 90 | fun alist_cons_list eq (k, v) = AList.map_default eq (k, []) (cons v) | 
| 91 | ||
| 92 | (* FIXME: use "prop_of_clause" defined below *) | |
| 93 | fun add_z3_hypotheses [] = I | |
| 94 | | add_z3_hypotheses hyps = | |
| 95 | HOLogic.dest_Trueprop | |
| 96 | #> curry HOLogic.mk_imp (Library.foldr1 HOLogic.mk_conj (map HOLogic.dest_Trueprop hyps)) | |
| 97 | #> HOLogic.mk_Trueprop | |
| 98 | ||
| 99 | fun inline_z3_hypotheses _ _ [] = [] | |
| 100 | | inline_z3_hypotheses hyp_names hyps ((name, role, t, rule, deps) :: lines) = | |
| 101 | if rule = z3_hypothesis_rule then | |
| 102 | inline_z3_hypotheses (name :: hyp_names) (alist_cons_list (op =) (t, name) hyps) lines | |
| 103 | else | |
| 104 | let val deps' = subtract (op =) hyp_names deps in | |
| 105 | if rule = z3_lemma_rule then | |
| 106 | (name, role, t, rule, deps') :: inline_z3_hypotheses hyp_names hyps lines | |
| 107 | else | |
| 108 | let | |
| 109 | val add_hyps = filter_out (null o inter (op =) deps o snd) hyps | |
| 110 | val t' = add_z3_hypotheses (map fst add_hyps) t | |
| 111 | val deps' = subtract (op =) hyp_names deps | |
| 112 | val hyps' = fold (AList.update (op =) o apsnd (insert (op =) name)) add_hyps hyps | |
| 113 | in | |
| 114 | (name, role, t', rule, deps') :: inline_z3_hypotheses hyp_names hyps' lines | |
| 115 | end | |
| 116 | end | |
| 54746 | 117 | |
| 49914 | 118 | (* No "real" literals means only type information (tfree_tcs, clsrel, or | 
| 119 | clsarity). *) | |
| 120 | fun is_only_type_information t = t aconv @{term True}
 | |
| 121 | ||
| 122 | (* Discard facts; consolidate adjacent lines that prove the same formula, since | |
| 123 | they differ only in type information.*) | |
| 54746 | 124 | fun add_line (line as (name as (_, ss), role, t, rule, [])) lines = | 
| 54700 | 125 | (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire) | 
| 126 | internal facts or definitions. *) | |
| 127 | if role = Lemma orelse role = Conjecture orelse role = Negated_Conjecture orelse | |
| 128 | role = Hypothesis then | |
| 54746 | 129 | line :: lines | 
| 54505 | 130 | else if role = Axiom then | 
| 49914 | 131 | (* Facts are not proof lines. *) | 
| 54507 | 132 | lines |> is_only_type_information t ? map (replace_dependencies_in_line (name, [])) | 
| 49914 | 133 | else | 
| 134 | map (replace_dependencies_in_line (name, [])) lines | |
| 54746 | 135 | | add_line line lines = line :: lines | 
| 49914 | 136 | |
| 137 | (* Recursively delete empty lines (type information) from the proof. *) | |
| 51201 | 138 | fun add_nontrivial_line (line as (name, _, t, _, [])) lines = | 
| 54507 | 139 | if is_only_type_information t then delete_dependency name lines else line :: lines | 
| 49914 | 140 | | add_nontrivial_line line lines = line :: lines | 
| 141 | and delete_dependency name lines = | |
| 54700 | 142 | fold_rev add_nontrivial_line (map (replace_dependencies_in_line (name, [])) lines) [] | 
| 49914 | 143 | |
| 54716 
55ed20a29a8c
reverse order in which lines are selected, to ensure that the number of dependencies is accurate
 blanchet parents: 
54715diff
changeset | 144 | fun add_desired_lines res [] = rev res | 
| 
55ed20a29a8c
reverse order in which lines are selected, to ensure that the number of dependencies is accurate
 blanchet parents: 
54715diff
changeset | 145 | | add_desired_lines res ((name as (_, ss), role, t, rule, deps) :: lines) = | 
| 
55ed20a29a8c
reverse order in which lines are selected, to ensure that the number of dependencies is accurate
 blanchet parents: 
54715diff
changeset | 146 | if role <> Plain orelse is_skolemize_rule rule orelse | 
| 
55ed20a29a8c
reverse order in which lines are selected, to ensure that the number of dependencies is accurate
 blanchet parents: 
54715diff
changeset | 147 | (* the last line must be kept *) | 
| 
55ed20a29a8c
reverse order in which lines are selected, to ensure that the number of dependencies is accurate
 blanchet parents: 
54715diff
changeset | 148 | null lines orelse | 
| 54746 | 149 | (not (is_only_type_information t) andalso null (Term.add_tvars t []) | 
| 150 | andalso length deps >= 2 andalso | |
| 54716 
55ed20a29a8c
reverse order in which lines are selected, to ensure that the number of dependencies is accurate
 blanchet parents: 
54715diff
changeset | 151 | (* don't keep next to last line, which usually results in a trivial step *) | 
| 
55ed20a29a8c
reverse order in which lines are selected, to ensure that the number of dependencies is accurate
 blanchet parents: 
54715diff
changeset | 152 | not (can the_single lines)) then | 
| 
55ed20a29a8c
reverse order in which lines are selected, to ensure that the number of dependencies is accurate
 blanchet parents: 
54715diff
changeset | 153 | add_desired_lines ((name, role, t, rule, deps) :: res) lines | 
| 
55ed20a29a8c
reverse order in which lines are selected, to ensure that the number of dependencies is accurate
 blanchet parents: 
54715diff
changeset | 154 | else | 
| 
55ed20a29a8c
reverse order in which lines are selected, to ensure that the number of dependencies is accurate
 blanchet parents: 
54715diff
changeset | 155 | add_desired_lines res (map (replace_dependencies_in_line (name, deps)) lines) | 
| 49914 | 156 | |
| 52454 | 157 | val add_labels_of_proof = | 
| 54700 | 158 | steps_of_proof | 
| 159 | #> fold_isar_steps (byline_of_step #> (fn SOME ((ls, _), _) => union (op =) ls | _ => I)) | |
| 49914 | 160 | |
| 161 | fun kill_useless_labels_in_proof proof = | |
| 162 | let | |
| 51178 | 163 | val used_ls = add_labels_of_proof proof [] | 
| 54700 | 164 | fun kill_label l = if member (op =) used_ls l then l else no_label | 
| 165 | fun kill_assms assms = map (apfst kill_label) assms | |
| 166 | fun kill_step (Prove (qs, xs, l, t, subproofs, by)) = | |
| 167 | Prove (qs, xs, kill_label l, t, map kill_proof subproofs, by) | |
| 168 | | kill_step step = step | |
| 169 | and kill_proof (Proof (fix, assms, steps)) = | |
| 170 | Proof (fix, kill_assms assms, map kill_step steps) | |
| 171 | in | |
| 172 | kill_proof proof | |
| 173 | end | |
| 49914 | 174 | |
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51976diff
changeset | 175 | fun prefix_of_depth n = replicate_string (n + 1) | 
| 49914 | 176 | |
| 54501 | 177 | val assume_prefix = "a" | 
| 178 | val have_prefix = "f" | |
| 179 | ||
| 49914 | 180 | val relabel_proof = | 
| 181 | let | |
| 51179 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 182 | fun fresh_label depth prefix (old as (l, subst, next)) = | 
| 50672 
ab5b8b5c9cbe
added "obtain" to Isar proof construction data structure
 blanchet parents: 
50671diff
changeset | 183 | if l = no_label then | 
| 
ab5b8b5c9cbe
added "obtain" to Isar proof construction data structure
 blanchet parents: 
50671diff
changeset | 184 | old | 
| 
ab5b8b5c9cbe
added "obtain" to Isar proof construction data structure
 blanchet parents: 
50671diff
changeset | 185 | else | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51976diff
changeset | 186 | let val l' = (prefix_of_depth depth prefix, next) in | 
| 51179 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 187 | (l', (l, l') :: subst, next + 1) | 
| 50672 
ab5b8b5c9cbe
added "obtain" to Isar proof construction data structure
 blanchet parents: 
50671diff
changeset | 188 | end | 
| 54700 | 189 | |
| 190 | fun relabel_facts subst = apfst (maps (the_list o AList.lookup (op =) subst)) | |
| 191 | ||
| 192 | fun relabel_assm depth (l, t) (subst, next) = | |
| 54507 | 193 | let val (l, subst, next) = (l, subst, next) |> fresh_label depth assume_prefix in | 
| 51179 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 194 | ((l, t), (subst, next)) | 
| 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 195 | end | 
| 54700 | 196 | |
| 197 | fun relabel_assms subst depth assms = fold_map (relabel_assm depth) assms (subst, 1) ||> fst | |
| 198 | ||
| 199 | fun relabel_steps _ _ _ [] = [] | |
| 200 | | relabel_steps subst depth next (Prove (qs, xs, l, t, sub, by) :: steps) = | |
| 50672 
ab5b8b5c9cbe
added "obtain" to Isar proof construction data structure
 blanchet parents: 
50671diff
changeset | 201 | let | 
| 54507 | 202 | val (l, subst, next) = (l, subst, next) |> fresh_label depth have_prefix | 
| 54700 | 203 | val sub = relabel_proofs subst depth sub | 
| 204 | val by = by |> relabel_byline subst | |
| 205 | in | |
| 206 | Prove (qs, xs, l, t, sub, by) :: relabel_steps subst depth next steps | |
| 207 | end | |
| 208 | | relabel_steps subst depth next (step :: steps) = | |
| 209 | step :: relabel_steps subst depth next steps | |
| 210 | and relabel_proof subst depth (Proof (fix, assms, steps)) = | |
| 211 | let val (assms, subst) = relabel_assms subst depth assms in | |
| 212 | Proof (fix, assms, relabel_steps subst depth 1 steps) | |
| 51179 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 213 | end | 
| 54700 | 214 | and relabel_byline subst byline = apfst (relabel_facts subst) byline | 
| 215 | and relabel_proofs subst depth = map (relabel_proof subst (depth + 1)) | |
| 216 | in | |
| 217 | relabel_proof [] 0 | |
| 218 | end | |
| 49914 | 219 | |
| 50004 
c96e8e40d789
several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
 blanchet parents: 
49994diff
changeset | 220 | val chain_direct_proof = | 
| 
c96e8e40d789
several improvements to Isar proof reconstruction, by Steffen Smolka (step merging in case splits, time measurements, etc.)
 blanchet parents: 
49994diff
changeset | 221 | let | 
| 54700 | 222 | fun chain_qs_lfs NONE lfs = ([], lfs) | 
| 223 | | chain_qs_lfs (SOME l0) lfs = | |
| 224 | if member (op =) lfs l0 then ([Then], lfs |> remove (op =) l0) else ([], lfs) | |
| 225 | fun chain_step lbl (Prove (qs, xs, l, t, subproofs, ((lfs, gfs), method))) = | |
| 226 | let val (qs', lfs) = chain_qs_lfs lbl lfs in | |
| 227 | Prove (qs' @ qs, xs, l, t, chain_proofs subproofs, ((lfs, gfs), method)) | |
| 228 | end | |
| 50672 
ab5b8b5c9cbe
added "obtain" to Isar proof construction data structure
 blanchet parents: 
50671diff
changeset | 229 | | chain_step _ step = step | 
| 51179 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 230 | and chain_steps _ [] = [] | 
| 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 231 | | chain_steps (prev as SOME _) (i :: is) = | 
| 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 232 | chain_step prev i :: chain_steps (label_of_step i) is | 
| 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 233 | | chain_steps _ (i :: is) = i :: chain_steps (label_of_step i) is | 
| 54700 | 234 | and chain_proof (Proof (fix, assms, steps)) = | 
| 235 | Proof (fix, assms, chain_steps (try (List.last #> fst) assms) steps) | |
| 51179 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 236 | and chain_proofs proofs = map (chain_proof) proofs | 
| 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 237 | in chain_proof end | 
| 49883 | 238 | |
| 49914 | 239 | type isar_params = | 
| 54505 | 240 | bool * bool * string * string * Time.time option * real * bool * (term, string) atp_step list * | 
| 241 | thm | |
| 49914 | 242 | |
| 53052 
a0db255af8c5
sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
 wenzelm parents: 
53047diff
changeset | 243 | fun isar_proof_text ctxt isar_proofs | 
| 54500 | 244 | (debug, verbose, metis_type_enc, metis_lam_trans, preplay_timeout, isar_compress, | 
| 54505 | 245 | isar_try0, atp_proof, goal) | 
| 49918 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 blanchet parents: 
49917diff
changeset | 246 | (one_line_params as (_, _, _, _, subgoal, subgoal_count)) = | 
| 49883 | 247 | let | 
| 52196 
2281f33e8da6
redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
 blanchet parents: 
52150diff
changeset | 248 | val (params, hyp_ts, concl_t) = strip_subgoal goal subgoal ctxt | 
| 
2281f33e8da6
redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
 blanchet parents: 
52150diff
changeset | 249 | val (_, ctxt) = | 
| 
2281f33e8da6
redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
 blanchet parents: 
52150diff
changeset | 250 | params | 
| 
2281f33e8da6
redid rac7830871177 to avoid duplicate fixed variable (e.g. lemma "P (a::nat)" proof - have "!!a::int. Q a" sledgehammer [e])
 blanchet parents: 
52150diff
changeset | 251 | |> map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) | 
| 54507 | 252 | |> (fn fixes => ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes) | 
| 53052 
a0db255af8c5
sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
 wenzelm parents: 
53047diff
changeset | 253 | val one_line_proof = one_line_proof_text 0 one_line_params | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52555diff
changeset | 254 | val do_preplay = preplay_timeout <> SOME Time.zeroTime | 
| 49883 | 255 | |
| 54700 | 256 | fun get_role keep_role ((num, _), role, t, _, _) = | 
| 257 | if keep_role role then SOME (raw_label_of_num num, t) else NONE | |
| 258 | ||
| 49883 | 259 | fun isar_proof_of () = | 
| 260 | let | |
| 261 | val atp_proof = | |
| 262 | atp_proof | |
| 54746 | 263 | |> inline_z3_defs [] | 
| 54750 | 264 | |> inline_z3_hypotheses [] [] | 
| 54505 | 265 | |> rpair [] |-> fold_rev add_line | 
| 49883 | 266 | |> rpair [] |-> fold_rev add_nontrivial_line | 
| 54716 
55ed20a29a8c
reverse order in which lines are selected, to ensure that the number of dependencies is accurate
 blanchet parents: 
54715diff
changeset | 267 | |> add_desired_lines [] | 
| 54700 | 268 | |
| 54535 
59737a43e044
support Negated_Conjecture as a TPTP role as well (e.g. for SMT proofs)
 blanchet parents: 
54507diff
changeset | 269 | val conjs = | 
| 54700 | 270 | map_filter (fn (name, role, _, _, _) => | 
| 271 | if member (op =) [Conjecture, Negated_Conjecture] role then SOME name else NONE) | |
| 272 | atp_proof | |
| 273 | val assms = map_filter (get_role (curry (op =) Hypothesis)) atp_proof | |
| 274 | val lems = | |
| 275 | map_filter (get_role (curry (op =) Lemma)) atp_proof | |
| 54712 | 276 | |> map (fn (l, t) => Prove ([], [], l, maybe_mk_Trueprop t, [], (([], []), ArithM))) | 
| 54700 | 277 | |
| 51212 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 278 | val bot = atp_proof |> List.last |> #1 | 
| 54700 | 279 | |
| 51145 | 280 | val refute_graph = | 
| 51212 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 281 | atp_proof | 
| 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 282 | |> map (fn (name, _, _, _, from) => (from, name)) | 
| 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 283 | |> make_refute_graph bot | 
| 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 284 | |> fold (Atom_Graph.default_node o rpair ()) conjs | 
| 54700 | 285 | |
| 51145 | 286 | val axioms = axioms_of_refute_graph refute_graph conjs | 
| 54700 | 287 | |
| 51145 | 288 | val tainted = tainted_atoms_of_refute_graph refute_graph conjs | 
| 51156 | 289 | val is_clause_tainted = exists (member (op =) tainted) | 
| 50676 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 290 | val steps = | 
| 49883 | 291 | Symtab.empty | 
| 51201 | 292 | |> fold (fn (name as (s, _), role, t, rule, _) => | 
| 50676 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 293 | Symtab.update_new (s, (rule, | 
| 51156 | 294 | t |> (if is_clause_tainted [name] then | 
| 54712 | 295 | role <> Conjecture ? (maybe_dest_Trueprop #> s_not) | 
| 50676 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 296 | #> fold exists_of (map Var (Term.add_vars t [])) | 
| 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 297 | else | 
| 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 298 | I)))) | 
| 49883 | 299 | atp_proof | 
| 54700 | 300 | |
| 51148 
2246a2e17f92
tuning -- refactoring in preparation for handling skolemization of conjecture
 blanchet parents: 
51147diff
changeset | 301 | fun is_clause_skolemize_rule [(s, _)] = | 
| 54505 | 302 | Option.map (is_skolemize_rule o fst) (Symtab.lookup steps s) = SOME true | 
| 50676 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 303 | | is_clause_skolemize_rule _ = false | 
| 54700 | 304 | |
| 54712 | 305 | (* The assumptions and conjecture are "prop"s; the other formulas are "bool"s (for ATPs) or | 
| 306 | "prop"s (for Z3). TODO: Always use "prop". *) | |
| 54505 | 307 | fun prop_of_clause [(num, _)] = | 
| 308 | Symtab.lookup steps num |> the |> snd |> maybe_mk_Trueprop |> close_form | |
| 50016 | 309 | | prop_of_clause names = | 
| 50676 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 310 | let | 
| 54712 | 311 | val lits = | 
| 312 | map (maybe_dest_Trueprop o snd) (map_filter (Symtab.lookup steps o fst) names) | |
| 50676 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 313 | in | 
| 50018 
4ea26c74d7ea
use implications rather than disjunctions to improve readability
 blanchet parents: 
50017diff
changeset | 314 | case List.partition (can HOLogic.dest_not) lits of | 
| 
4ea26c74d7ea
use implications rather than disjunctions to improve readability
 blanchet parents: 
50017diff
changeset | 315 | (negs as _ :: _, pos as _ :: _) => | 
| 54507 | 316 | s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos) | 
| 50018 
4ea26c74d7ea
use implications rather than disjunctions to improve readability
 blanchet parents: 
50017diff
changeset | 317 |               | _ => fold (curry s_disj) lits @{term False}
 | 
| 
4ea26c74d7ea
use implications rather than disjunctions to improve readability
 blanchet parents: 
50017diff
changeset | 318 | end | 
| 50016 | 319 | |> HOLogic.mk_Trueprop |> close_form | 
| 54700 | 320 | |
| 321 | val is_fixed = Variable.is_declared ctxt orf can Name.dest_skolem | |
| 322 | fun skolems_of t = Term.add_frees t [] |> filter_out (is_fixed o fst) |> rev | |
| 323 | ||
| 324 | fun maybe_show outer c = | |
| 325 | (outer andalso length c = 1 andalso subset (op =) (c, conjs)) ? cons Show | |
| 326 | ||
| 327 | fun isar_steps outer predecessor accum [] = | |
| 328 | accum | |
| 329 | |> (if tainted = [] then | |
| 330 | cons (Prove (if outer then [Show] else [], [], no_label, concl_t, [], | |
| 331 | (([the predecessor], []), MetisM))) | |
| 332 | else | |
| 333 | I) | |
| 334 | |> rev | |
| 335 | | isar_steps outer _ accum (Have (gamma, c) :: infs) = | |
| 336 | let | |
| 337 | val l = label_of_clause c | |
| 338 | val t = prop_of_clause c | |
| 339 | val by = (fold add_fact_of_dependencies gamma no_facts, MetisM) | |
| 340 | fun prove sub by = Prove (maybe_show outer c [], [], l, t, sub, by) | |
| 341 | fun do_rest l step = isar_steps outer (SOME l) (step :: accum) infs | |
| 342 | in | |
| 343 | if is_clause_tainted c then | |
| 54712 | 344 | (case gamma of | 
| 54700 | 345 | [g] => | 
| 346 | if is_clause_skolemize_rule g andalso is_clause_tainted g then | |
| 347 | let | |
| 348 | val subproof = Proof (skolems_of (prop_of_clause g), [], rev accum) | |
| 349 | in | |
| 350 | isar_steps outer (SOME l) [prove [subproof] (no_facts, MetisM)] [] | |
| 351 | end | |
| 51148 
2246a2e17f92
tuning -- refactoring in preparation for handling skolemization of conjecture
 blanchet parents: 
51147diff
changeset | 352 | else | 
| 54700 | 353 | do_rest l (prove [] by) | 
| 54712 | 354 | | _ => do_rest l (prove [] by)) | 
| 54700 | 355 | else | 
| 356 | if is_clause_skolemize_rule c then | |
| 357 | do_rest l (Prove ([], skolems_of t, l, t, [], by)) | |
| 358 | else | |
| 359 | do_rest l (prove [] by) | |
| 360 | end | |
| 361 | | isar_steps outer predecessor accum (Cases cases :: infs) = | |
| 362 | let | |
| 363 | fun isar_case (c, infs) = | |
| 364 | isar_proof false [] [(label_of_clause c, prop_of_clause c)] [] infs | |
| 365 | val c = succedent_of_cases cases | |
| 366 | val l = label_of_clause c | |
| 367 | val t = prop_of_clause c | |
| 368 | val step = | |
| 369 | Prove (maybe_show outer c [], [], l, t, map isar_case cases, | |
| 370 | ((the_list predecessor, []), MetisM)) | |
| 371 | in | |
| 372 | isar_steps outer (SOME l) (step :: accum) infs | |
| 373 | end | |
| 374 | and isar_proof outer fix assms lems infs = | |
| 375 | Proof (fix, assms, lems @ isar_steps outer NONE [] infs) | |
| 376 | ||
| 377 | val isar_proof_of_direct_proof = isar_proof true params assms lems | |
| 51179 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 378 | |
| 52592 
8a25b17e3d79
optimize isar-proofs by trying different proof methods
 smolkas parents: 
52556diff
changeset | 379 | (* 60 seconds seems like a good interpreation of "no timeout" *) | 
| 
8a25b17e3d79
optimize isar-proofs by trying different proof methods
 smolkas parents: 
52556diff
changeset | 380 | val preplay_timeout = preplay_timeout |> the_default (seconds 60.0) | 
| 
8a25b17e3d79
optimize isar-proofs by trying different proof methods
 smolkas parents: 
52556diff
changeset | 381 | |
| 51741 | 382 | val clean_up_labels_in_proof = | 
| 51165 | 383 | chain_direct_proof | 
| 384 | #> kill_useless_labels_in_proof | |
| 385 | #> relabel_proof | |
| 53761 
4d34e267fba9
use configuration mechanism for low-level tracing
 blanchet parents: 
53586diff
changeset | 386 |         val (preplay_interface as {overall_preplay_stats, ...}, isar_proof) =
 | 
| 51145 | 387 | refute_graph | 
| 54699 | 388 | (* | 
| 54712 | 389 | |> tap (tracing o prefix "REFUTE GRAPH: " o string_of_refute_graph) | 
| 54699 | 390 | *) | 
| 51031 
63d71b247323
more robustness in Isar proof reconstruction (cf. bug report by Ondrej)
 blanchet parents: 
51026diff
changeset | 391 | |> redirect_graph axioms tainted bot | 
| 54699 | 392 | (* | 
| 54712 | 393 | |> tap (tracing o prefix "DIRECT PROOF: " o string_of_direct_proof) | 
| 54699 | 394 | *) | 
| 51179 
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
 smolkas parents: 
51178diff
changeset | 395 | |> isar_proof_of_direct_proof | 
| 54712 | 396 | |> postprocess_remove_unreferenced_steps I | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52555diff
changeset | 397 | |> relabel_proof_canonically | 
| 54500 | 398 | |> `(proof_preplay_interface debug ctxt metis_type_enc metis_lam_trans do_preplay | 
| 53761 
4d34e267fba9
use configuration mechanism for low-level tracing
 blanchet parents: 
53586diff
changeset | 399 | preplay_timeout) | 
| 52626 | 400 | val ((preplay_time, preplay_fail), isar_proof) = | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52555diff
changeset | 401 | isar_proof | 
| 54501 | 402 | |> compress_proof (if isar_proofs = SOME true then isar_compress else 1000.0) | 
| 53762 | 403 | preplay_interface | 
| 52632 
23393c31c7fe
added sledgehammer parameters isar_try0 and isar_minimize and their negative aliases
 smolkas parents: 
52626diff
changeset | 404 | |> isar_try0 ? try0 preplay_timeout preplay_interface | 
| 54712 | 405 | |> postprocess_remove_unreferenced_steps (isar_try0 ? min_deps_of_step preplay_interface) | 
| 52626 | 406 | |> `overall_preplay_stats | 
| 407 | ||> clean_up_labels_in_proof | |
| 54500 | 408 | val isar_text = | 
| 409 | string_of_proof ctxt metis_type_enc metis_lam_trans subgoal subgoal_count isar_proof | |
| 49883 | 410 | in | 
| 49918 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 blanchet parents: 
49917diff
changeset | 411 | case isar_text of | 
| 49883 | 412 | "" => | 
| 51190 
2654b3965c8d
made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
 blanchet parents: 
51187diff
changeset | 413 | if isar_proofs = SOME true then | 
| 50671 | 414 | "\nNo structured proof available (proof too simple)." | 
| 49883 | 415 | else | 
| 416 | "" | |
| 417 | | _ => | |
| 50670 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
50557diff
changeset | 418 | let | 
| 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
50557diff
changeset | 419 | val msg = | 
| 51203 | 420 | (if verbose then | 
| 421 | let | |
| 52592 
8a25b17e3d79
optimize isar-proofs by trying different proof methods
 smolkas parents: 
52556diff
changeset | 422 | val num_steps = add_proof_steps (steps_of_proof isar_proof) 0 | 
| 51203 | 423 | in [string_of_int num_steps ^ " step" ^ plural_s num_steps] end | 
| 424 | else | |
| 425 | []) @ | |
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52555diff
changeset | 426 | (if do_preplay then | 
| 50924 | 427 | [(if preplay_fail then "may fail, " else "") ^ | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52555diff
changeset | 428 | string_of_preplay_time preplay_time] | 
| 50670 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
50557diff
changeset | 429 | else | 
| 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
50557diff
changeset | 430 | []) | 
| 50277 | 431 | in | 
| 54507 | 432 |             "\n\nStructured proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^
 | 
| 433 | Active.sendback_markup [Markup.padding_command] isar_text | |
| 50277 | 434 | end | 
| 49883 | 435 | end | 
| 436 | val isar_proof = | |
| 437 | if debug then | |
| 438 | isar_proof_of () | |
| 439 | else case try isar_proof_of () of | |
| 440 | SOME s => s | |
| 51190 
2654b3965c8d
made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
 blanchet parents: 
51187diff
changeset | 441 | | NONE => if isar_proofs = SOME true then | 
| 49883 | 442 | "\nWarning: The Isar proof construction failed." | 
| 443 | else | |
| 444 | "" | |
| 445 | in one_line_proof ^ isar_proof end | |
| 446 | ||
| 51187 
c344cf148e8f
avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
 blanchet parents: 
51179diff
changeset | 447 | fun isar_proof_would_be_a_good_idea preplay = | 
| 
c344cf148e8f
avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
 blanchet parents: 
51179diff
changeset | 448 | case preplay of | 
| 
c344cf148e8f
avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
 blanchet parents: 
51179diff
changeset | 449 | Played (reconstr, _) => reconstr = SMT | 
| 54093 | 450 | | Trust_Playable _ => false | 
| 51187 
c344cf148e8f
avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
 blanchet parents: 
51179diff
changeset | 451 | | Failed_to_Play _ => true | 
| 
c344cf148e8f
avoid using "smt" for minimization -- better use the prover itself, since then Sledgehammer gets to try metis again and gives the opportunity to output an Isar proof -- and show Isar proof as fallback for SMT proofs
 blanchet parents: 
51179diff
changeset | 452 | |
| 53052 
a0db255af8c5
sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
 wenzelm parents: 
53047diff
changeset | 453 | fun proof_text ctxt isar_proofs isar_params num_chained | 
| 49883 | 454 | (one_line_params as (preplay, _, _, _, _, _)) = | 
| 51190 
2654b3965c8d
made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
 blanchet parents: 
51187diff
changeset | 455 | (if isar_proofs = SOME true orelse | 
| 
2654b3965c8d
made "isar_proofs" a 3-way option, to provide a way to totally disable isar_proofs if desired
 blanchet parents: 
51187diff
changeset | 456 | (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea preplay) then | 
| 54500 | 457 | isar_proof_text ctxt isar_proofs (isar_params ()) | 
| 49883 | 458 | else | 
| 53052 
a0db255af8c5
sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
 wenzelm parents: 
53047diff
changeset | 459 | one_line_proof_text num_chained) one_line_params | 
| 49883 | 460 | |
| 461 | end; |