| author | nipkow | 
| Mon, 19 Feb 2018 13:56:16 +0100 | |
| changeset 67656 | 59feb83c6ab9 | 
| parent 67560 | 0fa87bd86566 | 
| child 68668 | c9570658e8f1 | 
| permissions | -rw-r--r-- | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 1 | (* Title: HOL/Tools/Sledgehammer/sledgehammer_isar.ML | 
| 49883 | 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 | ||
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 8 | signature SLEDGEHAMMER_ISAR = | 
| 49883 | 9 | sig | 
| 54771 
85879aa61334
move some Z3 specifics out (and into private repository with the rest of the Z3-specific code)
 blanchet parents: 
54770diff
changeset | 10 | type atp_step_name = ATP_Proof.atp_step_name | 
| 54495 | 11 |   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 | 12 | type 'a atp_proof = 'a ATP_Proof.atp_proof | 
| 49914 | 13 | type stature = ATP_Problem_Generate.stature | 
| 55287 | 14 | type one_line_params = Sledgehammer_Proof_Methods.one_line_params | 
| 49914 | 15 | |
| 55222 | 16 | val trace : bool Config.T | 
| 17 | ||
| 49914 | 18 | type isar_params = | 
| 57783 | 19 | bool * (string option * string option) * Time.time * real option * bool * bool | 
| 55288 
1a4358d14ce2
added 'smt' option to control generation of 'by smt' proofs
 blanchet parents: 
55287diff
changeset | 20 | * (term, string) atp_step list * thm | 
| 49914 | 21 | |
| 55296 | 22 | val proof_text : Proof.context -> bool -> bool option -> bool option -> (unit -> isar_params) -> | 
| 23 | int -> one_line_params -> string | |
| 49883 | 24 | end; | 
| 25 | ||
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 26 | structure Sledgehammer_Isar : SLEDGEHAMMER_ISAR = | 
| 49883 | 27 | struct | 
| 28 | ||
| 29 | open ATP_Util | |
| 49914 | 30 | open ATP_Problem | 
| 49883 | 31 | open ATP_Proof | 
| 32 | open ATP_Proof_Reconstruct | |
| 58516 
1edba0152491
insert lemmas closer to where they are needed, both for esthetics and (primarily) for correctness in case the lemma refers to a skolem
 blanchet parents: 
58514diff
changeset | 33 | open ATP_Waldmeister | 
| 49918 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 blanchet parents: 
49917diff
changeset | 34 | open Sledgehammer_Util | 
| 55287 | 35 | open Sledgehammer_Proof_Methods | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 36 | open Sledgehammer_Isar_Proof | 
| 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 37 | open Sledgehammer_Isar_Preplay | 
| 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 38 | open Sledgehammer_Isar_Compress | 
| 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 39 | open Sledgehammer_Isar_Minimize | 
| 49914 | 40 | |
| 41 | structure String_Redirect = ATP_Proof_Redirect( | |
| 53586 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
53052diff
changeset | 42 | type key = atp_step_name | 
| 49914 | 43 | val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s') | 
| 44 | val string_of = fst) | |
| 45 | ||
| 49883 | 46 | open String_Redirect | 
| 47 | ||
| 55222 | 48 | val trace = Attrib.setup_config_bool @{binding sledgehammer_isar_trace} (K false)
 | 
| 49 | ||
| 57785 | 50 | val e_definition_rule = "definition" | 
| 57654 
f89c0749533d
'shift_quantors' is not an E skolemization rule (cf. 3ab503b04bdb)
 blanchet parents: 
57288diff
changeset | 51 | val e_skolemize_rule = "skolemize" | 
| 57759 | 52 | val leo2_extcnf_forall_neg_rule = "extcnf_forall_neg" | 
| 59577 | 53 | val pirate_datatype_rule = "DT" | 
| 57709 | 54 | val satallax_skolemize_rule = "tab_ex" | 
| 54746 | 55 | val vampire_skolemisation_rule = "skolemisation" | 
| 57761 | 56 | val veriT_la_generic_rule = "la_generic" | 
| 57708 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 57 | val veriT_simp_arith_rule = "simp_arith" | 
| 57761 | 58 | val veriT_tmp_ite_elim_rule = "tmp_ite_elim" | 
| 59 | val veriT_tmp_skolemize_rule = "tmp_skolemize" | |
| 58061 | 60 | val z3_skolemize_rule = Z3_Proof.string_of_rule Z3_Proof.Skolemize | 
| 61 | val z3_th_lemma_rule_prefix = Z3_Proof.string_of_rule (Z3_Proof.Th_Lemma "") | |
| 57702 
dfc834e39c1f
Skolemization support for leo-II and Zipperposition.
 fleury parents: 
57699diff
changeset | 62 | val zipperposition_cnf_rule = "cnf" | 
| 54746 | 63 | |
| 54772 | 64 | val skolemize_rules = | 
| 57785 | 65 | [e_definition_rule, e_skolemize_rule, leo2_extcnf_forall_neg_rule, satallax_skolemize_rule, | 
| 66 | spass_skolemize_rule, vampire_skolemisation_rule, veriT_tmp_ite_elim_rule, | |
| 58142 | 67 | veriT_tmp_skolemize_rule, waldmeister_skolemize_rule, z3_skolemize_rule, zipperposition_cnf_rule] | 
| 54746 | 68 | |
| 58653 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58517diff
changeset | 69 | fun is_ext_rule rule = (rule = leo2_extcnf_equal_neg_rule) | 
| 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58517diff
changeset | 70 | val is_maybe_ext_rule = is_ext_rule orf String.isPrefix satallax_tab_rule_prefix | 
| 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58517diff
changeset | 71 | |
| 54769 
3d6ac2f68bf3
correcly recognize E skolemization steps that are wrapped in a 'shift_quantors' inference
 blanchet parents: 
54768diff
changeset | 72 | val is_skolemize_rule = member (op =) skolemize_rules | 
| 57761 | 73 | fun is_arith_rule rule = | 
| 57762 | 74 | String.isPrefix z3_th_lemma_rule_prefix rule orelse rule = veriT_simp_arith_rule orelse | 
| 57761 | 75 | rule = veriT_la_generic_rule | 
| 59577 | 76 | val is_datatype_rule = String.isPrefix pirate_datatype_rule | 
| 54755 | 77 | |
| 54501 | 78 | fun raw_label_of_num num = (num, 0) | 
| 49914 | 79 | |
| 54501 | 80 | fun label_of_clause [(num, _)] = raw_label_of_num num | 
| 81 | | label_of_clause c = (space_implode "___" (map (fst o raw_label_of_num o fst) c), 0) | |
| 50005 | 82 | |
| 58653 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58517diff
changeset | 83 | fun add_global_fact ss = apsnd (union (op =) ss) | 
| 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58517diff
changeset | 84 | |
| 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58517diff
changeset | 85 | fun add_fact_of_dependency [(_, ss as _ :: _)] = add_global_fact ss | 
| 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58517diff
changeset | 86 | | add_fact_of_dependency names = apfst (insert (op =) (label_of_clause names)) | 
| 49914 | 87 | |
| 54799 | 88 | fun add_line_pass1 (line as (name, role, t, rule, [])) lines = | 
| 54770 | 89 | (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire) internal facts or | 
| 90 | definitions. *) | |
| 57708 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 91 | if role = Conjecture orelse role = Negated_Conjecture then | 
| 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 92 | line :: lines | 
| 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 93 |     else if t aconv @{prop True} then
 | 
| 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 94 | map (replace_dependencies_in_line (name, [])) lines | 
| 57714 
4856a7b8b9c3
Changing the role of rule "tmp_ite_elim" of the SMT solver veriT to Lemma.
 fleury parents: 
57713diff
changeset | 95 | else if role = Lemma orelse role = Hypothesis orelse is_arith_rule rule then | 
| 57708 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 96 | line :: lines | 
| 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 97 | else if role = Axiom then | 
| 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 98 | lines (* axioms (facts) need no proof lines *) | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 99 | else | 
| 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 100 | map (replace_dependencies_in_line (name, [])) lines | 
| 54755 | 101 | | add_line_pass1 line lines = line :: lines | 
| 49914 | 102 | |
| 57771 | 103 | fun add_lines_pass2 res [] = rev res | 
| 104 | | add_lines_pass2 res ((line as (name, role, t, rule, deps)) :: lines) = | |
| 55184 
6e2295db4cf8
keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
 blanchet parents: 
55183diff
changeset | 105 | let | 
| 57791 | 106 | fun normalize role = | 
| 107 | role = Conjecture ? (HOLogic.dest_Trueprop #> s_not #> HOLogic.mk_Trueprop) | |
| 108 | ||
| 57771 | 109 | val norm_t = normalize role t | 
| 110 | val is_duplicate = | |
| 111 | exists (fn (prev_name, prev_role, prev_t, _, _) => | |
| 58514 
1fc93ea5136b
eliminate duplicate hypotheses (which can arise due to (un)clausification)
 blanchet parents: 
58506diff
changeset | 112 | (prev_role = Hypothesis andalso prev_t aconv t) orelse | 
| 
1fc93ea5136b
eliminate duplicate hypotheses (which can arise due to (un)clausification)
 blanchet parents: 
58506diff
changeset | 113 | (member (op =) deps prev_name andalso | 
| 
1fc93ea5136b
eliminate duplicate hypotheses (which can arise due to (un)clausification)
 blanchet parents: 
58506diff
changeset | 114 | Term.aconv_untyped (normalize prev_role prev_t, norm_t))) | 
| 57771 | 115 | res | 
| 57770 
6c4ab6f0a6fc
normalize conjectures vs. negated conjectures when comparing terms
 blanchet parents: 
57769diff
changeset | 116 | |
| 57787 | 117 |       fun looks_boring () = t aconv @{prop False} orelse length deps < 2
 | 
| 55184 
6e2295db4cf8
keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
 blanchet parents: 
55183diff
changeset | 118 | |
| 
6e2295db4cf8
keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
 blanchet parents: 
55183diff
changeset | 119 | fun is_skolemizing_line (_, _, _, rule', deps') = | 
| 
6e2295db4cf8
keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
 blanchet parents: 
55183diff
changeset | 120 | is_skolemize_rule rule' andalso member (op =) deps' name | 
| 57770 
6c4ab6f0a6fc
normalize conjectures vs. negated conjectures when comparing terms
 blanchet parents: 
57769diff
changeset | 121 | |
| 55184 
6e2295db4cf8
keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
 blanchet parents: 
55183diff
changeset | 122 | fun is_before_skolemize_rule () = exists is_skolemizing_line lines | 
| 
6e2295db4cf8
keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
 blanchet parents: 
55183diff
changeset | 123 | in | 
| 57771 | 124 | if is_duplicate orelse | 
| 57770 
6c4ab6f0a6fc
normalize conjectures vs. negated conjectures when comparing terms
 blanchet parents: 
57769diff
changeset | 125 | (role = Plain andalso not (is_skolemize_rule rule) andalso | 
| 58653 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58517diff
changeset | 126 | not (is_ext_rule rule) andalso not (is_arith_rule rule) andalso | 
| 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58517diff
changeset | 127 | not (is_datatype_rule rule) andalso not (null lines) andalso looks_boring () andalso | 
| 
4b44c227c0e0
improved handling of extensionality in Isar proofs generated from LEO-II and Satallax
 blanchet parents: 
58517diff
changeset | 128 | not (is_before_skolemize_rule ())) then | 
| 57771 | 129 | add_lines_pass2 res (map (replace_dependencies_in_line (name, deps)) lines) | 
| 55184 
6e2295db4cf8
keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
 blanchet parents: 
55183diff
changeset | 130 | else | 
| 57771 | 131 | add_lines_pass2 (line :: res) lines | 
| 55184 
6e2295db4cf8
keep formula right before skolemization, because the universal variables might be different (or differently ordered) as in the original axiom or negated conjecture from which it was skolemized
 blanchet parents: 
55183diff
changeset | 132 | end | 
| 49914 | 133 | |
| 134 | type isar_params = | |
| 57783 | 135 | bool * (string option * string option) * Time.time * real option * bool * bool | 
| 55288 
1a4358d14ce2
added 'smt' option to control generation of 'by smt' proofs
 blanchet parents: 
55287diff
changeset | 136 | * (term, string) atp_step list * thm | 
| 49914 | 137 | |
| 56852 
b38c5b9cf590
added 'satx' to Sledgehammer's portfolio (cf. 'isar_try0')
 blanchet parents: 
56130diff
changeset | 138 | val basic_systematic_methods = [Metis_Method (NONE, NONE), Meson_Method, Blast_Method, SATx_Method] | 
| 58092 
4ae52c60603a
renamed 'skolem' to 'moura' (to suggest Z3-style skolemization); reintroduced 'fastforce' to the mix of tested proof methods
 blanchet parents: 
58083diff
changeset | 139 | val basic_simp_based_methods = [Auto_Method, Simp_Method, Fastforce_Method, Force_Method] | 
| 55323 
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
 blanchet parents: 
55311diff
changeset | 140 | val basic_arith_methods = [Linarith_Method, Presburger_Method, Algebra_Method] | 
| 55311 | 141 | |
| 57744 
a68b8db8691d
added appropriate method for skolemization of Z3 steps to the mix
 blanchet parents: 
57741diff
changeset | 142 | val arith_methods = basic_arith_methods @ basic_simp_based_methods @ basic_systematic_methods | 
| 55311 | 143 | val datatype_methods = [Simp_Method, Simp_Size_Method] | 
| 57744 
a68b8db8691d
added appropriate method for skolemization of Z3 steps to the mix
 blanchet parents: 
57741diff
changeset | 144 | val systematic_methods = | 
| 
a68b8db8691d
added appropriate method for skolemization of Z3 steps to the mix
 blanchet parents: 
57741diff
changeset | 145 | basic_systematic_methods @ basic_arith_methods @ basic_simp_based_methods @ | 
| 57699 | 146 | [Metis_Method (SOME full_typesN, NONE), Metis_Method (SOME no_typesN, NONE)] | 
| 57744 
a68b8db8691d
added appropriate method for skolemization of Z3 steps to the mix
 blanchet parents: 
57741diff
changeset | 147 | val rewrite_methods = basic_simp_based_methods @ basic_systematic_methods @ basic_arith_methods | 
| 58517 | 148 | val skolem_methods = Moura_Method :: systematic_methods | 
| 54766 
6ac273f176cd
store alternative proof methods in Isar data structure
 blanchet parents: 
54765diff
changeset | 149 | |
| 60612 
79d71bfea310
removed chained facts from preplaying -- and careful about extra chained facts when removing 'proof -' and 'qed' from one-line Isar proofs
 blanchet parents: 
59577diff
changeset | 150 | fun isar_proof_text ctxt debug num_chained isar_proofs smt_proofs isar_params | 
| 57739 | 151 | (one_line_params as ((used_facts, (_, one_line_play)), banner, subgoal, subgoal_count)) = | 
| 49883 | 152 | let | 
| 58843 | 153 | val _ = if debug then writeln "Constructing Isar proof..." else () | 
| 56097 | 154 | |
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 155 | fun generate_proof_text () = | 
| 49883 | 156 | let | 
| 57791 | 157 | val (verbose, alt_metis_args, preplay_timeout, compress, try0, minimize, atp_proof0, goal) = | 
| 57245 | 158 | isar_params () | 
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 159 | in | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 160 | if null atp_proof0 then | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 161 | one_line_proof_text ctxt 0 one_line_params | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 162 | else | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 163 | let | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 164 | val systematic_methods' = insert (op =) (Metis_Method alt_metis_args) systematic_methods | 
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 165 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 166 | fun massage_methods (meths as meth :: _) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 167 | if not try0 then [meth] | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 168 | else if smt_proofs = SOME true then SMT_Method :: meths | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 169 | else meths | 
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 170 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 171 | val (params, _, concl_t) = strip_subgoal goal subgoal ctxt | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 172 | val fixes = map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 173 | val ctxt = ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes |> snd | 
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 174 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 175 | val do_preplay = preplay_timeout <> Time.zeroTime | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 176 | val compress = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 177 | (case compress of | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 178 | NONE => if isar_proofs = NONE andalso do_preplay then 1000.0 else 10.0 | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 179 | | SOME n => n) | 
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 180 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 181 | fun is_fixed ctxt = Variable.is_declared ctxt orf Name.is_skolem | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 182 | fun skolems_of ctxt t = Term.add_frees t [] |> filter_out (is_fixed ctxt o fst) |> rev | 
| 54700 | 183 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 184 | fun get_role keep_role ((num, _), role, t, rule, _) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 185 | if keep_role role then SOME ((raw_label_of_num num, t), rule) else NONE | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 186 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 187 | val atp_proof = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 188 | fold_rev add_line_pass1 atp_proof0 [] | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 189 | |> add_lines_pass2 [] | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 190 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 191 | val conjs = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 192 | map_filter (fn (name, role, _, _, _) => | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 193 | if member (op =) [Conjecture, Negated_Conjecture] role then SOME name else NONE) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 194 | atp_proof | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 195 | val assms = map_filter (Option.map fst o get_role (curry (op =) Hypothesis)) atp_proof | 
| 57288 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 196 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 197 | fun add_lemma ((l, t), rule) ctxt = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 198 | let | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 199 | val (skos, meths) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 200 | (if is_skolemize_rule rule then (skolems_of ctxt t, skolem_methods) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 201 | else if is_arith_rule rule then ([], arith_methods) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 202 | else ([], rewrite_methods)) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 203 | ||> massage_methods | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 204 | in | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 205 | (Prove ([], skos, l, t, [], ([], []), meths, ""), | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 206 | ctxt |> not (null skos) ? (Variable.add_fixes (map fst skos) #> snd)) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 207 | end | 
| 57288 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 208 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 209 | val (lems, _) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 210 | fold_map add_lemma (map_filter (get_role (curry (op =) Lemma)) atp_proof) ctxt | 
| 54700 | 211 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 212 | val bot = #1 (List.last atp_proof) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 213 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 214 | val refute_graph = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 215 | atp_proof | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 216 | |> map (fn (name, _, _, _, from) => (from, name)) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 217 | |> make_refute_graph bot | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 218 | |> fold (Atom_Graph.default_node o rpair ()) conjs | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 219 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 220 | val axioms = axioms_of_refute_graph refute_graph conjs | 
| 54700 | 221 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 222 | val tainted = tainted_atoms_of_refute_graph refute_graph conjs | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 223 | val is_clause_tainted = exists (member (op =) tainted) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 224 | val steps = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 225 | Symtab.empty | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 226 | |> fold (fn (name as (s, _), role, t, rule, _) => | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 227 | Symtab.update_new (s, (rule, t | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 228 | |> (if is_clause_tainted [name] then | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 229 | HOLogic.dest_Trueprop | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 230 | #> role <> Conjecture ? s_not | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 231 | #> fold exists_of (map Var (Term.add_vars t [])) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 232 | #> HOLogic.mk_Trueprop | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 233 | else | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 234 | I)))) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 235 | atp_proof | 
| 58516 
1edba0152491
insert lemmas closer to where they are needed, both for esthetics and (primarily) for correctness in case the lemma refers to a skolem
 blanchet parents: 
58514diff
changeset | 236 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 237 | fun is_referenced_in_step _ (Let _) = false | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 238 | | is_referenced_in_step l (Prove (_, _, _, _, subs, (ls, _), _, _)) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 239 | member (op =) ls l orelse exists (is_referenced_in_proof l) subs | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 240 | and is_referenced_in_proof l (Proof (_, _, steps)) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 241 | exists (is_referenced_in_step l) steps | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 242 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 243 | fun insert_lemma_in_step lem | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 244 | (step as Prove (qs, fix, l, t, subs, (ls, gs), meths, comment)) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 245 | let val l' = the (label_of_isar_step lem) in | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 246 | if member (op =) ls l' then | 
| 58517 | 247 | [lem, step] | 
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 248 | else | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 249 | let val refs = map (is_referenced_in_proof l') subs in | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 250 | if length (filter I refs) = 1 then | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 251 | let | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 252 | val subs' = map2 (fn false => I | true => insert_lemma_in_proof lem) refs | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 253 | subs | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 254 | in | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 255 | [Prove (qs, fix, l, t, subs', (ls, gs), meths, comment)] | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 256 | end | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 257 | else | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 258 | [lem, step] | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 259 | end | 
| 58517 | 260 | end | 
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 261 | and insert_lemma_in_steps lem [] = [lem] | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 262 | | insert_lemma_in_steps lem (step :: steps) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 263 | if is_referenced_in_step (the (label_of_isar_step lem)) step then | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 264 | insert_lemma_in_step lem step @ steps | 
| 54700 | 265 | else | 
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 266 | step :: insert_lemma_in_steps lem steps | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 267 | and insert_lemma_in_proof lem (Proof (fix, assms, steps)) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 268 | Proof (fix, assms, insert_lemma_in_steps lem steps) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 269 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 270 | val rule_of_clause_id = fst o the o Symtab.lookup steps o fst | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 271 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 272 | val finish_off = close_form #> rename_bound_vars | 
| 54755 | 273 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 274 | fun prop_of_clause [(num, _)] = Symtab.lookup steps num |> the |> snd |> finish_off | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 275 | | prop_of_clause names = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 276 | let | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 277 | val lits = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 278 | map (HOLogic.dest_Trueprop o snd) (map_filter (Symtab.lookup steps o fst) names) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 279 | in | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 280 | (case List.partition (can HOLogic.dest_not) lits of | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 281 | (negs as _ :: _, pos as _ :: _) => | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 282 | s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 283 | Library.foldr1 s_disj pos) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 284 |                   | _ => fold (curry s_disj) lits @{term False})
 | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 285 | end | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 286 | |> HOLogic.mk_Trueprop |> finish_off | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 287 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 288 | fun maybe_show outer c = if outer andalso eq_set (op =) (c, conjs) then [Show] else [] | 
| 55280 | 289 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 290 | fun isar_steps outer predecessor accum [] = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 291 | accum | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 292 | |> (if tainted = [] then | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 293 | (* e.g., trivial, empty proof by Z3 *) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 294 | cons (Prove (if outer then [Show] else [], [], no_label, concl_t, [], | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 295 | sort_facts (the_list predecessor, []), massage_methods systematic_methods', | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 296 | "")) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 297 | else | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 298 | I) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 299 | |> rev | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 300 | | isar_steps outer _ accum (Have (id, (gamma, c)) :: infs) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 301 | let | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 302 | val l = label_of_clause c | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 303 | val t = prop_of_clause c | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 304 | val rule = rule_of_clause_id id | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 305 | val skolem = is_skolemize_rule rule | 
| 54700 | 306 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 307 | val deps = ([], []) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 308 | |> fold add_fact_of_dependency gamma | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 309 | |> is_maybe_ext_rule rule ? add_global_fact [short_thm_name ctxt ext] | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 310 | |> sort_facts | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 311 | val meths = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 312 | (if skolem then skolem_methods | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 313 | else if is_arith_rule rule then arith_methods | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 314 | else if is_datatype_rule rule then datatype_methods | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 315 | else systematic_methods') | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 316 | |> massage_methods | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 317 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 318 | fun prove sub facts = Prove (maybe_show outer c, [], l, t, sub, facts, meths, "") | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 319 | fun steps_of_rest step = isar_steps outer (SOME l) (step :: accum) infs | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 320 | in | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 321 | if is_clause_tainted c then | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 322 | (case gamma of | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 323 | [g] => | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 324 | if skolem andalso is_clause_tainted g then | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 325 | let | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 326 | val skos = skolems_of ctxt (prop_of_clause g) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 327 | val subproof = Proof (skos, [], rev accum) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 328 | in | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 329 | isar_steps outer (SOME l) [prove [subproof] ([], [])] infs | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 330 | end | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 331 | else | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 332 | steps_of_rest (prove [] deps) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 333 | | _ => steps_of_rest (prove [] deps)) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 334 | else | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 335 | steps_of_rest | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 336 | (if skolem then | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 337 | (case skolems_of ctxt t of | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 338 | [] => prove [] deps | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 339 | | skos => Prove ([], skos, l, t, [], deps, meths, "")) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 340 | else | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 341 | prove [] deps) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 342 | end | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 343 | | isar_steps outer predecessor accum (Cases cases :: infs) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 344 | let | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 345 | fun isar_case (c, subinfs) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 346 | isar_proof false [] [(label_of_clause c, prop_of_clause c)] [] subinfs | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 347 | val c = succedent_of_cases cases | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 348 | val l = label_of_clause c | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 349 | val t = prop_of_clause c | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 350 | val step = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 351 | Prove (maybe_show outer c, [], l, t, | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 352 | map isar_case (filter_out (null o snd) cases), | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 353 | sort_facts (the_list predecessor, []), massage_methods systematic_methods', | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 354 | "") | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 355 | in | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 356 | isar_steps outer (SOME l) (step :: accum) infs | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 357 | end | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 358 | and isar_proof outer fix assms lems infs = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 359 | Proof (fix, assms, | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 360 | fold_rev insert_lemma_in_steps lems (isar_steps outer NONE [] infs)) | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 361 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 362 | val trace = Config.get ctxt trace | 
| 55222 | 363 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 364 | val canonical_isar_proof = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 365 | refute_graph | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 366 | |> trace ? tap (tracing o prefix "Refute graph:\n" o string_of_refute_graph) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 367 | |> redirect_graph axioms tainted bot | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 368 | |> trace ? tap (tracing o prefix "Direct proof:\n" o string_of_direct_proof) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 369 | |> isar_proof true params assms lems | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 370 | |> postprocess_isar_proof_remove_show_stuttering | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 371 | |> postprocess_isar_proof_remove_unreferenced_steps I | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 372 | |> relabel_isar_proof_canonically | 
| 54754 | 373 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 374 | val ctxt = ctxt |> enrich_context_with_local_facts canonical_isar_proof | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 375 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 376 | val preplay_data = Unsynchronized.ref Canonical_Label_Tab.empty | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 377 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 378 | val _ = fold_isar_steps (fn meth => | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 379 | K (set_preplay_outcomes_of_isar_step ctxt preplay_timeout preplay_data meth [])) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 380 | (steps_of_isar_proof canonical_isar_proof) () | 
| 55299 | 381 | |
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 382 | fun str_of_preplay_outcome outcome = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 383 | if Lazy.is_finished outcome then string_of_play_outcome (Lazy.force outcome) else "?" | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 384 | fun str_of_meth l meth = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 385 | string_of_proof_method ctxt [] meth ^ " " ^ | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 386 | str_of_preplay_outcome | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 387 | (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 388 | fun comment_of l = map (str_of_meth l) #> commas | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 389 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 390 | fun trace_isar_proof label proof = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 391 | if trace then | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 392 | tracing (timestamp () ^ "\n" ^ label ^ ":\n\n" ^ | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 393 | string_of_isar_proof ctxt subgoal subgoal_count | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 394 | (comment_isar_proof comment_of proof) ^ "\n") | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 395 | else | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 396 | () | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 397 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 398 | fun comment_of l (meth :: _) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 399 | (case (verbose, | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 400 | Lazy.force (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth)) of | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 401 | (false, Played _) => "" | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 402 | | (_, outcome) => string_of_play_outcome outcome) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 403 | |
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 404 | val (play_outcome, isar_proof) = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 405 | canonical_isar_proof | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 406 | |> tap (trace_isar_proof "Original") | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 407 | |> compress_isar_proof ctxt compress preplay_timeout preplay_data | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 408 | |> tap (trace_isar_proof "Compressed") | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 409 | |> postprocess_isar_proof_remove_unreferenced_steps | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 410 | (keep_fastest_method_of_isar_step (!preplay_data) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 411 | #> minimize ? minimize_isar_step_dependencies ctxt preplay_data) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 412 | |> tap (trace_isar_proof "Minimized") | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 413 | |> `(preplay_outcome_of_isar_proof (!preplay_data)) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 414 | ||> (comment_isar_proof comment_of | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 415 | #> chain_isar_proof | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 416 | #> kill_useless_labels_in_isar_proof | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 417 | #> relabel_isar_proof_nicely | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 418 | #> rationalize_obtains_in_isar_proofs ctxt) | 
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 419 | in | 
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 420 | (case (num_chained, add_isar_steps (steps_of_isar_proof isar_proof) 0) of | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 421 | (0, 1) => | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 422 | one_line_proof_text ctxt 0 | 
| 67560 | 423 | (if is_less (play_outcome_ord (play_outcome, one_line_play)) then | 
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 424 | (case isar_proof of | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 425 | Proof (_, _, [Prove (_, _, _, _, _, (_, gfs), meth :: _, _)]) => | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 426 | let | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 427 | val used_facts' = filter (fn (s, (sc, _)) => | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 428 | member (op =) gfs s andalso sc <> Chained) used_facts | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 429 | in | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 430 | ((used_facts', (meth, play_outcome)), banner, subgoal, subgoal_count) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 431 | end) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 432 | else | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 433 | one_line_params) ^ | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 434 | (if isar_proofs = SOME true then "\n(No Isar proof available.)" else "") | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 435 | | (_, num_steps) => | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 436 | let | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 437 | val msg = | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 438 | (if verbose then [string_of_int num_steps ^ " step" ^ plural_s num_steps] | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 439 | else []) @ | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 440 | (if do_preplay then [string_of_play_outcome play_outcome] else []) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 441 | in | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 442 | one_line_proof_text ctxt 0 one_line_params ^ | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 443 |                 "\n\nIsar proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^
 | 
| 63518 | 444 | Active.sendback_markup_command | 
| 62220 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 445 | (string_of_isar_proof ctxt subgoal subgoal_count isar_proof) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 446 | end) | 
| 
0e17a97234bd
avoid error in Isar proof reconstruction if no ATP proof is available
 blanchet parents: 
60612diff
changeset | 447 | end | 
| 49883 | 448 | end | 
| 57056 | 449 | in | 
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 450 | if debug then | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 451 | generate_proof_text () | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 452 | else | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 453 | (case try generate_proof_text () of | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 454 | SOME s => s | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 455 | | NONE => | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 456 | one_line_proof_text ctxt 0 one_line_params ^ | 
| 63692 | 457 | (if isar_proofs = SOME true then "\nWarning: Isar proof construction failed" else "")) | 
| 57056 | 458 | end | 
| 49883 | 459 | |
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 460 | fun isar_proof_would_be_a_good_idea smt_proofs (meth, play) = | 
| 54824 | 461 | (case play of | 
| 58061 | 462 | Played _ => meth = SMT_Method andalso smt_proofs <> SOME true | 
| 62826 | 463 | | Play_Timed_Out time => time > Time.zeroTime | 
| 56093 | 464 | | Play_Failed => true) | 
| 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 | 465 | |
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 466 | fun proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained | 
| 57739 | 467 | (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 | 468 | (if isar_proofs = SOME true orelse | 
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 469 | (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea smt_proofs preplay) then | 
| 60612 
79d71bfea310
removed chained facts from preplaying -- and careful about extra chained facts when removing 'proof -' and 'qed' from one-line Isar proofs
 blanchet parents: 
59577diff
changeset | 470 | isar_proof_text ctxt debug num_chained isar_proofs smt_proofs isar_params | 
| 49883 | 471 | else | 
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56983diff
changeset | 472 | one_line_proof_text ctxt num_chained) one_line_params | 
| 49883 | 473 | |
| 474 | end; |