| author | wenzelm | 
| Sun, 22 Jul 2018 20:01:03 +0200 | |
| changeset 68681 | 14167c321d22 | 
| parent 67379 | c2dfc510a38c | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 44651 
5d6a11e166cf
renamed "Metis_Tactics" to "Metis_Tactic", now that there is only one Metis tactic ("metisFT" is legacy)
 blanchet parents: 
44634diff
changeset | 1 | (* Title: HOL/Tools/Metis/metis_tactic.ML | 
| 38027 | 2 | Author: Kong W. Susanto, Cambridge University Computer Laboratory | 
| 3 | Author: Lawrence C. Paulson, Cambridge University Computer Laboratory | |
| 4 | Author: Jasmin Blanchette, TU Muenchen | |
| 23442 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 5 | Copyright Cambridge University 2007 | 
| 23447 | 6 | |
| 29266 | 7 | HOL setup for the Metis prover. | 
| 23442 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 8 | *) | 
| 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 9 | |
| 44651 
5d6a11e166cf
renamed "Metis_Tactics" to "Metis_Tactic", now that there is only one Metis tactic ("metisFT" is legacy)
 blanchet parents: 
44634diff
changeset | 10 | signature METIS_TACTIC = | 
| 23442 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 11 | sig | 
| 39979 
b13515940b53
added "trace_meson" configuration option, replacing old-fashioned reference
 blanchet parents: 
39978diff
changeset | 12 | val trace : bool Config.T | 
| 40665 
1a65f0c74827
added "verbose" option to Metis to shut up its warnings if necessary
 blanchet parents: 
40262diff
changeset | 13 | val verbose : bool Config.T | 
| 50705 
0e943b33d907
use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
 blanchet parents: 
50694diff
changeset | 14 | val new_skolem : bool Config.T | 
| 47039 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 15 | val advisory_simp : bool Config.T | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 16 | val metis_tac_unused : string list -> string -> Proof.context -> thm list -> int -> thm -> | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 17 | thm list * thm Seq.seq | 
| 54756 | 18 | val metis_tac : string list -> string -> Proof.context -> thm list -> int -> tactic | 
| 61302 | 19 | val metis_method : (string list option * string option) * thm list -> Proof.context -> thm list -> | 
| 62219 
dbac573b27e7
preplaying of 'smt' and 'metis' more in sync with actual method
 blanchet parents: 
61841diff
changeset | 20 | tactic | 
| 45521 | 21 | val metis_lam_transs : string list | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 22 | val parse_metis_options : (string list option * string option) parser | 
| 23442 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 23 | end | 
| 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 24 | |
| 44651 
5d6a11e166cf
renamed "Metis_Tactics" to "Metis_Tactic", now that there is only one Metis tactic ("metisFT" is legacy)
 blanchet parents: 
44634diff
changeset | 25 | structure Metis_Tactic : METIS_TACTIC = | 
| 23442 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 26 | struct | 
| 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 27 | |
| 46320 | 28 | open ATP_Problem_Generate | 
| 29 | open ATP_Proof_Reconstruct | |
| 30 | open Metis_Generate | |
| 39497 
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
 blanchet parents: 
39494diff
changeset | 31 | open Metis_Reconstruct | 
| 35826 | 32 | |
| 54756 | 33 | val new_skolem = Attrib.setup_config_bool @{binding metis_new_skolem} (K false)
 | 
| 34 | val advisory_simp = Attrib.setup_config_bool @{binding metis_advisory_simp} (K true)
 | |
| 23442 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 35 | |
| 43134 
0c82e00ba63e
make sure no warnings are given for polymorphic facts where we use a monomorphic instance
 blanchet parents: 
43133diff
changeset | 36 | (* Designed to work also with monomorphic instances of polymorphic theorems. *) | 
| 60358 | 37 | fun have_common_thm ctxt ths1 ths2 = | 
| 38 | exists (member (Term.aconv_untyped o apply2 Thm.prop_of) ths1) | |
| 39 | (map (Meson.make_meta_clause ctxt) ths2) | |
| 23442 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 40 | |
| 32956 | 41 | (*Determining which axiom clauses are actually used*) | 
| 39419 
c9accfd621a5
"Metis." -> "Metis_" to reflect change in "metis.ML"
 blanchet parents: 
39376diff
changeset | 42 | fun used_axioms axioms (th, Metis_Proof.Axiom _) = SOME (lookth axioms th) | 
| 43128 | 43 | | used_axioms _ _ = NONE | 
| 24855 | 44 | |
| 43129 | 45 | (* Lightweight predicate type information comes in two flavors, "t = t'" and | 
| 46 | "t => t'", where "t" and "t'" are the same term modulo type tags. | |
| 47 | In Isabelle, type tags are stripped away, so we are left with "t = t" or | |
| 43159 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43136diff
changeset | 48 | "t => t". Type tag idempotence is also handled this way. *) | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51717diff
changeset | 49 | fun reflexive_or_trivial_of_metis ctxt type_enc sym_tab concealed mth = | 
| 59632 | 50 | (case hol_clause_of_metis ctxt type_enc sym_tab concealed mth of | 
| 51 |     Const (@{const_name HOL.eq}, _) $ _ $ t =>
 | |
| 52 | let | |
| 53 | val ct = Thm.cterm_of ctxt t | |
| 54 | val cT = Thm.ctyp_of_cterm ct | |
| 60801 | 55 | in refl |> Thm.instantiate' [SOME cT] [SOME ct] end | 
| 59632 | 56 |   | Const (@{const_name disj}, _) $ t1 $ t2 =>
 | 
| 57 | (if can HOLogic.dest_not t1 then t2 else t1) | |
| 58 | |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt |> Thm.trivial | |
| 59 | | _ => raise Fail "expected reflexive or trivial clause") | |
| 60358 | 60 | |> Meson.make_meta_clause ctxt | 
| 43129 | 61 | |
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51717diff
changeset | 62 | fun lam_lifted_of_metis ctxt type_enc sym_tab concealed mth = | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45508diff
changeset | 63 | let | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 64 |     val tac = rewrite_goals_tac ctxt @{thms lambda_def [abs_def]} THEN resolve_tac ctxt [refl] 1
 | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51717diff
changeset | 65 | val t = hol_clause_of_metis ctxt type_enc sym_tab concealed mth | 
| 59632 | 66 | val ct = Thm.cterm_of ctxt (HOLogic.mk_Trueprop t) | 
| 60358 | 67 | in Goal.prove_internal ctxt [] ct (K tac) |> Meson.make_meta_clause ctxt end | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45508diff
changeset | 68 | |
| 45570 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 69 | fun add_vars_and_frees (t $ u) = fold (add_vars_and_frees) [t, u] | 
| 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 70 | | add_vars_and_frees (Abs (_, _, t)) = add_vars_and_frees t | 
| 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 71 | | add_vars_and_frees (t as Var _) = insert (op =) t | 
| 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 72 | | add_vars_and_frees (t as Free _) = insert (op =) t | 
| 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 73 | | add_vars_and_frees _ = I | 
| 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 74 | |
| 45569 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 blanchet parents: 
45568diff
changeset | 75 | fun introduce_lam_wrappers ctxt th = | 
| 59582 | 76 | if Meson_Clausify.is_quasi_lambda_free (Thm.prop_of th) then | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45508diff
changeset | 77 | th | 
| 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45508diff
changeset | 78 | else | 
| 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45508diff
changeset | 79 | let | 
| 45570 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 80 | fun conv first ctxt ct = | 
| 59582 | 81 | if Meson_Clausify.is_quasi_lambda_free (Thm.term_of ct) then | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45508diff
changeset | 82 | Thm.reflexive ct | 
| 57408 | 83 | else | 
| 59582 | 84 | (case Thm.term_of ct of | 
| 57408 | 85 | Abs (_, _, u) => | 
| 86 | if first then | |
| 87 | (case add_vars_and_frees u [] of | |
| 88 | [] => | |
| 89 | Conv.abs_conv (conv false o snd) ctxt ct | |
| 60362 | 90 |                 |> (fn th => Meson.first_order_resolve ctxt th @{thm Metis.eq_lambdaI})
 | 
| 57408 | 91 | | v :: _ => | 
| 59582 | 92 | Abs (Name.uu, fastype_of v, abstract_over (v, Thm.term_of ct)) $ v | 
| 59632 | 93 | |> Thm.cterm_of ctxt | 
| 57408 | 94 | |> Conv.comb_conv (conv true ctxt)) | 
| 95 | else | |
| 45570 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 96 | Conv.abs_conv (conv false o snd) ctxt ct | 
| 57408 | 97 |           | Const (@{const_name Meson.skolem}, _) $ _ => Thm.reflexive ct
 | 
| 98 | | _ => Conv.comb_conv (conv true ctxt) ct) | |
| 59582 | 99 | val eq_th = conv true ctxt (Thm.cprop_of th) | 
| 45570 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 100 | (* We replace the equation's left-hand side with a beta-equivalent term | 
| 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 101 | so that "Thm.equal_elim" works below. *) | 
| 59582 | 102 | val t0 $ _ $ t2 = Thm.prop_of eq_th | 
| 59632 | 103 | val eq_ct = t0 $ Thm.prop_of th $ t2 |> Thm.cterm_of ctxt | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 104 | val eq_th' = Goal.prove_internal ctxt [] eq_ct (K (resolve_tac ctxt [eq_th] 1)) | 
| 45570 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 105 | in Thm.equal_elim eq_th' th end | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45508diff
changeset | 106 | |
| 47039 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 107 | fun clause_params ordering = | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 108 |   {ordering = ordering,
 | 
| 44492 
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
 blanchet parents: 
44411diff
changeset | 109 | orderLiterals = Metis_Clause.UnsignedLiteralOrder, | 
| 39450 
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
 blanchet parents: 
39419diff
changeset | 110 | orderTerms = true} | 
| 47039 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 111 | fun active_params ordering = | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 112 |   {clause = clause_params ordering,
 | 
| 39450 
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
 blanchet parents: 
39419diff
changeset | 113 | prefactor = #prefactor Metis_Active.default, | 
| 
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
 blanchet parents: 
39419diff
changeset | 114 | postfactor = #postfactor Metis_Active.default} | 
| 
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
 blanchet parents: 
39419diff
changeset | 115 | val waiting_params = | 
| 
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
 blanchet parents: 
39419diff
changeset | 116 |   {symbolsWeight = 1.0,
 | 
| 47047 
10bece4ac87e
more conservative Metis defaults, for backward compatiblity (as illustrated by one "metis" call in "Auth/KerberosV")
 blanchet parents: 
47045diff
changeset | 117 | variablesWeight = 0.05, | 
| 
10bece4ac87e
more conservative Metis defaults, for backward compatiblity (as illustrated by one "metis" call in "Auth/KerberosV")
 blanchet parents: 
47045diff
changeset | 118 | literalsWeight = 0.01, | 
| 39450 
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
 blanchet parents: 
39419diff
changeset | 119 | models = []} | 
| 47039 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 120 | fun resolution_params ordering = | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 121 |   {active = active_params ordering, waiting = waiting_params}
 | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 122 | |
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 123 | fun kbo_advisory_simp_ordering ord_info = | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 124 | let | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 125 | fun weight (m, _) = | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 126 | AList.lookup (op =) ord_info (Metis_Name.toString m) |> the_default 1 | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 127 | fun precedence p = | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58839diff
changeset | 128 | (case int_ord (apply2 weight p) of | 
| 47039 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 129 | EQUAL => #precedence Metis_KnuthBendixOrder.default p | 
| 57408 | 130 | | ord => ord) | 
| 47039 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 131 |   in {weight = weight, precedence = precedence} end
 | 
| 37573 | 132 | |
| 55285 | 133 | fun metis_call type_enc lam_trans = | 
| 134 | let | |
| 135 | val type_enc = | |
| 136 | (case AList.find (fn (enc, encs) => enc = hd encs) type_enc_aliases type_enc of | |
| 137 | [alias] => alias | |
| 138 | | _ => type_enc) | |
| 139 | val opts = | |
| 140 | [] |> type_enc <> partial_typesN ? cons type_enc | |
| 141 | |> lam_trans <> default_metis_lam_trans ? cons lam_trans | |
| 142 |   in metisN ^ (if null opts then "" else " (" ^ commas opts ^ ")") end
 | |
| 143 | ||
| 50875 
bfb626265782
less brutal Metis failure -- the brutality was accidentally introduced by df8ae0590be2
 blanchet parents: 
50705diff
changeset | 144 | exception METIS_UNPROVABLE of unit | 
| 
bfb626265782
less brutal Metis failure -- the brutality was accidentally introduced by df8ae0590be2
 blanchet parents: 
50705diff
changeset | 145 | |
| 37516 
c81c86bfc18a
have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
 blanchet parents: 
37509diff
changeset | 146 | (* Main function to start Metis proof and reconstruction *) | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 147 | fun FOL_SOLVE unused (type_enc :: fallback_type_encs) lam_trans ctxt cls ths0 = | 
| 42361 | 148 | let val thy = Proof_Context.theory_of ctxt | 
| 50705 
0e943b33d907
use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
 blanchet parents: 
50694diff
changeset | 149 | val new_skolem = | 
| 
0e943b33d907
use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
 blanchet parents: 
50694diff
changeset | 150 | Config.get ctxt new_skolem orelse null (Meson.choice_theorems thy) | 
| 46365 | 151 | val do_lams = | 
| 152 | (lam_trans = liftingN orelse lam_trans = lam_liftingN) | |
| 153 | ? introduce_lam_wrappers ctxt | |
| 35826 | 154 | val th_cls_pairs = | 
| 39894 
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
 blanchet parents: 
39892diff
changeset | 155 | map2 (fn j => fn th => | 
| 
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
 blanchet parents: 
39892diff
changeset | 156 | (Thm.get_name_hint th, | 
| 45570 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 157 | th |> Drule.eta_contraction_rule | 
| 57263 | 158 | |> Meson_Clausify.cnf_axiom ctxt new_skolem (lam_trans = combsN) j | 
| 45570 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 159 | ||> map do_lams)) | 
| 39894 
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
 blanchet parents: 
39892diff
changeset | 160 | (0 upto length ths0 - 1) ths0 | 
| 43092 
93ec303e1917
more work on new metis that exploits the powerful new type encodings
 blanchet parents: 
43091diff
changeset | 161 | val ths = maps (snd o snd) th_cls_pairs | 
| 39938 
0a2091f86eb4
fixed two bugs in new skolemizer: instantiations now take types into consideration, and rotate_tac is given the proper offset
 blanchet parents: 
39937diff
changeset | 162 | val dischargers = map (fst o snd) th_cls_pairs | 
| 45570 
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
 blanchet parents: 
45569diff
changeset | 163 | val cls = cls |> map (Drule.eta_contraction_rule #> do_lams) | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 164 | val _ = trace_msg ctxt (K "FOL_SOLVE: CONJECTURE CLAUSES") | 
| 67379 | 165 | val _ = List.app (fn th => trace_msg ctxt (fn () => Thm.string_of_thm ctxt th)) cls | 
| 44411 
e3629929b171
change Metis's default settings if type information axioms are generated
 blanchet parents: 
44408diff
changeset | 166 | val _ = trace_msg ctxt (fn () => "type_enc = " ^ type_enc) | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51717diff
changeset | 167 | val type_enc = type_enc_of_string Strict type_enc | 
| 47039 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 168 | val (sym_tab, axioms, ord_info, concealed) = | 
| 57263 | 169 | generate_metis_problem ctxt type_enc lam_trans cls ths | 
| 43159 
29b55f292e0b
added support for helpers in new Metis, so far only for polymorphic type encodings
 blanchet parents: 
43136diff
changeset | 170 | fun get_isa_thm mth Isa_Reflexive_or_Trivial = | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51717diff
changeset | 171 | reflexive_or_trivial_of_metis ctxt type_enc sym_tab concealed mth | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45508diff
changeset | 172 | | get_isa_thm mth Isa_Lambda_Lifted = | 
| 52031 
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
 blanchet parents: 
51717diff
changeset | 173 | lam_lifted_of_metis ctxt type_enc sym_tab concealed mth | 
| 45569 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 blanchet parents: 
45568diff
changeset | 174 | | get_isa_thm _ (Isa_Raw ith) = ith | 
| 
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
 blanchet parents: 
45568diff
changeset | 175 | val axioms = axioms |> map (fn (mth, ith) => (mth, get_isa_thm mth ith)) | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 176 | val _ = trace_msg ctxt (K "ISABELLE CLAUSES") | 
| 67379 | 177 | val _ = List.app (fn (_, ith) => trace_msg ctxt (fn () => Thm.string_of_thm ctxt ith)) axioms | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 178 | val _ = trace_msg ctxt (K "METIS CLAUSES") | 
| 67379 | 179 | val _ = List.app (fn (mth, _) => trace_msg ctxt (fn () => Metis_Thm.toString mth)) axioms | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 180 | val _ = trace_msg ctxt (K "START METIS PROVE PROCESS") | 
| 47039 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 181 | val ordering = | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 182 | if Config.get ctxt advisory_simp then | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 183 | kbo_advisory_simp_ordering (ord_info ()) | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 184 | else | 
| 
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
 blanchet parents: 
47015diff
changeset | 185 | Metis_KnuthBendixOrder.default | 
| 50875 
bfb626265782
less brutal Metis failure -- the brutality was accidentally introduced by df8ae0590be2
 blanchet parents: 
50705diff
changeset | 186 | fun fall_back () = | 
| 
bfb626265782
less brutal Metis failure -- the brutality was accidentally introduced by df8ae0590be2
 blanchet parents: 
50705diff
changeset | 187 | (verbose_warning ctxt | 
| 55257 | 188 |            ("Falling back on " ^ quote (metis_call (hd fallback_type_encs) lam_trans) ^ "...");
 | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 189 | FOL_SOLVE unused fallback_type_encs lam_trans ctxt cls ths0) | 
| 32956 | 190 | in | 
| 59582 | 191 |     (case filter (fn t => Thm.prop_of t aconv @{prop False}) cls of
 | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 192 |        false_th :: _ => [false_th RS @{thm FalseE}]
 | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 193 | | [] => | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 194 | (case Metis_Resolution.loop (Metis_Resolution.new (resolution_params ordering) | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 195 |          {axioms = axioms |> map fst, conjecture = []}) of
 | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 196 | Metis_Resolution.Contradiction mth => | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 197 | let | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 198 | val _ = trace_msg ctxt (fn () => "METIS RECONSTRUCTION START: " ^ Metis_Thm.toString mth) | 
| 59582 | 199 | val ctxt' = fold Variable.declare_constraints (map Thm.prop_of cls) ctxt | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 200 | (*add constraints arising from converting goal to clause form*) | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 201 | val proof = Metis_Proof.proof mth | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 202 | val result = fold (replay_one_inference ctxt' type_enc concealed sym_tab) proof axioms | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 203 | val used = map_filter (used_axioms axioms) proof | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 204 | val _ = trace_msg ctxt (K "METIS COMPLETED; clauses actually used:") | 
| 67379 | 205 | val _ = List.app (fn th => trace_msg ctxt (fn () => Thm.string_of_thm ctxt th)) used | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 206 | val (used_th_cls_pairs, unused_th_cls_pairs) = | 
| 60358 | 207 | List.partition (have_common_thm ctxt used o snd o snd) th_cls_pairs | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 208 | val unused_ths = maps (snd o snd) unused_th_cls_pairs | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 209 | val unused_names = map fst unused_th_cls_pairs | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 210 | in | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 211 | unused := unused_ths; | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 212 | if not (null unused_names) then | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 213 |            verbose_warning ctxt ("Unused theorems: " ^ commas_quote unused_names)
 | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 214 | else | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 215 | (); | 
| 60358 | 216 | if not (null cls) andalso not (have_common_thm ctxt used cls) then | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 217 | verbose_warning ctxt "The assumptions are inconsistent" | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 218 | else | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 219 | (); | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 220 | (case result of | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 221 | (_, ith) :: _ => | 
| 61268 | 222 | (trace_msg ctxt (fn () => "Success: " ^ Thm.string_of_thm ctxt ith); | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 223 | [discharge_skolem_premises ctxt dischargers ith]) | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 224 | | _ => (trace_msg ctxt (K "Metis: No result"); [])) | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 225 | end | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 226 | | Metis_Resolution.Satisfiable _ => | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 227 | (trace_msg ctxt (K "Metis: No first-order proof with the supplied lemmas"); | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 228 | raise METIS_UNPROVABLE ())) | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 229 | handle METIS_UNPROVABLE () => if null fallback_type_encs then [] else fall_back () | 
| 50875 
bfb626265782
less brutal Metis failure -- the brutality was accidentally introduced by df8ae0590be2
 blanchet parents: 
50705diff
changeset | 230 | | METIS_RECONSTRUCT (loc, msg) => | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 231 | if null fallback_type_encs then | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 232 |              (verbose_warning ctxt ("Failed to replay Metis proof\n" ^ loc ^ ": " ^ msg); [])
 | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 233 | else | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 234 | fall_back ()) | 
| 42733 
01ef1c3d9cfd
more robust exception handling in Metis (also works if there are several subgoals)
 blanchet parents: 
42650diff
changeset | 235 | end | 
| 23442 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 236 | |
| 45508 | 237 | fun neg_clausify ctxt combinators = | 
| 38028 | 238 | single | 
| 43964 
9338aa218f09
thread proper context through, to make sure that "using [[meson_max_clauses = 200]]" is not ignored when clausifying the conjecture
 blanchet parents: 
43963diff
changeset | 239 | #> Meson.make_clauses_unsorted ctxt | 
| 55236 | 240 | #> combinators ? map (Meson_Clausify.introduce_combinators_in_theorem ctxt) | 
| 38028 | 241 | #> Meson.finish_cnf | 
| 242 | ||
| 39269 
c2795d8a2461
use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
 blanchet parents: 
39267diff
changeset | 243 | fun preskolem_tac ctxt st0 = | 
| 
c2795d8a2461
use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
 blanchet parents: 
39267diff
changeset | 244 | (if exists (Meson.has_too_many_clauses ctxt) | 
| 59582 | 245 | (Logic.prems_of_goal (Thm.prop_of st0) 1) then | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50875diff
changeset | 246 |      Simplifier.full_simp_tac (Meson_Clausify.ss_only @{thms not_all not_ex} ctxt) 1
 | 
| 55239 | 247 | THEN CNF.cnfx_rewrite_tac ctxt 1 | 
| 39269 
c2795d8a2461
use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
 blanchet parents: 
39267diff
changeset | 248 | else | 
| 
c2795d8a2461
use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
 blanchet parents: 
39267diff
changeset | 249 | all_tac) st0 | 
| 
c2795d8a2461
use definitional CNF for the goal if at least one of the premisses would lead to too many clauses in Meson
 blanchet parents: 
39267diff
changeset | 250 | |
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 251 | fun metis_tac_unused type_encs0 lam_trans ctxt ths i st0 = | 
| 37926 
e6ff246c0cdb
renamings + only need second component of name pool to reconstruct proofs
 blanchet parents: 
37925diff
changeset | 252 | let | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 253 | val unused = Unsynchronized.ref [] | 
| 55520 | 254 | val type_encs = if null type_encs0 then partial_type_encs else type_encs0 | 
| 39978 
11bfb7e7cc86
added "trace_metis" configuration option, replacing old-fashioned references
 blanchet parents: 
39964diff
changeset | 255 | val _ = trace_msg ctxt (fn () => | 
| 61268 | 256 | "Metis called with theorems\n" ^ cat_lines (map (Thm.string_of_thm ctxt) ths)) | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 257 | val type_encs = type_encs |> maps unalias_type_enc | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 258 | val combs = (lam_trans = combsN) | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59058diff
changeset | 259 | fun tac clause = resolve_tac ctxt (FOL_SOLVE unused type_encs lam_trans ctxt clause ths) 1 | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 260 | val seq = Meson.MESON (preskolem_tac ctxt) (maps (neg_clausify ctxt combs)) tac ctxt i st0 | 
| 32956 | 261 | in | 
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 262 | (!unused, seq) | 
| 32956 | 263 | end | 
| 23442 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 264 | |
| 55521 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 265 | fun metis_tac type_encs lam_trans ctxt ths i = snd o metis_tac_unused type_encs lam_trans ctxt ths i | 
| 
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
 blanchet parents: 
55520diff
changeset | 266 | |
| 55520 | 267 | (* Whenever "X" has schematic type variables, we treat "using X by metis" as "by (metis X)" to | 
| 268 | prevent "Subgoal.FOCUS" from freezing the type variables. We don't do it for nonschematic facts | |
| 269 | "X" because this breaks a few proofs (in the rare and subtle case where a proof relied on | |
| 270 | extensionality not being applied) and brings few benefits. *) | |
| 59582 | 271 | val has_tvar = exists_type (exists_subtype (fn TVar _ => true | _ => false)) o Thm.prop_of | 
| 43034 
18259246abb5
try both "metis" and (on failure) "metisFT" in replay
 blanchet parents: 
42847diff
changeset | 272 | |
| 55315 | 273 | fun metis_method ((override_type_encs, lam_trans), ths) ctxt facts = | 
| 55520 | 274 | let val (schem_facts, nonschem_facts) = List.partition has_tvar facts in | 
| 61841 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61302diff
changeset | 275 | HEADGOAL (Method.insert_tac ctxt nonschem_facts THEN' | 
| 55520 | 276 | CHANGED_PROP o metis_tac (these override_type_encs) | 
| 277 | (the_default default_metis_lam_trans lam_trans) ctxt (schem_facts @ ths)) | |
| 43099 | 278 | end | 
| 43100 | 279 | |
| 46365 | 280 | val metis_lam_transs = [hide_lamsN, liftingN, combsN] | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 281 | |
| 45578 | 282 | fun set_opt _ x NONE = SOME x | 
| 283 | | set_opt get x (SOME x0) = | |
| 55523 
9429e7b5b827
removed final periods in messages for proof methods
 blanchet parents: 
55521diff
changeset | 284 |     error ("Cannot specify both " ^ quote (get x0) ^ " and " ^ quote (get x))
 | 
| 54756 | 285 | |
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 286 | fun consider_opt s = | 
| 54756 | 287 | if member (op =) metis_lam_transs s then apsnd (set_opt I s) else apfst (set_opt hd [s]) | 
| 45514 | 288 | |
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 289 | val parse_metis_options = | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 290 | Scan.optional | 
| 58831 
aa8cf5eed06e
proper syntax categery "name" -- as usual and as documented;
 wenzelm parents: 
58818diff
changeset | 291 |       (Args.parens (Args.name -- Scan.option (@{keyword ","} |-- Args.name))
 | 
| 45519 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 292 | >> (fn (s, s') => | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 293 | (NONE, NONE) |> consider_opt s | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 294 | |> (case s' of SOME s' => consider_opt s' | _ => I))) | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 295 | (NONE, NONE) | 
| 
cd6e78cb6ee8
make metis reconstruction handling more flexible
 blanchet parents: 
45514diff
changeset | 296 | |
| 58818 | 297 | val _ = | 
| 298 | Theory.setup | |
| 299 |     (Method.setup @{binding metis}
 | |
| 300 | (Scan.lift parse_metis_options -- Attrib.thms >> (METHOD oo metis_method)) | |
| 301 | "Metis for FOL and HOL problems") | |
| 23442 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 302 | |
| 
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
 wenzelm parents: diff
changeset | 303 | end; |