| author | wenzelm | 
| Tue, 08 Apr 2014 23:05:21 +0200 | |
| changeset 56477 | 57b5c8db55f1 | 
| parent 56130 | 1ef77430713b | 
| child 56852 | b38c5b9cf590 | 
| 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 = | 
| 55296 | 19 | bool * (string option * string option) * Time.time * real * 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 | ||
| 54769 
3d6ac2f68bf3
correcly recognize E skolemization steps that are wrapped in a 'shift_quantors' inference
 blanchet parents: 
54768diff
changeset | 49 | val e_skolemize_rules = ["skolemize", "shift_quantors"] | 
| 54836 | 50 | val spass_pirate_datatype_rule = "DT" | 
| 54746 | 51 | val vampire_skolemisation_rule = "skolemisation" | 
| 52 | (* TODO: Use "Z3_Proof.string_of_rule" once it is moved to Isabelle *) | |
| 54753 | 53 | val z3_skolemize_rule = "sk" | 
| 54 | val z3_th_lemma_rule = "th-lemma" | |
| 54746 | 55 | |
| 54772 | 56 | val skolemize_rules = | 
| 57 | e_skolemize_rules @ [spass_skolemize_rule, vampire_skolemisation_rule, z3_skolemize_rule] | |
| 54746 | 58 | |
| 54769 
3d6ac2f68bf3
correcly recognize E skolemization steps that are wrapped in a 'shift_quantors' inference
 blanchet parents: 
54768diff
changeset | 59 | val is_skolemize_rule = member (op =) skolemize_rules | 
| 54755 | 60 | val is_arith_rule = String.isPrefix z3_th_lemma_rule | 
| 54836 | 61 | val is_datatype_rule = String.isPrefix spass_pirate_datatype_rule | 
| 54755 | 62 | |
| 54501 | 63 | fun raw_label_of_num num = (num, 0) | 
| 49914 | 64 | |
| 54501 | 65 | fun label_of_clause [(num, _)] = raw_label_of_num num | 
| 66 | | label_of_clause c = (space_implode "___" (map (fst o raw_label_of_num o fst) c), 0) | |
| 50005 | 67 | |
| 54505 | 68 | fun add_fact_of_dependencies [(_, ss as _ :: _)] = apsnd (union (op =) ss) | 
| 69 | | add_fact_of_dependencies names = apfst (insert (op =) (label_of_clause names)) | |
| 49914 | 70 | |
| 56126 | 71 | fun is_True_prop t = t aconv @{prop True}
 | 
| 54758 | 72 | |
| 54799 | 73 | fun add_line_pass1 (line as (name, role, t, rule, [])) lines = | 
| 54770 | 74 | (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire) internal facts or | 
| 75 | definitions. *) | |
| 56126 | 76 | if role = Conjecture orelse role = Negated_Conjecture then | 
| 54746 | 77 | line :: lines | 
| 56130 
1ef77430713b
consolidate consecutive steps that prove the same formula
 blanchet parents: 
56128diff
changeset | 78 | else if is_True_prop t then | 
| 
1ef77430713b
consolidate consecutive steps that prove the same formula
 blanchet parents: 
56128diff
changeset | 79 | map (replace_dependencies_in_line (name, [])) lines | 
| 56126 | 80 | else if role = Lemma orelse role = Hypothesis orelse is_arith_rule rule then | 
| 56130 
1ef77430713b
consolidate consecutive steps that prove the same formula
 blanchet parents: 
56128diff
changeset | 81 | line :: lines | 
| 54505 | 82 | else if role = Axiom then | 
| 56130 
1ef77430713b
consolidate consecutive steps that prove the same formula
 blanchet parents: 
56128diff
changeset | 83 | lines (* axioms (facts) need no proof lines *) | 
| 49914 | 84 | else | 
| 85 | map (replace_dependencies_in_line (name, [])) lines | |
| 54755 | 86 | | add_line_pass1 line lines = line :: lines | 
| 49914 | 87 | |
| 56130 
1ef77430713b
consolidate consecutive steps that prove the same formula
 blanchet parents: 
56128diff
changeset | 88 | fun add_lines_pass2 res _ [] = rev res | 
| 
1ef77430713b
consolidate consecutive steps that prove the same formula
 blanchet parents: 
56128diff
changeset | 89 | | add_lines_pass2 res prev_t ((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 | 90 | let | 
| 
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 | 91 | val is_last_line = null 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 | 92 | |
| 
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 | 93 | fun looks_interesting () = | 
| 56130 
1ef77430713b
consolidate consecutive steps that prove the same formula
 blanchet parents: 
56128diff
changeset | 94 | not (is_True_prop t) andalso not (t aconv prev_t) andalso null (Term.add_tvars t []) andalso | 
| 
1ef77430713b
consolidate consecutive steps that prove the same formula
 blanchet parents: 
56128diff
changeset | 95 | length deps >= 2 andalso not (can the_single 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 | 96 | |
| 
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 | 97 | 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 | 98 | is_skolemize_rule rule' andalso member (op =) deps' name | 
| 
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 | 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 | 100 | in | 
| 
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 | 101 | if role <> Plain orelse is_skolemize_rule rule orelse is_arith_rule rule orelse | 
| 
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 | 102 | is_datatype_rule rule orelse is_last_line orelse looks_interesting () orelse | 
| 
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 | 103 | is_before_skolemize_rule () then | 
| 56130 
1ef77430713b
consolidate consecutive steps that prove the same formula
 blanchet parents: 
56128diff
changeset | 104 | add_lines_pass2 (line :: res) t 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 | else | 
| 56130 
1ef77430713b
consolidate consecutive steps that prove the same formula
 blanchet parents: 
56128diff
changeset | 106 | add_lines_pass2 res t (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 | 107 | end | 
| 49914 | 108 | |
| 109 | type isar_params = | |
| 55296 | 110 | bool * (string option * string option) * Time.time * real * bool * bool | 
| 55288 
1a4358d14ce2
added 'smt' option to control generation of 'by smt' proofs
 blanchet parents: 
55287diff
changeset | 111 | * (term, string) atp_step list * thm | 
| 49914 | 112 | |
| 55323 
253a029335a2
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
 blanchet parents: 
55311diff
changeset | 113 | val basic_systematic_methods = [Metis_Method (NONE, NONE), Meson_Method, Blast_Method] | 
| 56097 | 114 | val 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 | 115 | val basic_arith_methods = [Linarith_Method, Presburger_Method, Algebra_Method] | 
| 55311 | 116 | |
| 117 | val arith_methods = basic_arith_methods @ simp_based_methods @ basic_systematic_methods | |
| 118 | val datatype_methods = [Simp_Method, Simp_Size_Method] | |
| 119 | val systematic_methods0 = basic_systematic_methods @ basic_arith_methods @ simp_based_methods @ | |
| 120 | [Metis_Method (SOME no_typesN, NONE)] | |
| 121 | val rewrite_methods = simp_based_methods @ basic_systematic_methods @ basic_arith_methods | |
| 122 | val skolem_methods = basic_systematic_methods | |
| 54766 
6ac273f176cd
store alternative proof methods in Isar data structure
 blanchet parents: 
54765diff
changeset | 123 | |
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 124 | fun isar_proof_text ctxt debug isar_proofs smt_proofs isar_params | 
| 49918 
cf441f4a358b
renamed Isar-proof related options + changed semantics of Isar shrinking
 blanchet parents: 
49917diff
changeset | 125 | (one_line_params as (_, _, _, _, subgoal, subgoal_count)) = | 
| 49883 | 126 | let | 
| 56097 | 127 | val _ = if debug then Output.urgent_message "Constructing Isar proof..." else () | 
| 128 | ||
| 49883 | 129 | fun isar_proof_of () = | 
| 130 | let | |
| 55296 | 131 | val SOME (verbose, alt_metis_args, preplay_timeout, compress_isar, try0_isar, minimize, | 
| 55267 | 132 | atp_proof, goal) = try isar_params () | 
| 55257 | 133 | |
| 55311 | 134 | val systematic_methods = insert (op =) (Metis_Method alt_metis_args) systematic_methods0 | 
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 135 | |
| 55311 | 136 | fun massage_methods (meths as meth :: _) = | 
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 137 | if not try0_isar then [meth] | 
| 56081 | 138 | else if smt_proofs = SOME true then SMT2_Method :: meths | 
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 139 | else meths | 
| 55273 | 140 | |
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 141 | val (params, _, concl_t) = strip_subgoal goal subgoal ctxt | 
| 55264 | 142 | val fixes = map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params | 
| 143 | 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 | 144 | |
| 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 145 | val do_preplay = preplay_timeout <> Time.zeroTime | 
| 55253 | 146 | val compress_isar = if isar_proofs = NONE andalso do_preplay then 1000.0 else compress_isar | 
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 147 | |
| 55948 | 148 | val is_fixed = Variable.is_declared ctxt orf Name.is_skolem | 
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 149 | fun skolems_of t = Term.add_frees t [] |> filter_out (is_fixed o fst) |> rev | 
| 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 150 | |
| 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 151 | fun get_role keep_role ((num, _), role, t, rule, _) = | 
| 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 152 | 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 | 153 | |
| 49883 | 154 | val atp_proof = | 
| 155 | atp_proof | |
| 54755 | 156 | |> rpair [] |-> fold_rev add_line_pass1 | 
| 56130 
1ef77430713b
consolidate consecutive steps that prove the same formula
 blanchet parents: 
56128diff
changeset | 157 | |> add_lines_pass2 [] Term.dummy | 
| 54700 | 158 | |
| 54535 
59737a43e044
support Negated_Conjecture as a TPTP role as well (e.g. for SMT proofs)
 blanchet parents: 
54507diff
changeset | 159 | val conjs = | 
| 54700 | 160 | map_filter (fn (name, role, _, _, _) => | 
| 161 | if member (op =) [Conjecture, Negated_Conjecture] role then SOME name else NONE) | |
| 162 | atp_proof | |
| 54751 | 163 | val assms = map_filter (Option.map fst o get_role (curry (op =) Hypothesis)) atp_proof | 
| 54700 | 164 | val lems = | 
| 165 | map_filter (get_role (curry (op =) Lemma)) atp_proof | |
| 54751 | 166 | |> map (fn ((l, t), rule) => | 
| 54753 | 167 | let | 
| 55244 
12e1a5d8ee48
simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
 blanchet parents: 
55223diff
changeset | 168 | val (skos, meths) = | 
| 55273 | 169 | (if is_skolemize_rule rule then (skolems_of t, skolem_methods) | 
| 170 | else if is_arith_rule rule then ([], arith_methods) | |
| 171 | else ([], rewrite_methods)) | |
| 55311 | 172 | ||> massage_methods | 
| 54753 | 173 | in | 
| 55299 | 174 | Prove ([], skos, l, t, [], ([], []), meths, "") | 
| 54753 | 175 | end) | 
| 54700 | 176 | |
| 51212 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 177 | val bot = atp_proof |> List.last |> #1 | 
| 54700 | 178 | |
| 51145 | 179 | val refute_graph = | 
| 51212 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 180 | atp_proof | 
| 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 181 | |> map (fn (name, _, _, _, from) => (from, name)) | 
| 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 182 | |> make_refute_graph bot | 
| 
2bbcc9cc12b4
ensure all conjecture clauses are in the graph -- to prevent exceptions later
 blanchet parents: 
51208diff
changeset | 183 | |> fold (Atom_Graph.default_node o rpair ()) conjs | 
| 54700 | 184 | |
| 51145 | 185 | val axioms = axioms_of_refute_graph refute_graph conjs | 
| 54700 | 186 | |
| 51145 | 187 | val tainted = tainted_atoms_of_refute_graph refute_graph conjs | 
| 51156 | 188 | val is_clause_tainted = exists (member (op =) tainted) | 
| 50676 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 189 | val steps = | 
| 49883 | 190 | Symtab.empty | 
| 51201 | 191 | |> fold (fn (name as (s, _), role, t, rule, _) => | 
| 54758 | 192 | Symtab.update_new (s, (rule, t | 
| 193 | |> (if is_clause_tainted [name] then | |
| 54768 
ee0881a54c72
fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
 blanchet parents: 
54767diff
changeset | 194 | HOLogic.dest_Trueprop | 
| 
ee0881a54c72
fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
 blanchet parents: 
54767diff
changeset | 195 | #> role <> Conjecture ? s_not | 
| 54758 | 196 | #> fold exists_of (map Var (Term.add_vars t [])) | 
| 54768 
ee0881a54c72
fixed confusion between 'prop' and 'bool' introduced in 4960647932ec
 blanchet parents: 
54767diff
changeset | 197 | #> HOLogic.mk_Trueprop | 
| 54758 | 198 | else | 
| 199 | I)))) | |
| 200 | atp_proof | |
| 54700 | 201 | |
| 54755 | 202 | val rule_of_clause_id = fst o the o Symtab.lookup steps o fst | 
| 54700 | 203 | |
| 54757 
4960647932ec
use 'prop' rather than 'bool' systematically in Isar reconstruction code
 blanchet parents: 
54756diff
changeset | 204 | fun prop_of_clause [(num, _)] = Symtab.lookup steps num |> the |> snd |> close_form | 
| 50016 | 205 | | prop_of_clause names = | 
| 50676 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 206 | let | 
| 54758 | 207 | val lits = map (HOLogic.dest_Trueprop o snd) | 
| 208 | (map_filter (Symtab.lookup steps o fst) names) | |
| 50676 
83b8a5a39709
generate "obtain" steps corresponding to skolemization inferences
 blanchet parents: 
50675diff
changeset | 209 | in | 
| 54754 | 210 | (case List.partition (can HOLogic.dest_not) lits of | 
| 50018 
4ea26c74d7ea
use implications rather than disjunctions to improve readability
 blanchet parents: 
50017diff
changeset | 211 | (negs as _ :: _, pos as _ :: _) => | 
| 54507 | 212 | s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos) | 
| 54754 | 213 |               | _ => fold (curry s_disj) lits @{term False})
 | 
| 50018 
4ea26c74d7ea
use implications rather than disjunctions to improve readability
 blanchet parents: 
50017diff
changeset | 214 | end | 
| 50016 | 215 | |> HOLogic.mk_Trueprop |> close_form | 
| 54700 | 216 | |
| 55169 | 217 | fun maybe_show outer c = (outer andalso eq_set (op =) (c, conjs)) ? cons Show | 
| 54700 | 218 | |
| 219 | fun isar_steps outer predecessor accum [] = | |
| 220 | accum | |
| 221 | |> (if tainted = [] then | |
| 222 | cons (Prove (if outer then [Show] else [], [], no_label, concl_t, [], | |
| 55311 | 223 | (the_list predecessor, []), massage_methods systematic_methods, "")) | 
| 54700 | 224 | else | 
| 225 | I) | |
| 226 | |> rev | |
| 54755 | 227 | | isar_steps outer _ accum (Have (id, (gamma, c)) :: infs) = | 
| 54700 | 228 | let | 
| 229 | val l = label_of_clause c | |
| 230 | val t = prop_of_clause c | |
| 54755 | 231 | val rule = rule_of_clause_id id | 
| 232 | val skolem = is_skolemize_rule rule | |
| 233 | ||
| 55279 | 234 | val deps = fold add_fact_of_dependencies gamma ([], []) | 
| 55244 
12e1a5d8ee48
simplified data structure -- eliminated distinction between 'first-class' and 'second-class' proof methods
 blanchet parents: 
55223diff
changeset | 235 | val meths = | 
| 55273 | 236 | (if skolem then skolem_methods | 
| 237 | else if is_arith_rule rule then arith_methods | |
| 238 | else if is_datatype_rule rule then datatype_methods | |
| 55311 | 239 | else systematic_methods) | 
| 240 | |> massage_methods | |
| 55280 | 241 | |
| 55299 | 242 | fun prove sub facts = Prove (maybe_show outer c [], [], l, t, sub, facts, meths, "") | 
| 55280 | 243 | fun steps_of_rest step = isar_steps outer (SOME l) (step :: accum) infs | 
| 54700 | 244 | in | 
| 245 | if is_clause_tainted c then | |
| 54712 | 246 | (case gamma of | 
| 54700 | 247 | [g] => | 
| 54755 | 248 | if skolem andalso is_clause_tainted g then | 
| 54751 | 249 | let val subproof = Proof (skolems_of (prop_of_clause g), [], rev accum) in | 
| 55280 | 250 | isar_steps outer (SOME l) [prove [subproof] ([], [])] infs | 
| 54700 | 251 | end | 
| 51148 
2246a2e17f92
tuning -- refactoring in preparation for handling skolemization of conjecture
 blanchet parents: 
51147diff
changeset | 252 | else | 
| 55280 | 253 | steps_of_rest (prove [] deps) | 
| 254 | | _ => steps_of_rest (prove [] deps)) | |
| 54700 | 255 | else | 
| 55299 | 256 | steps_of_rest (if skolem then Prove ([], skolems_of t, l, t, [], deps, meths, "") | 
| 55280 | 257 | else prove [] deps) | 
| 54700 | 258 | end | 
| 259 | | isar_steps outer predecessor accum (Cases cases :: infs) = | |
| 260 | let | |
| 55186 
fafdf2424c57
don't forget the last inference(s) after conjecture skolemization
 blanchet parents: 
55184diff
changeset | 261 | fun isar_case (c, subinfs) = | 
| 
fafdf2424c57
don't forget the last inference(s) after conjecture skolemization
 blanchet parents: 
55184diff
changeset | 262 | isar_proof false [] [(label_of_clause c, prop_of_clause c)] [] subinfs | 
| 54700 | 263 | val c = succedent_of_cases cases | 
| 264 | val l = label_of_clause c | |
| 265 | val t = prop_of_clause c | |
| 266 | val step = | |
| 54760 
a1ac3eaa0d11
generate proper succedent for cases with trivial branches
 blanchet parents: 
54759diff
changeset | 267 | Prove (maybe_show outer c [], [], l, t, | 
| 
a1ac3eaa0d11
generate proper succedent for cases with trivial branches
 blanchet parents: 
54759diff
changeset | 268 | map isar_case (filter_out (null o snd) cases), | 
| 55311 | 269 | (the_list predecessor, []), massage_methods systematic_methods, "") | 
| 54700 | 270 | in | 
| 271 | isar_steps outer (SOME l) (step :: accum) infs | |
| 272 | end | |
| 273 | and isar_proof outer fix assms lems infs = | |
| 274 | Proof (fix, assms, lems @ isar_steps outer NONE [] infs) | |
| 275 | ||
| 55222 | 276 | val trace = Config.get ctxt trace | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 277 | |
| 55256 | 278 | val canonical_isar_proof = | 
| 51145 | 279 | refute_graph | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 280 | |> 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 | 281 | |> redirect_graph axioms tainted bot | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 282 | |> trace ? tap (tracing o prefix "Direct proof: " o string_of_direct_proof) | 
| 54754 | 283 | |> isar_proof true params assms lems | 
| 55213 | 284 | |> postprocess_isar_proof_remove_unreferenced_steps I | 
| 285 | |> relabel_isar_proof_canonically | |
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 286 | |
| 55286 | 287 | val ctxt = ctxt |> enrich_context_with_local_facts canonical_isar_proof | 
| 55256 | 288 | |
| 55260 | 289 | val preplay_data = Unsynchronized.ref Canonical_Label_Tab.empty | 
| 290 | ||
| 55264 | 291 | val _ = fold_isar_steps (fn meth => | 
| 55452 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55329diff
changeset | 292 | K (set_preplay_outcomes_of_isar_step ctxt debug preplay_timeout preplay_data meth [])) | 
| 55260 | 293 | (steps_of_isar_proof canonical_isar_proof) () | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 294 | |
| 55223 
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
 blanchet parents: 
55222diff
changeset | 295 | fun str_of_preplay_outcome outcome = | 
| 
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
 blanchet parents: 
55222diff
changeset | 296 | 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 | 297 | fun str_of_meth l meth = | 
| 55260 | 298 | string_of_proof_method meth ^ " " ^ | 
| 55266 | 299 | 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 | 300 | fun comment_of l = map (str_of_meth l) #> commas | 
| 55222 | 301 | |
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 302 | fun trace_isar_proof label proof = | 
| 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 303 | if trace then | 
| 55299 | 304 | tracing (timestamp () ^ "\n" ^ label ^ ":\n\n" ^ | 
| 56097 | 305 | string_of_isar_proof ctxt subgoal subgoal_count | 
| 306 | (comment_isar_proof comment_of proof) ^ "\n") | |
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 307 | else | 
| 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 308 | () | 
| 54754 | 309 | |
| 55299 | 310 | fun comment_of l (meth :: _) = | 
| 311 | (case (verbose, | |
| 312 | Lazy.force (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth)) of | |
| 313 | (false, Played _) => "" | |
| 314 | | (_, outcome) => string_of_play_outcome outcome) | |
| 315 | ||
| 54828 | 316 | val (play_outcome, isar_proof) = | 
| 55256 | 317 | canonical_isar_proof | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 318 | |> tap (trace_isar_proof "Original") | 
| 55452 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55329diff
changeset | 319 | |> compress_isar_proof ctxt debug compress_isar preplay_timeout preplay_data | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 320 | |> tap (trace_isar_proof "Compressed") | 
| 55213 | 321 | |> postprocess_isar_proof_remove_unreferenced_steps | 
| 55266 | 322 | (keep_fastest_method_of_isar_step (!preplay_data) | 
| 55452 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55329diff
changeset | 323 | #> minimize ? minimize_isar_step_dependencies ctxt debug preplay_data) | 
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 324 | |> tap (trace_isar_proof "Minimized") | 
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55327diff
changeset | 325 | (* It's not clear whether this is worth the trouble (and if so, "isar_compress" has an | 
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55327diff
changeset | 326 | unnatural semantics): *) | 
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55327diff
changeset | 327 | (* | 
| 55327 
3c7f3122ccdc
do a second phase of proof compression after minimization
 blanchet parents: 
55325diff
changeset | 328 | |> minimize | 
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55327diff
changeset | 329 | ? (compress_isar_proof ctxt compress_isar preplay_timeout preplay_data | 
| 55327 
3c7f3122ccdc
do a second phase of proof compression after minimization
 blanchet parents: 
55325diff
changeset | 330 | #> tap (trace_isar_proof "Compressed again")) | 
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55327diff
changeset | 331 | *) | 
| 55260 | 332 | |> `(preplay_outcome_of_isar_proof (!preplay_data)) | 
| 55325 | 333 | ||> (comment_isar_proof comment_of | 
| 334 | #> chain_isar_proof | |
| 335 | #> kill_useless_labels_in_isar_proof | |
| 336 | #> relabel_isar_proof_nicely) | |
| 49883 | 337 | in | 
| 56097 | 338 | (case string_of_isar_proof ctxt subgoal subgoal_count isar_proof of | 
| 49883 | 339 | "" => | 
| 55213 | 340 | if isar_proofs = SOME true then "\nNo structured proof available (proof too simple)." | 
| 341 | else "" | |
| 55214 
48a347b40629
better tracing + syntactically correct 'metis' calls
 blanchet parents: 
55213diff
changeset | 342 | | isar_text => | 
| 50670 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
50557diff
changeset | 343 | let | 
| 
eaa540986291
properly take the existential closure of skolems
 blanchet parents: 
50557diff
changeset | 344 | val msg = | 
| 51203 | 345 | (if verbose then | 
| 55260 | 346 | let val num_steps = add_isar_steps (steps_of_isar_proof isar_proof) 0 in | 
| 55213 | 347 | [string_of_int num_steps ^ " step" ^ plural_s num_steps] | 
| 348 | end | |
| 51203 | 349 | else | 
| 350 | []) @ | |
| 54828 | 351 | (if do_preplay then [string_of_play_outcome play_outcome] else []) | 
| 50277 | 352 | in | 
| 54507 | 353 |             "\n\nStructured proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^
 | 
| 354 | Active.sendback_markup [Markup.padding_command] isar_text | |
| 54754 | 355 | end) | 
| 49883 | 356 | end | 
| 54760 
a1ac3eaa0d11
generate proper succedent for cases with trivial branches
 blanchet parents: 
54759diff
changeset | 357 | |
| 55168 
948e8b7ea82f
correctly handle exceptions arising from (experimental) Isar proof code
 blanchet parents: 
54838diff
changeset | 358 | val one_line_proof = one_line_proof_text 0 one_line_params | 
| 49883 | 359 | val isar_proof = | 
| 360 | if debug then | |
| 361 | isar_proof_of () | |
| 54754 | 362 | else | 
| 363 | (case try isar_proof_of () of | |
| 364 | SOME s => s | |
| 365 | | NONE => | |
| 56097 | 366 | if isar_proofs = SOME true then "\nWarning: Isar proof construction failed." else "") | 
| 49883 | 367 | in one_line_proof ^ isar_proof end | 
| 368 | ||
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 369 | fun isar_proof_would_be_a_good_idea smt_proofs (meth, play) = | 
| 54824 | 370 | (case play of | 
| 56081 | 371 | Played _ => meth = SMT2_Method andalso smt_proofs <> SOME true | 
| 56093 | 372 | | Play_Timed_Out time => Time.> (time, Time.zeroTime) | 
| 373 | | 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 | 374 | |
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 375 | fun proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained | 
| 55256 | 376 | (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 | 377 | (if isar_proofs = SOME true orelse | 
| 55297 
1dfcd49f5dcb
renamed 'smt' option 'smt_proofs' to avoid clash with 'smt' prover
 blanchet parents: 
55296diff
changeset | 378 | (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 | 379 | isar_proof_text ctxt debug isar_proofs smt_proofs isar_params | 
| 49883 | 380 | else | 
| 53052 
a0db255af8c5
sledgehammer sendback always uses Markup.padding_command: sensible default for most practical applications -- old-style in-line replacement is superseded by auto mode or panel;
 wenzelm parents: 
53047diff
changeset | 381 | one_line_proof_text num_chained) one_line_params | 
| 49883 | 382 | |
| 383 | end; |