author | wenzelm |
Wed, 25 Jun 2025 16:35:25 +0200 | |
changeset 82768 | 8f866fd6fae1 |
parent 82574 | 318526b74e6f |
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:
44634
diff
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:
44634
diff
changeset
|
10 |
signature METIS_TACTIC = |
23442
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff
changeset
|
11 |
sig |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
12 |
type inst |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
13 |
|
39979
b13515940b53
added "trace_meson" configuration option, replacing old-fashioned reference
blanchet
parents:
39978
diff
changeset
|
14 |
val trace : bool Config.T |
40665
1a65f0c74827
added "verbose" option to Metis to shut up its warnings if necessary
blanchet
parents:
40262
diff
changeset
|
15 |
val verbose : bool Config.T |
81746 | 16 |
val instantiate : bool Config.T |
17 |
val instantiate_undefined : 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:
50694
diff
changeset
|
18 |
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:
47015
diff
changeset
|
19 |
val advisory_simp : bool Config.T |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
20 |
val pretty_name_inst : Proof.context -> string * inst -> Pretty.T |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
21 |
val string_of_name_inst : Proof.context -> string * inst -> string |
54756 | 22 |
val metis_tac : string list -> string -> Proof.context -> thm list -> int -> tactic |
61302 | 23 |
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:
61841
diff
changeset
|
24 |
tactic |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
25 |
val metis_infer_thm_insts : (string list option * string option) * thm list -> Proof.context -> |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
26 |
thm list -> int -> thm -> (thm * inst) list option |
45521 | 27 |
val metis_lam_transs : string list |
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
28 |
val parse_metis_options : (string list option * string option) parser |
23442
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff
changeset
|
29 |
end |
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff
changeset
|
30 |
|
44651
5d6a11e166cf
renamed "Metis_Tactics" to "Metis_Tactic", now that there is only one Metis tactic ("metisFT" is legacy)
blanchet
parents:
44634
diff
changeset
|
31 |
structure Metis_Tactic : METIS_TACTIC = |
23442
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff
changeset
|
32 |
struct |
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff
changeset
|
33 |
|
46320 | 34 |
open ATP_Problem_Generate |
35 |
open ATP_Proof_Reconstruct |
|
36 |
open Metis_Generate |
|
39497
fa16349939b7
complete refactoring of Metis along the lines of Sledgehammer
blanchet
parents:
39494
diff
changeset
|
37 |
open Metis_Reconstruct |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
38 |
open Metis_Instantiations |
35826 | 39 |
|
69593 | 40 |
val new_skolem = Attrib.setup_config_bool \<^binding>\<open>metis_new_skolem\<close> (K false) |
41 |
val advisory_simp = Attrib.setup_config_bool \<^binding>\<open>metis_advisory_simp\<close> (K true) |
|
23442
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff
changeset
|
42 |
|
32956 | 43 |
(*Determining which axiom clauses are actually used*) |
39419
c9accfd621a5
"Metis." -> "Metis_" to reflect change in "metis.ML"
blanchet
parents:
39376
diff
changeset
|
44 |
fun used_axioms axioms (th, Metis_Proof.Axiom _) = SOME (lookth axioms th) |
43128 | 45 |
| used_axioms _ _ = NONE |
24855 | 46 |
|
43129 | 47 |
(* Lightweight predicate type information comes in two flavors, "t = t'" and |
48 |
"t => t'", where "t" and "t'" are the same term modulo type tags. |
|
49 |
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:
43136
diff
changeset
|
50 |
"t => t". Type tag idempotence is also handled this way. *) |
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51717
diff
changeset
|
51 |
fun reflexive_or_trivial_of_metis ctxt type_enc sym_tab concealed mth = |
59632 | 52 |
(case hol_clause_of_metis ctxt type_enc sym_tab concealed mth of |
74347 | 53 |
\<^Const_>\<open>HOL.eq _ for _ t\<close> => |
70489 | 54 |
let |
55 |
val ct = Thm.cterm_of ctxt t |
|
56 |
val cT = Thm.ctyp_of_cterm ct |
|
57 |
in refl |> Thm.instantiate' [SOME cT] [SOME ct] end |
|
74347 | 58 |
| \<^Const_>\<open>disj for t1 t2\<close> => |
70489 | 59 |
(if can HOLogic.dest_not t1 then t2 else t1) |
60 |
|> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt |> Thm.trivial |
|
59632 | 61 |
| _ => raise Fail "expected reflexive or trivial clause") |
60358 | 62 |
|> Meson.make_meta_clause ctxt |
43129 | 63 |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51717
diff
changeset
|
64 |
fun lam_lifted_of_metis ctxt type_enc sym_tab concealed mth = |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45508
diff
changeset
|
65 |
let |
74530
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
wenzelm
parents:
74526
diff
changeset
|
66 |
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:
51717
diff
changeset
|
67 |
val t = hol_clause_of_metis ctxt type_enc sym_tab concealed mth |
59632 | 68 |
val ct = Thm.cterm_of ctxt (HOLogic.mk_Trueprop t) |
70489 | 69 |
in |
74530
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
wenzelm
parents:
74526
diff
changeset
|
70 |
Goal.prove_internal ctxt [] ct (K tac) |
70489 | 71 |
|> Meson.make_meta_clause ctxt |
72 |
end |
|
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45508
diff
changeset
|
73 |
|
45570
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
blanchet
parents:
45569
diff
changeset
|
74 |
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:
45569
diff
changeset
|
75 |
| 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:
45569
diff
changeset
|
76 |
| 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:
45569
diff
changeset
|
77 |
| 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:
45569
diff
changeset
|
78 |
| 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:
45569
diff
changeset
|
79 |
|
45569
eb30a5490543
wrap lambdas earlier, to get more control over beta/eta
blanchet
parents:
45568
diff
changeset
|
80 |
fun introduce_lam_wrappers ctxt th = |
70489 | 81 |
if Meson_Clausify.is_quasi_lambda_free (Thm.prop_of th) then th |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45508
diff
changeset
|
82 |
else |
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45508
diff
changeset
|
83 |
let |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
84 |
(* increment indexes to prevent (type) variable clashes *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
85 |
fun resolve_lambdaI th = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
86 |
Meson.first_order_resolve ctxt th |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
87 |
(Thm.incr_indexes (Thm.maxidx_of th + 1) @{thm Metis.eq_lambdaI}) |
45570
6d95a66cce00
pull variables (Var) out of lambdas, so that the Isabelle theorems closely mirror the Metis lambda-lifted ones
blanchet
parents:
45569
diff
changeset
|
88 |
fun conv first ctxt ct = |
70489 | 89 |
if Meson_Clausify.is_quasi_lambda_free (Thm.term_of ct) then Thm.reflexive ct |
57408 | 90 |
else |
59582 | 91 |
(case Thm.term_of ct of |
57408 | 92 |
Abs (_, _, u) => |
70489 | 93 |
if first then |
94 |
(case add_vars_and_frees u [] of |
|
95 |
[] => |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
96 |
Conv.abs_conv (conv false o snd) ctxt ct |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
97 |
|> resolve_lambdaI |
70489 | 98 |
| v :: _ => |
99 |
Abs (Name.uu, fastype_of v, abstract_over (v, Thm.term_of ct)) $ v |
|
100 |
|> Thm.cterm_of ctxt |
|
101 |
|> Conv.comb_conv (conv true ctxt)) |
|
102 |
else Conv.abs_conv (conv false o snd) ctxt ct |
|
74347 | 103 |
| \<^Const_>\<open>Meson.skolem _ for _\<close> => Thm.reflexive ct |
57408 | 104 |
| _ => Conv.comb_conv (conv true ctxt) ct) |
59582 | 105 |
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:
45569
diff
changeset
|
106 |
(* 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:
45569
diff
changeset
|
107 |
so that "Thm.equal_elim" works below. *) |
59582 | 108 |
val t0 $ _ $ t2 = Thm.prop_of eq_th |
59632 | 109 |
val eq_ct = t0 $ Thm.prop_of th $ t2 |> Thm.cterm_of ctxt |
74530
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
wenzelm
parents:
74526
diff
changeset
|
110 |
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:
45569
diff
changeset
|
111 |
in Thm.equal_elim eq_th' th end |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45508
diff
changeset
|
112 |
|
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:
47015
diff
changeset
|
113 |
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:
47015
diff
changeset
|
114 |
{ordering = ordering, |
44492
a330c0608da8
avoid using ":" for anything but systematic type tag annotations, because Hurd's Metis gives it that special semantics
blanchet
parents:
44411
diff
changeset
|
115 |
orderLiterals = Metis_Clause.UnsignedLiteralOrder, |
39450
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents:
39419
diff
changeset
|
116 |
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:
47015
diff
changeset
|
117 |
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:
47015
diff
changeset
|
118 |
{clause = clause_params ordering, |
39450
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents:
39419
diff
changeset
|
119 |
prefactor = #prefactor Metis_Active.default, |
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents:
39419
diff
changeset
|
120 |
postfactor = #postfactor Metis_Active.default} |
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents:
39419
diff
changeset
|
121 |
val waiting_params = |
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents:
39419
diff
changeset
|
122 |
{symbolsWeight = 1.0, |
47047
10bece4ac87e
more conservative Metis defaults, for backward compatiblity (as illustrated by one "metis" call in "Auth/KerberosV")
blanchet
parents:
47045
diff
changeset
|
123 |
variablesWeight = 0.05, |
10bece4ac87e
more conservative Metis defaults, for backward compatiblity (as illustrated by one "metis" call in "Auth/KerberosV")
blanchet
parents:
47045
diff
changeset
|
124 |
literalsWeight = 0.01, |
39450
7e9879fbb7c5
supply the Metis parameter defaults as argument, instead of patching the Metis sources;
blanchet
parents:
39419
diff
changeset
|
125 |
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:
47015
diff
changeset
|
126 |
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:
47015
diff
changeset
|
127 |
{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:
47015
diff
changeset
|
128 |
|
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:
47015
diff
changeset
|
129 |
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:
47015
diff
changeset
|
130 |
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:
47015
diff
changeset
|
131 |
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:
47015
diff
changeset
|
132 |
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:
47015
diff
changeset
|
133 |
fun precedence p = |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58839
diff
changeset
|
134 |
(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:
47015
diff
changeset
|
135 |
EQUAL => #precedence Metis_KnuthBendixOrder.default p |
57408 | 136 |
| 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:
47015
diff
changeset
|
137 |
in {weight = weight, precedence = precedence} end |
37573 | 138 |
|
50875
bfb626265782
less brutal Metis failure -- the brutality was accidentally introduced by df8ae0590be2
blanchet
parents:
50705
diff
changeset
|
139 |
exception METIS_UNPROVABLE of unit |
bfb626265782
less brutal Metis failure -- the brutality was accidentally introduced by df8ae0590be2
blanchet
parents:
50705
diff
changeset
|
140 |
|
37516
c81c86bfc18a
have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet
parents:
37509
diff
changeset
|
141 |
(* Main function to start Metis proof and reconstruction *) |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
142 |
fun FOL_SOLVE infer_params th_name type_encs lam_trans clausify_refl ctxt cls0 ths0 = |
70489 | 143 |
let |
144 |
val thy = Proof_Context.theory_of ctxt |
|
145 |
||
146 |
val new_skolem = Config.get ctxt new_skolem orelse null (Meson.choice_theorems thy) |
|
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
70514
diff
changeset
|
147 |
val do_lams = lam_trans = liftingN ? introduce_lam_wrappers ctxt |
70489 | 148 |
val th_cls_pairs = |
74904 | 149 |
map_index (fn (j, th) => |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
150 |
(th, |
70489 | 151 |
th |
152 |
|> Drule.eta_contraction_rule |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
153 |
|> Meson_Clausify.cnf_axiom Meson.simp_options_all_true ctxt |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
154 |
{new_skolem = new_skolem, combs = (lam_trans = combsN), refl = clausify_refl} j |
70489 | 155 |
||> map do_lams)) |
74904 | 156 |
ths0 |
70489 | 157 |
val ths = maps (snd o snd) th_cls_pairs |
158 |
val dischargers = map (fst o snd) th_cls_pairs |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
159 |
val cls = cls0 |> map (Drule.eta_contraction_rule #> do_lams) |
70489 | 160 |
val _ = trace_msg ctxt (K "FOL_SOLVE: CONJECTURE CLAUSES") |
161 |
val _ = List.app (fn th => trace_msg ctxt (fn () => Thm.string_of_thm ctxt th)) cls |
|
162 |
||
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
163 |
val type_enc_str :: fallback_type_encs = type_encs |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
164 |
val _ = trace_msg ctxt (fn () => "type_enc = " ^ type_enc_str) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
165 |
val type_enc = type_enc_of_string Strict type_enc_str |
70489 | 166 |
|
167 |
val (sym_tab, axioms, ord_info, concealed) = |
|
168 |
generate_metis_problem ctxt type_enc lam_trans cls ths |
|
169 |
fun get_isa_thm mth Isa_Reflexive_or_Trivial = |
|
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51717
diff
changeset
|
170 |
reflexive_or_trivial_of_metis ctxt type_enc sym_tab concealed mth |
70489 | 171 |
| get_isa_thm mth Isa_Lambda_Lifted = |
52031
9a9238342963
tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents:
51717
diff
changeset
|
172 |
lam_lifted_of_metis ctxt type_enc sym_tab concealed mth |
70489 | 173 |
| get_isa_thm _ (Isa_Raw ith) = ith |
70514 | 174 |
val axioms = axioms |> map (fn (mth, ith) => |
175 |
(mth, get_isa_thm mth ith |> Thm.close_derivation \<^here>)) |
|
70489 | 176 |
val _ = trace_msg ctxt (K "ISABELLE CLAUSES") |
177 |
val _ = List.app (fn (_, ith) => trace_msg ctxt (fn () => Thm.string_of_thm ctxt ith)) axioms |
|
178 |
val _ = trace_msg ctxt (K "METIS CLAUSES") |
|
179 |
val _ = List.app (fn (mth, _) => trace_msg ctxt (fn () => Metis_Thm.toString mth)) axioms |
|
180 |
||
181 |
val _ = trace_msg ctxt (K "START METIS PROVE PROCESS") |
|
182 |
val ordering = |
|
183 |
if Config.get ctxt advisory_simp |
|
184 |
then kbo_advisory_simp_ordering (ord_info ()) |
|
185 |
else Metis_KnuthBendixOrder.default |
|
186 |
||
50875
bfb626265782
less brutal Metis failure -- the brutality was accidentally introduced by df8ae0590be2
blanchet
parents:
50705
diff
changeset
|
187 |
fun fall_back () = |
bfb626265782
less brutal Metis failure -- the brutality was accidentally introduced by df8ae0590be2
blanchet
parents:
50705
diff
changeset
|
188 |
(verbose_warning ctxt |
70489 | 189 |
("Falling back on " ^ quote (metis_call (hd fallback_type_encs) lam_trans) ^ "..."); |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
190 |
FOL_SOLVE infer_params th_name fallback_type_encs lam_trans clausify_refl ctxt cls0 ths0) |
32956 | 191 |
in |
69593 | 192 |
(case filter (fn t => Thm.prop_of t aconv \<^prop>\<open>False\<close>) cls of |
70489 | 193 |
false_th :: _ => [false_th RS @{thm FalseE}] |
194 |
| [] => |
|
195 |
(case Metis_Resolution.loop (Metis_Resolution.new (resolution_params ordering) |
|
196 |
{axioms = axioms |> map fst, conjecture = []}) of |
|
197 |
Metis_Resolution.Contradiction mth => |
|
198 |
let |
|
199 |
val _ = trace_msg ctxt (fn () => "METIS RECONSTRUCTION START: " ^ Metis_Thm.toString mth) |
|
200 |
val ctxt' = fold Variable.declare_constraints (map Thm.prop_of cls) ctxt |
|
201 |
(*add constraints arising from converting goal to clause form*) |
|
202 |
val proof = Metis_Proof.proof mth |
|
203 |
val result = fold (replay_one_inference ctxt' type_enc concealed sym_tab) proof axioms |
|
204 |
val used = map_filter (used_axioms axioms) proof |
|
205 |
val _ = trace_msg ctxt (K "METIS COMPLETED; clauses actually used:") |
|
206 |
val _ = List.app (fn th => trace_msg ctxt (fn () => Thm.string_of_thm ctxt th)) used |
|
207 |
||
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
208 |
val (used_ths, unused_ths) = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
209 |
List.partition (have_common_thm ctxt used o #2 o #2) th_cls_pairs |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
210 |
|> apply2 (map #1) |
70489 | 211 |
val _ = |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
212 |
if exists is_some (map th_name used_ths) then |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
213 |
infer_params := SOME ({ |
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81746
diff
changeset
|
214 |
infer_ctxt = ctxt', |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
215 |
type_enc = type_enc_str, |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
216 |
lam_trans = lam_trans, |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
217 |
th_name = th_name, |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
218 |
new_skolem = new_skolem, |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
219 |
th_cls_pairs = map (apsnd snd) th_cls_pairs, |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
220 |
lifted = fst concealed, |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
221 |
sym_tab = sym_tab, |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
222 |
axioms = axioms, |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
223 |
mth = mth |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
224 |
}) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
225 |
else (); |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
226 |
val _ = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
227 |
if not (null unused_ths) then |
80306 | 228 |
verbose_warning ctxt ("Unused theorems: " ^ |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
229 |
commas_quote (unused_ths |> map (fn th => |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
230 |
th_name th |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
231 |
|> the_default (Proof_Context.print_thm_name ctxt |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
232 |
(Thm.get_name_hint th))))) |
70489 | 233 |
else (); |
234 |
val _ = |
|
235 |
if not (null cls) andalso not (have_common_thm ctxt used cls) then |
|
236 |
verbose_warning ctxt "The assumptions are inconsistent" |
|
237 |
else (); |
|
238 |
in |
|
239 |
(case result of |
|
240 |
(_, ith) :: _ => |
|
241 |
(trace_msg ctxt (fn () => "Success: " ^ Thm.string_of_thm ctxt ith); |
|
242 |
[discharge_skolem_premises ctxt dischargers ith]) |
|
243 |
| _ => (trace_msg ctxt (K "Metis: No result"); [])) |
|
244 |
end |
|
245 |
| Metis_Resolution.Satisfiable _ => |
|
246 |
(trace_msg ctxt (K "Metis: No first-order proof with the supplied lemmas"); |
|
247 |
raise METIS_UNPROVABLE ())) |
|
248 |
handle METIS_UNPROVABLE () => if null fallback_type_encs then [] else fall_back () |
|
249 |
| METIS_RECONSTRUCT (loc, msg) => |
|
250 |
if null fallback_type_encs then |
|
251 |
(verbose_warning ctxt ("Failed to replay Metis proof\n" ^ loc ^ ": " ^ msg); []) |
|
252 |
else fall_back ()) |
|
42733
01ef1c3d9cfd
more robust exception handling in Metis (also works if there are several subgoals)
blanchet
parents:
42650
diff
changeset
|
253 |
end |
23442
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff
changeset
|
254 |
|
45508 | 255 |
fun neg_clausify ctxt combinators = |
38028 | 256 |
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:
43963
diff
changeset
|
257 |
#> Meson.make_clauses_unsorted ctxt |
55236 | 258 |
#> combinators ? map (Meson_Clausify.introduce_combinators_in_theorem ctxt) |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
259 |
#> Meson.finish_cnf true |
38028 | 260 |
|
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:
39267
diff
changeset
|
261 |
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:
39267
diff
changeset
|
262 |
(if exists (Meson.has_too_many_clauses ctxt) |
59582 | 263 |
(Logic.prems_of_goal (Thm.prop_of st0) 1) then |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
50875
diff
changeset
|
264 |
Simplifier.full_simp_tac (Meson_Clausify.ss_only @{thms not_all not_ex} ctxt) 1 |
55239 | 265 |
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:
39267
diff
changeset
|
266 |
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:
39267
diff
changeset
|
267 |
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:
39267
diff
changeset
|
268 |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
269 |
fun metis_tac_infer_params th_name type_encs0 lam_trans clausify_refl ctxt ths i = |
37926
e6ff246c0cdb
renamings + only need second component of name pool to reconstruct proofs
blanchet
parents:
37925
diff
changeset
|
270 |
let |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
271 |
val infer_params = Unsynchronized.ref NONE |
55520 | 272 |
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:
39964
diff
changeset
|
273 |
val _ = trace_msg ctxt (fn () => |
61268 | 274 |
"Metis called with theorems\n" ^ cat_lines (map (Thm.string_of_thm ctxt) ths)) |
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
275 |
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:
55520
diff
changeset
|
276 |
val combs = (lam_trans = combsN) |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
277 |
fun tac clause = resolve_tac ctxt |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
278 |
(FOL_SOLVE infer_params th_name type_encs lam_trans clausify_refl ctxt clause ths) 1 |
32956 | 279 |
in |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
280 |
Meson.MESON (preskolem_tac ctxt) (maps (neg_clausify ctxt combs)) tac ctxt i |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
281 |
#> Seq.map (fn st => (!infer_params, st)) |
32956 | 282 |
end |
23442
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff
changeset
|
283 |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
284 |
(* Theorem name functions *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
285 |
(* No theorem name (infer_thm_insts won't be called) *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
286 |
val no_th_name = K NONE |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
287 |
(* Name hint as theorem name (theorems in ths get a name but perhaps "??.unknown") *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
288 |
fun th_name_hint ths = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
289 |
Option.filter (member Thm.eq_thm ths) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
290 |
#> Option.map (Thm_Name.print o Thm.get_name_hint) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
291 |
(* Use multi_thm with input string to get theorem name (only if part of multi_ths) *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
292 |
fun th_name_multi_thm multi_ths th = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
293 |
let |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
294 |
fun index_of_th ths = find_index (curry Thm.eq_thm th) ths + 1 |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
295 |
fun make_name ([_], name) = name |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
296 |
(* This doesn't work if name already has attributes or is indexed, |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
297 |
* but provides much more information than "??.unknown" name hint *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
298 |
| make_name (ths, name) = name ^ "(" ^ string_of_int (index_of_th ths) ^ ")" |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
299 |
in |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
300 |
List.find (fn (ths, _) => member Thm.eq_thm ths th) multi_ths |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
301 |
|> Option.map make_name |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
302 |
end |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
303 |
|
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
304 |
(* Metis tactic to prove a subgoal and optionally suggest of-instantiations *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
305 |
fun metis_tac' th_name type_encs lam_trans ctxt ths i = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
306 |
let |
81746 | 307 |
val instantiate = Config.get ctxt instantiate |
308 |
val instantiate_tac = |
|
309 |
if instantiate then |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
310 |
(fn (NONE, st) => st |
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81746
diff
changeset
|
311 |
| (SOME infer_params, st) => tap (instantiate_call ctxt infer_params) st) |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
312 |
else |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
313 |
snd |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
314 |
in |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
315 |
(* "clausify_refl" (removal of redundant equalities) renames variables, |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
316 |
* which is bad for inferring variable instantiations. Metis doesn't need |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
317 |
* it, so we set it to "false" when we want to infer variable instantiations. |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
318 |
* We don't do it always because we don't want to break existing proofs. *) |
81746 | 319 |
metis_tac_infer_params th_name type_encs lam_trans (not instantiate) ctxt ths i |
320 |
#> Seq.map instantiate_tac |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
321 |
end |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
322 |
|
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
323 |
(* Metis tactic without theorem name, therefore won't suggest of-instantiations *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
324 |
val metis_tac = metis_tac' no_th_name |
55521
241c6a2fdda1
added a version of the Metis tactic that returns the unused facts
blanchet
parents:
55520
diff
changeset
|
325 |
|
55520 | 326 |
(* Whenever "X" has schematic type variables, we treat "using X by metis" as "by (metis X)" to |
327 |
prevent "Subgoal.FOCUS" from freezing the type variables. We don't do it for nonschematic facts |
|
328 |
"X" because this breaks a few proofs (in the rare and subtle case where a proof relied on |
|
329 |
extensionality not being applied) and brings few benefits. *) |
|
59582 | 330 |
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:
42847
diff
changeset
|
331 |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
332 |
(* Metis method to prove the goal and optionally suggest of-instantiations *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
333 |
fun metis_method' th_name ((override_type_encs, lam_trans), ths) ctxt facts = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
334 |
let |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
335 |
val (schem_facts, nonschem_facts) = List.partition has_tvar facts |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
336 |
in |
61841
4d3527b94f2a
more general types Proof.method / context_tactic;
wenzelm
parents:
61302
diff
changeset
|
337 |
HEADGOAL (Method.insert_tac ctxt nonschem_facts THEN' |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
338 |
CHANGED_PROP o metis_tac' th_name (these override_type_encs) |
55520 | 339 |
(the_default default_metis_lam_trans lam_trans) ctxt (schem_facts @ ths)) |
43099 | 340 |
end |
43100 | 341 |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
342 |
(* Metis method without theorem name, therefore won't suggest of-instantiations *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
343 |
val metis_method = metis_method' no_th_name |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
344 |
|
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
345 |
(* Metis method with input strings for better theorem names in of-instantiations *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
346 |
fun metis_method_multi_thms (opts, multi_ths) = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
347 |
metis_method' (th_name_multi_thm multi_ths) (opts, maps fst multi_ths) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
348 |
|
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
349 |
(* Use Metis to infer variable instantiations of theorems *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
350 |
fun metis_infer_thm_insts ((override_type_encs, lam_trans), ths) ctxt facts i = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
351 |
let |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
352 |
val (schem_facts, nonschem_facts) = List.partition has_tvar facts |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
353 |
val insert_tac = Method.insert_tac ctxt nonschem_facts i |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
354 |
val metis_tac = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
355 |
metis_tac_infer_params (th_name_hint ths) (these override_type_encs) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
356 |
(the_default default_metis_lam_trans lam_trans) false ctxt (schem_facts @ ths) i |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
357 |
in |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
358 |
Seq.THEN (insert_tac, metis_tac) |
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81746
diff
changeset
|
359 |
#> Seq.map (Option.mapPartial (infer_thm_insts ctxt) o fst) |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
360 |
#> Option.mapPartial fst o Seq.pull |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
361 |
end |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
362 |
|
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
70514
diff
changeset
|
363 |
val metis_lam_transs = [opaque_liftingN, liftingN, combsN] |
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
364 |
|
45578 | 365 |
fun set_opt _ x NONE = SOME x |
366 |
| set_opt get x (SOME x0) = |
|
55523
9429e7b5b827
removed final periods in messages for proof methods
blanchet
parents:
55521
diff
changeset
|
367 |
error ("Cannot specify both " ^ quote (get x0) ^ " and " ^ quote (get x)) |
54756 | 368 |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
369 |
fun consider_opt s = |
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
70514
diff
changeset
|
370 |
if s = "hide_lams" then |
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
70514
diff
changeset
|
371 |
error "\"hide_lams\" has been renamed \"opaque_lifting\"" |
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
70514
diff
changeset
|
372 |
else if member (op =) metis_lam_transs s then |
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
70514
diff
changeset
|
373 |
apsnd (set_opt I s) |
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
70514
diff
changeset
|
374 |
else |
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
70514
diff
changeset
|
375 |
apfst (set_opt hd [s]) |
45514 | 376 |
|
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
377 |
val parse_metis_options = |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
378 |
Scan.optional |
69593 | 379 |
(Args.parens (Args.name -- Scan.option (\<^keyword>\<open>,\<close> |-- Args.name)) |
45519
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
380 |
>> (fn (s, s') => |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
381 |
(NONE, NONE) |> consider_opt s |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
382 |
|> (case s' of SOME s' => consider_opt s' | _ => I))) |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
383 |
(NONE, NONE) |
cd6e78cb6ee8
make metis reconstruction handling more flexible
blanchet
parents:
45514
diff
changeset
|
384 |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
385 |
(* Parse multi_thm with input string *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
386 |
val parse_multi_thm = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
387 |
let |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
388 |
(* Remove spaces before and after "[]():," *) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
389 |
val strip_spaces = |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
390 |
ATP_Util.strip_spaces false (not o member (op =) (String.explode "[]():,")) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
391 |
in |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
392 |
ATP_Util.scan_and_trace_multi_thm |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
393 |
>> apsnd (map Token.unparse #> implode_space #> strip_spaces) |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
394 |
end |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
395 |
|
58818 | 396 |
val _ = |
397 |
Theory.setup |
|
69593 | 398 |
(Method.setup \<^binding>\<open>metis\<close> |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
399 |
(Scan.lift parse_metis_options -- Scan.repeat parse_multi_thm |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
80306
diff
changeset
|
400 |
>> (METHOD oo metis_method_multi_thms)) |
58818 | 401 |
"Metis for FOL and HOL problems") |
23442
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff
changeset
|
402 |
|
028e39e5e8f3
The Metis prover (slightly modified version from Larry);
wenzelm
parents:
diff
changeset
|
403 |
end; |