| author | blanchet | 
| Thu, 28 Aug 2014 20:05:39 +0200 | |
| changeset 58089 | 20e76da3a0ef | 
| parent 58083 | ca7ec8728348 | 
| child 58092 | 4ae52c60603a | 
| 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 | |
| 49918 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 blanchet parents: 
49917diff
changeset | 33 | open Sledgehammer_Util | 
| 55287 | 34 | open Sledgehammer_Proof_Methods | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 35 | open Sledgehammer_Isar_Proof | 
| 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 36 | open Sledgehammer_Isar_Preplay | 
| 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 37 | open Sledgehammer_Isar_Compress | 
| 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 38 | open Sledgehammer_Isar_Minimize | 
| 49914 | 39 | |
| 40 | structure String_Redirect = ATP_Proof_Redirect( | |
| 53586 
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
 blanchet parents: 
53052diff
changeset | 41 | type key = atp_step_name | 
| 49914 | 42 | val ord = fn ((s, _ : string list), (s', _)) => fast_string_ord (s, s') | 
| 43 | val string_of = fst) | |
| 44 | ||
| 49883 | 45 | open String_Redirect | 
| 46 | ||
| 55222 | 47 | val trace = Attrib.setup_config_bool @{binding sledgehammer_isar_trace} (K false)
 | 
| 48 | ||
| 57785 | 49 | val e_definition_rule = "definition" | 
| 57654 
f89c0749533d
'shift_quantors' is not an E skolemization rule (cf. 3ab503b04bdb)
 blanchet parents: 
57288diff
changeset | 50 | val e_skolemize_rule = "skolemize" | 
| 57759 | 51 | val leo2_extcnf_forall_neg_rule = "extcnf_forall_neg" | 
| 57709 | 52 | val satallax_skolemize_rule = "tab_ex" | 
| 54836 | 53 | val spass_pirate_datatype_rule = "DT" | 
| 54746 | 54 | val vampire_skolemisation_rule = "skolemisation" | 
| 57761 | 55 | val veriT_la_generic_rule = "la_generic" | 
| 57708 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 56 | val veriT_simp_arith_rule = "simp_arith" | 
| 57761 | 57 | val veriT_tmp_ite_elim_rule = "tmp_ite_elim" | 
| 58 | val veriT_tmp_skolemize_rule = "tmp_skolemize" | |
| 58061 | 59 | val z3_skolemize_rule = Z3_Proof.string_of_rule Z3_Proof.Skolemize | 
| 60 | 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 | 61 | val zipperposition_cnf_rule = "cnf" | 
| 54746 | 62 | |
| 54772 | 63 | val skolemize_rules = | 
| 57785 | 64 | [e_definition_rule, e_skolemize_rule, leo2_extcnf_forall_neg_rule, satallax_skolemize_rule, | 
| 65 | spass_skolemize_rule, vampire_skolemisation_rule, veriT_tmp_ite_elim_rule, | |
| 66 | veriT_tmp_skolemize_rule, z3_skolemize_rule, zipperposition_cnf_rule] | |
| 54746 | 67 | |
| 54769 
3d6ac2f68bf3
correcly recognize E skolemization steps that are wrapped in a 'shift_quantors' inference
 blanchet parents: 
54768diff
changeset | 68 | val is_skolemize_rule = member (op =) skolemize_rules | 
| 57761 | 69 | fun is_arith_rule rule = | 
| 57762 | 70 | String.isPrefix z3_th_lemma_rule_prefix rule orelse rule = veriT_simp_arith_rule orelse | 
| 57761 | 71 | rule = veriT_la_generic_rule | 
| 54836 | 72 | val is_datatype_rule = String.isPrefix spass_pirate_datatype_rule | 
| 54755 | 73 | |
| 54501 | 74 | fun raw_label_of_num num = (num, 0) | 
| 49914 | 75 | |
| 54501 | 76 | fun label_of_clause [(num, _)] = raw_label_of_num num | 
| 77 | | label_of_clause c = (space_implode "___" (map (fst o raw_label_of_num o fst) c), 0) | |
| 50005 | 78 | |
| 54505 | 79 | fun add_fact_of_dependencies [(_, ss as _ :: _)] = apsnd (union (op =) ss) | 
| 80 | | add_fact_of_dependencies names = apfst (insert (op =) (label_of_clause names)) | |
| 49914 | 81 | |
| 54799 | 82 | fun add_line_pass1 (line as (name, role, t, rule, [])) lines = | 
| 54770 | 83 | (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire) internal facts or | 
| 84 | definitions. *) | |
| 57708 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 85 | if role = Conjecture orelse role = Negated_Conjecture then | 
| 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 86 | line :: lines | 
| 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 87 |     else if t aconv @{prop True} then
 | 
| 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 88 | 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 | 89 | 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 | 90 | line :: lines | 
| 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 91 | else if role = Axiom then | 
| 
4b52c1b319ce
veriT changes for lifted terms, and ite_elim rules.
 fleury parents: 
57705diff
changeset | 92 | lines (* axioms (facts) need no proof lines *) | 
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 93 | else | 
| 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 94 | map (replace_dependencies_in_line (name, [])) lines | 
| 54755 | 95 | | add_line_pass1 line lines = line :: lines | 
| 49914 | 96 | |
| 57771 | 97 | fun add_lines_pass2 res [] = rev res | 
| 98 | | 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 | 99 | let | 
| 57791 | 100 | fun normalize role = | 
| 101 | role = Conjecture ? (HOLogic.dest_Trueprop #> s_not #> HOLogic.mk_Trueprop) | |
| 102 | ||
| 57771 | 103 | val norm_t = normalize role t | 
| 104 | val is_duplicate = | |
| 105 | exists (fn (prev_name, prev_role, prev_t, _, _) => | |
| 106 | member (op =) deps prev_name andalso | |
| 107 | Term.aconv_untyped (normalize prev_role prev_t, norm_t)) | |
| 108 | res | |
| 57770 
6c4ab6f0a6fc
normalize conjectures vs. negated conjectures when comparing terms
 blanchet parents: 
57769diff
changeset | 109 | |
| 57787 | 110 |       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 | 111 | |
| 
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 | 112 | 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 | 113 | is_skolemize_rule rule' andalso member (op =) deps' name | 
| 57770 
6c4ab6f0a6fc
normalize conjectures vs. negated conjectures when comparing terms
 blanchet parents: 
57769diff
changeset | 114 | |
| 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 | 115 | 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 | 116 | in | 
| 57771 | 117 | if is_duplicate orelse | 
| 57770 
6c4ab6f0a6fc
normalize conjectures vs. negated conjectures when comparing terms
 blanchet parents: 
57769diff
changeset | 118 | (role = Plain andalso not (is_skolemize_rule rule) andalso | 
| 
6c4ab6f0a6fc
normalize conjectures vs. negated conjectures when comparing terms
 blanchet parents: 
57769diff
changeset | 119 | not (is_arith_rule rule) andalso not (is_datatype_rule rule) andalso | 
| 
6c4ab6f0a6fc
normalize conjectures vs. negated conjectures when comparing terms
 blanchet parents: 
57769diff
changeset | 120 | not (null lines) andalso looks_boring () andalso not (is_before_skolemize_rule ())) then | 
| 57771 | 121 | 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 | 122 | else | 
| 57771 | 123 | 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 | 124 | end | 
| 49914 | 125 | |
| 126 | type isar_params = | |
| 57783 | 127 | 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 | 128 | * (term, string) atp_step list * thm | 
| 49914 | 129 | |
| 56852 
b38c5b9cf590
added 'satx' to Sledgehammer's portfolio (cf. 'isar_try0')
 blanchet parents: 
56130diff
changeset | 130 | val basic_systematic_methods = [Metis_Method (NONE, NONE), Meson_Method, Blast_Method, SATx_Method] | 
| 57744 
a68b8db8691d
added appropriate method for skolemization of Z3 steps to the mix
 blanchet parents: 
57741diff
changeset | 131 | val basic_simp_based_methods = [Auto_Method, Simp_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 | 132 | val basic_arith_methods = [Linarith_Method, Presburger_Method, Algebra_Method] | 
| 55311 | 133 | |
| 57744 
a68b8db8691d
added appropriate method for skolemization of Z3 steps to the mix
 blanchet parents: 
57741diff
changeset | 134 | val arith_methods = basic_arith_methods @ basic_simp_based_methods @ basic_systematic_methods | 
| 55311 | 135 | 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 | 136 | val systematic_methods = | 
| 
a68b8db8691d
added appropriate method for skolemization of Z3 steps to the mix
 blanchet parents: 
57741diff
changeset | 137 | basic_systematic_methods @ basic_arith_methods @ basic_simp_based_methods @ | 
| 57699 | 138 | [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 | 139 | val rewrite_methods = basic_simp_based_methods @ basic_systematic_methods @ basic_arith_methods | 
| 58077 | 140 | fun skolem_methods_of z3 = (basic_systematic_methods, [Skolem_Method]) |> z3 ? swap |> op @ | 
| 54766 
6ac273f176cd
store alternative proof methods in Isar data structure
 blanchet parents: 
54765diff
changeset | 141 | |
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 142 | fun isar_proof_text ctxt debug isar_proofs smt_proofs isar_params | 
| 57739 | 143 | (one_line_params as ((used_facts, (_, one_line_play)), banner, subgoal, subgoal_count)) = | 
| 49883 | 144 | let | 
| 56097 | 145 | val _ = if debug then Output.urgent_message "Constructing Isar proof..." else () | 
| 146 | ||
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 147 | fun generate_proof_text () = | 
| 49883 | 148 | let | 
| 57791 | 149 | val (verbose, alt_metis_args, preplay_timeout, compress, try0, minimize, atp_proof0, goal) = | 
| 57245 | 150 | isar_params () | 
| 55257 | 151 | |
| 57744 
a68b8db8691d
added appropriate method for skolemization of Z3 steps to the mix
 blanchet parents: 
57741diff
changeset | 152 | 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 | 153 | |
| 55311 | 154 | fun massage_methods (meths as meth :: _) = | 
| 57245 | 155 | if not try0 then [meth] | 
| 58061 | 156 | else if smt_proofs = SOME true then SMT_Method :: meths | 
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 157 | else meths | 
| 55273 | 158 | |
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 159 | val (params, _, concl_t) = strip_subgoal goal subgoal ctxt | 
| 55264 | 160 | val fixes = map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params | 
| 161 | 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 | 162 | |
| 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 163 | val do_preplay = preplay_timeout <> Time.zeroTime | 
| 57783 | 164 | val compress = | 
| 165 | (case compress of | |
| 166 | NONE => if isar_proofs = NONE andalso do_preplay then 1000.0 else 10.0 | |
| 167 | | SOME n => n) | |
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 168 | |
| 57288 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 169 | fun is_fixed ctxt = Variable.is_declared ctxt orf Name.is_skolem | 
| 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 170 | fun skolems_of ctxt t = Term.add_frees t [] |> filter_out (is_fixed ctxt o fst) |> rev | 
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 171 | |
| 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 172 | fun get_role keep_role ((num, _), role, t, rule, _) = | 
| 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 173 | if keep_role role then SOME ((raw_label_of_num num, t), rule) else NONE | 
| 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 174 | |
| 49883 | 175 | val atp_proof = | 
| 57791 | 176 | fold_rev add_line_pass1 atp_proof0 [] | 
| 57771 | 177 | |> add_lines_pass2 [] | 
| 54700 | 178 | |
| 54535 
59737a43e044
support Negated_Conjecture as a TPTP role as well (e.g. for SMT proofs)
 blanchet parents: 
54507diff
changeset | 179 | val conjs = | 
| 54700 | 180 | map_filter (fn (name, role, _, _, _) => | 
| 181 | if member (op =) [Conjecture, Negated_Conjecture] role then SOME name else NONE) | |
| 182 | atp_proof | |
| 54751 | 183 | 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 | 184 | |
| 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 185 | fun add_lemma ((l, t), rule) ctxt = | 
| 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 186 | let | 
| 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 187 | val (skos, meths) = | 
| 58077 | 188 | (if is_skolemize_rule rule then | 
| 189 | (skolems_of ctxt t, skolem_methods_of (rule = z3_skolemize_rule)) | |
| 190 | else if is_arith_rule rule then | |
| 191 | ([], arith_methods) | |
| 192 | else | |
| 193 | ([], rewrite_methods)) | |
| 57288 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 194 | ||> massage_methods | 
| 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 195 | in | 
| 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 196 | (Prove ([], skos, l, t, [], ([], []), meths, ""), | 
| 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 197 | ctxt |> not (null skos) ? (Variable.add_fixes (map fst skos) #> snd)) | 
| 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 198 | end | 
| 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 199 | |
| 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 200 | val (lems, _) = | 
| 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 201 | fold_map add_lemma (map_filter (get_role (curry (op =) Lemma)) atp_proof) ctxt | 
| 54700 | 202 | |
| 57791 | 203 | val bot = #1 (List.last atp_proof) | 
| 54700 | 204 | |
| 51145 | 205 | val refute_graph = | 
| 51212 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 206 | atp_proof | 
| 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 207 | |> map (fn (name, _, _, _, from) => (from, name)) | 
| 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 208 | |> make_refute_graph bot | 
| 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 209 | |> fold (Atom_Graph.default_node o rpair ()) conjs | 
| 54700 | 210 | |
| 51145 | 211 | val axioms = axioms_of_refute_graph refute_graph conjs | 
| 54700 | 212 | |
| 51145 | 213 | val tainted = tainted_atoms_of_refute_graph refute_graph conjs | 
| 51156 | 214 | val is_clause_tainted = exists (member (op =) tainted) | 
| 50676 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 215 | val steps = | 
| 49883 | 216 | Symtab.empty | 
| 51201 | 217 | |> fold (fn (name as (s, _), role, t, rule, _) => | 
| 54758 | 218 | Symtab.update_new (s, (rule, t | 
| 219 | |> (if is_clause_tainted [name] then | |
| 54768 
ee0881a54c72
fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
 blanchet parents: 
54767diff
changeset | 220 | HOLogic.dest_Trueprop | 
| 
ee0881a54c72
fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
 blanchet parents: 
54767diff
changeset | 221 | #> role <> Conjecture ? s_not | 
| 54758 | 222 | #> fold exists_of (map Var (Term.add_vars t [])) | 
| 54768 
ee0881a54c72
fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
 blanchet parents: 
54767diff
changeset | 223 | #> HOLogic.mk_Trueprop | 
| 54758 | 224 | else | 
| 225 | I)))) | |
| 226 | atp_proof | |
| 54700 | 227 | |
| 54755 | 228 | val rule_of_clause_id = fst o the o Symtab.lookup steps o fst | 
| 54700 | 229 | |
| 57768 
a63f14f1214c
fine-tuned Isar reconstruction, esp. boolean simplifications
 blanchet parents: 
57767diff
changeset | 230 | val finish_off = close_form #> rename_bound_vars | 
| 57765 | 231 | |
| 57767 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 blanchet parents: 
57765diff
changeset | 232 | fun prop_of_clause [(num, _)] = Symtab.lookup steps num |> the |> snd |> finish_off | 
| 50016 | 233 | | prop_of_clause names = | 
| 50676 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 234 | let | 
| 57765 | 235 | val lits = | 
| 236 | map (HOLogic.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 | 237 | in | 
| 54754 | 238 | (case List.partition (can HOLogic.dest_not) lits of | 
| 50018 
4ea26c74d7ea
use implications rather than disjunctions to improve readability
 blanchet parents: 
50017diff
changeset | 239 | (negs as _ :: _, pos as _ :: _) => | 
| 54507 | 240 | s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos) | 
| 54754 | 241 |               | _ => fold (curry s_disj) lits @{term False})
 | 
| 50018 
4ea26c74d7ea
use implications rather than disjunctions to improve readability
 blanchet parents: 
50017diff
changeset | 242 | end | 
| 57767 
5bc204ca27ce
centralized boolean simplification so that e.g. LEO-II benefits from it
 blanchet parents: 
57765diff
changeset | 243 | |> HOLogic.mk_Trueprop |> finish_off | 
| 54700 | 244 | |
| 58083 | 245 | fun maybe_show outer c = if outer andalso eq_set (op =) (c, conjs) then [Show] else [] | 
| 54700 | 246 | |
| 247 | fun isar_steps outer predecessor accum [] = | |
| 248 | accum | |
| 249 | |> (if tainted = [] then | |
| 58083 | 250 | (* e.g., trivial, empty proof by Z3 *) | 
| 54700 | 251 | cons (Prove (if outer then [Show] else [], [], no_label, concl_t, [], | 
| 57776 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57771diff
changeset | 252 | sort_facts (the_list predecessor, []), massage_methods systematic_methods', "")) | 
| 54700 | 253 | else | 
| 254 | I) | |
| 255 | |> rev | |
| 54755 | 256 | | isar_steps outer _ accum (Have (id, (gamma, c)) :: infs) = | 
| 54700 | 257 | let | 
| 258 | val l = label_of_clause c | |
| 259 | val t = prop_of_clause c | |
| 54755 | 260 | val rule = rule_of_clause_id id | 
| 261 | val skolem = is_skolemize_rule rule | |
| 262 | ||
| 57776 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57771diff
changeset | 263 | val deps = ([], []) | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57771diff
changeset | 264 | |> fold add_fact_of_dependencies gamma | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57771diff
changeset | 265 | |> sort_facts | 
| 55244 
12e1a5d8ee48
simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
 blanchet parents: 
55223diff
changeset | 266 | val meths = | 
| 58077 | 267 | (if skolem then skolem_methods_of (rule = z3_skolemize_rule) | 
| 55273 | 268 | else if is_arith_rule rule then arith_methods | 
| 269 | else if is_datatype_rule rule then datatype_methods | |
| 57744 
a68b8db8691d
added appropriate method for skolemization of Z3 steps to the mix
 blanchet parents: 
57741diff
changeset | 270 | else systematic_methods') | 
| 55311 | 271 | |> massage_methods | 
| 55280 | 272 | |
| 58083 | 273 | fun prove sub facts = Prove (maybe_show outer c, [], l, t, sub, facts, meths, "") | 
| 55280 | 274 | fun steps_of_rest step = isar_steps outer (SOME l) (step :: accum) infs | 
| 54700 | 275 | in | 
| 276 | if is_clause_tainted c then | |
| 54712 | 277 | (case gamma of | 
| 54700 | 278 | [g] => | 
| 54755 | 279 | if skolem andalso is_clause_tainted g then | 
| 57758 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 280 | (case skolems_of ctxt (prop_of_clause g) of | 
| 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 281 | [] => steps_of_rest (prove [] deps) | 
| 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 282 | | skos => | 
| 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 283 | let val subproof = Proof (skos, [], rev accum) in | 
| 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 284 | isar_steps outer (SOME l) [prove [subproof] ([], [])] infs | 
| 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 285 | end) | 
| 51148 
2246a2e17f92
tuning -- refactoring in preparation for handling skolemization of conjecture
 blanchet parents: 
51147diff
changeset | 286 | else | 
| 55280 | 287 | steps_of_rest (prove [] deps) | 
| 288 | | _ => steps_of_rest (prove [] deps)) | |
| 54700 | 289 | else | 
| 57288 
ffd928316c75
supports gradual skolemization (cf. Z3) by threading context through correctly
 blanchet parents: 
57287diff
changeset | 290 | steps_of_rest | 
| 57758 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 291 | (if skolem then | 
| 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 292 | (case skolems_of ctxt t of | 
| 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 293 | [] => prove [] deps | 
| 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 294 | | skos => Prove ([], skos, l, t, [], deps, meths, "")) | 
| 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 295 | else | 
| 
b2e6166bf487
beware of 'skolem' rules that do not skolemize (e.g. LEO-II)
 blanchet parents: 
57744diff
changeset | 296 | prove [] deps) | 
| 54700 | 297 | end | 
| 298 | | isar_steps outer predecessor accum (Cases cases :: infs) = | |
| 299 | let | |
| 55186 
fafdf2424c57
don't forget the last inference(s) after conjecture skolemization
 blanchet parents: 
55184diff
changeset | 300 | fun isar_case (c, subinfs) = | 
| 
fafdf2424c57
don't forget the last inference(s) after conjecture skolemization
 blanchet parents: 
55184diff
changeset | 301 | isar_proof false [] [(label_of_clause c, prop_of_clause c)] [] subinfs | 
| 54700 | 302 | val c = succedent_of_cases cases | 
| 303 | val l = label_of_clause c | |
| 304 | val t = prop_of_clause c | |
| 305 | val step = | |
| 58083 | 306 | Prove (maybe_show outer c, [], l, t, | 
| 54760 
a1ac3eaa0d11
generate proper succedent for cases with trivial branches
 blanchet parents: 
54759diff
changeset | 307 | map isar_case (filter_out (null o snd) cases), | 
| 57776 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57771diff
changeset | 308 | sort_facts (the_list predecessor, []), massage_methods systematic_methods', "") | 
| 54700 | 309 | in | 
| 310 | isar_steps outer (SOME l) (step :: accum) infs | |
| 311 | end | |
| 312 | and isar_proof outer fix assms lems infs = | |
| 313 | Proof (fix, assms, lems @ isar_steps outer NONE [] infs) | |
| 314 | ||
| 55222 | 315 | val trace = Config.get ctxt trace | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 316 | |
| 55256 | 317 | val canonical_isar_proof = | 
| 51145 | 318 | refute_graph | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 319 | |> trace ? tap (tracing o prefix "Refute graph: " o string_of_refute_graph) | 
| 51031 
63d71b247323
more robustness in Isar proof reconstruction (cf. bug report by Ondrej)
 blanchet parents: 
51026diff
changeset | 320 | |> redirect_graph axioms tainted bot | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 321 | |> trace ? tap (tracing o prefix "Direct proof: " o string_of_direct_proof) | 
| 54754 | 322 | |> isar_proof true params assms lems | 
| 58083 | 323 | |> postprocess_isar_proof_remove_show_stuttering | 
| 55213 | 324 | |> postprocess_isar_proof_remove_unreferenced_steps I | 
| 325 | |> relabel_isar_proof_canonically | |
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 326 | |
| 55286 | 327 | val ctxt = ctxt |> enrich_context_with_local_facts canonical_isar_proof | 
| 55256 | 328 | |
| 55260 | 329 | val preplay_data = Unsynchronized.ref Canonical_Label_Tab.empty | 
| 330 | ||
| 55264 | 331 | val _ = fold_isar_steps (fn meth => | 
| 57054 | 332 | K (set_preplay_outcomes_of_isar_step ctxt preplay_timeout preplay_data meth [])) | 
| 55260 | 333 | (steps_of_isar_proof canonical_isar_proof) () | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 334 | |
| 55223 
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
 blanchet parents: 
55222diff
changeset | 335 | fun str_of_preplay_outcome outcome = | 
| 
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
 blanchet parents: 
55222diff
changeset | 336 | if Lazy.is_finished outcome then string_of_play_outcome (Lazy.force outcome) else "?" | 
| 
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
 blanchet parents: 
55222diff
changeset | 337 | fun str_of_meth l meth = | 
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56983diff
changeset | 338 | string_of_proof_method ctxt [] meth ^ " " ^ | 
| 55266 | 339 | str_of_preplay_outcome (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth) | 
| 55244 
12e1a5d8ee48
simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
 blanchet parents: 
55223diff
changeset | 340 | fun comment_of l = map (str_of_meth l) #> commas | 
| 55222 | 341 | |
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 342 | fun trace_isar_proof label proof = | 
| 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 343 | if trace then | 
| 55299 | 344 | tracing (timestamp () ^ "\n" ^ label ^ ":\n\n" ^ | 
| 56097 | 345 | string_of_isar_proof ctxt subgoal subgoal_count | 
| 346 | (comment_isar_proof comment_of proof) ^ "\n") | |
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 347 | else | 
| 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 348 | () | 
| 54754 | 349 | |
| 55299 | 350 | fun comment_of l (meth :: _) = | 
| 351 | (case (verbose, | |
| 352 | Lazy.force (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth)) of | |
| 353 | (false, Played _) => "" | |
| 354 | | (_, outcome) => string_of_play_outcome outcome) | |
| 355 | ||
| 54828 | 356 | val (play_outcome, isar_proof) = | 
| 55256 | 357 | canonical_isar_proof | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 358 | |> tap (trace_isar_proof "Original") | 
| 57245 | 359 | |> compress_isar_proof ctxt compress preplay_timeout preplay_data | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 360 | |> tap (trace_isar_proof "Compressed") | 
| 55213 | 361 | |> postprocess_isar_proof_remove_unreferenced_steps | 
| 55266 | 362 | (keep_fastest_method_of_isar_step (!preplay_data) | 
| 57054 | 363 | #> minimize ? minimize_isar_step_dependencies ctxt preplay_data) | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 364 | |> tap (trace_isar_proof "Minimized") | 
| 55260 | 365 | |> `(preplay_outcome_of_isar_proof (!preplay_data)) | 
| 55325 | 366 | ||> (comment_isar_proof comment_of | 
| 367 | #> chain_isar_proof | |
| 368 | #> kill_useless_labels_in_isar_proof | |
| 57792 | 369 | #> relabel_isar_proof_nicely | 
| 370 | #> rationalize_obtains_in_isar_proofs ctxt) | |
| 49883 | 371 | in | 
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 372 | (case add_isar_steps (steps_of_isar_proof isar_proof) 0 of | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 373 | 1 => | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 374 | one_line_proof_text ctxt 0 | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 375 | (if play_outcome_ord (play_outcome, one_line_play) = LESS then | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 376 | (case isar_proof of | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 377 | Proof (_, _, [Prove (_, _, _, _, _, (_, gfs), meth :: _, _)]) => | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 378 | let val used_facts' = filter (member (op =) gfs o fst) used_facts in | 
| 57739 | 379 | ((used_facts', (meth, play_outcome)), banner, subgoal, subgoal_count) | 
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 380 | end) | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 381 | else | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 382 | one_line_params) ^ | 
| 57780 | 383 | (if isar_proofs = SOME true then "\n(No Isar proof available.)" | 
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 384 | else "") | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 385 | | num_steps => | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 386 | let | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 387 | val msg = | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 388 | (if verbose then [string_of_int num_steps ^ " step" ^ plural_s num_steps] else []) @ | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 389 | (if do_preplay then [string_of_play_outcome play_outcome] else []) | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 390 | in | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 391 | one_line_proof_text ctxt 0 one_line_params ^ | 
| 57780 | 392 |             "\n\nIsar proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^
 | 
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 393 | Active.sendback_markup [Markup.padding_command] | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 394 | (string_of_isar_proof ctxt subgoal subgoal_count isar_proof) | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 395 | end) | 
| 49883 | 396 | end | 
| 57056 | 397 | in | 
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 398 | if debug then | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 399 | generate_proof_text () | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 400 | else | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 401 | (case try generate_proof_text () of | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 402 | SOME s => s | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 403 | | NONE => | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 404 | one_line_proof_text ctxt 0 one_line_params ^ | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 405 | (if isar_proofs = SOME true then "\nWarning: Isar proof construction failed." else "")) | 
| 57056 | 406 | end | 
| 49883 | 407 | |
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 408 | fun isar_proof_would_be_a_good_idea smt_proofs (meth, play) = | 
| 54824 | 409 | (case play of | 
| 58061 | 410 | Played _ => meth = SMT_Method andalso smt_proofs <> SOME true | 
| 56093 | 411 | | Play_Timed_Out time => Time.> (time, Time.zeroTime) | 
| 412 | | 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 | 413 | |
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 414 | fun proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained | 
| 57739 | 415 | (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 | 416 | (if isar_proofs = SOME true orelse | 
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 417 | (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea smt_proofs preplay) then | 
| 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 418 | isar_proof_text ctxt debug isar_proofs smt_proofs isar_params | 
| 49883 | 419 | else | 
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56983diff
changeset | 420 | one_line_proof_text ctxt num_chained) one_line_params | 
| 49883 | 421 | |
| 422 | end; |