author | wenzelm |
Wed, 25 Jun 2025 16:35:25 +0200 | |
changeset 82768 | 8f866fd6fae1 |
parent 82574 | 318526b74e6f |
permissions | -rw-r--r-- |
81254 | 1 |
(* Title: HOL/Tools/Metis/metis_instantiations.ML |
81757
4d15005da582
tuned documentation and order of instantiated facts
Lukas Bartl
parents:
81746
diff
changeset
|
2 |
Author: Lukas Bartl, Universitaet Augsburg |
81254 | 3 |
|
4 |
Inference of Variable Instantiations for Metis. |
|
5 |
*) |
|
6 |
||
7 |
signature METIS_INSTANTIATIONS = |
|
8 |
sig |
|
9 |
type inst |
|
10 |
||
11 |
type infer_params = { |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
12 |
infer_ctxt : Proof.context, |
81254 | 13 |
type_enc : string, |
14 |
lam_trans : string, |
|
15 |
th_name : thm -> string option, |
|
16 |
new_skolem : bool, |
|
17 |
th_cls_pairs : (thm * thm list) list, |
|
18 |
lifted : (string * term) list, |
|
19 |
sym_tab : int Symtab.table, |
|
20 |
axioms : (Metis_Thm.thm * thm) list, |
|
21 |
mth : Metis_Thm.thm |
|
22 |
} |
|
23 |
||
81746 | 24 |
val instantiate : bool Config.T |
25 |
val instantiate_undefined : bool Config.T |
|
81254 | 26 |
val metis_call : string -> string -> string |
27 |
val table_of_thm_inst : thm * inst -> cterm Vars.table |
|
28 |
val pretty_name_inst : Proof.context -> string * inst -> Pretty.T |
|
29 |
val string_of_name_inst : Proof.context -> string * inst -> string |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
30 |
val infer_thm_insts : Proof.context -> infer_params -> (thm * inst) list option |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
31 |
val instantiate_call : Proof.context -> infer_params -> thm -> unit |
81254 | 32 |
end; |
33 |
||
34 |
structure Metis_Instantiations : METIS_INSTANTIATIONS = |
|
35 |
struct |
|
36 |
||
37 |
open ATP_Problem_Generate |
|
38 |
open ATP_Proof_Reconstruct |
|
39 |
open Metis_Generate |
|
40 |
open Metis_Reconstruct |
|
41 |
||
42 |
(* Type of variable instantiations of a theorem. This is a list of (certified) |
|
43 |
* terms suitably ordered for an of-instantiation of the theorem. *) |
|
44 |
type inst = cterm list |
|
45 |
||
46 |
(* Params needed for inference of variable instantiations *) |
|
47 |
type infer_params = { |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
48 |
infer_ctxt : Proof.context, |
81254 | 49 |
type_enc : string, |
50 |
lam_trans : string, |
|
51 |
th_name : thm -> string option, |
|
52 |
new_skolem : bool, |
|
53 |
th_cls_pairs : (thm * thm list) list, |
|
54 |
lifted : (string * term) list, |
|
55 |
sym_tab : int Symtab.table, |
|
56 |
axioms : (Metis_Thm.thm * thm) list, |
|
57 |
mth : Metis_Thm.thm |
|
58 |
} |
|
59 |
||
60 |
(* Config option to enable suggestion of of-instantiations (e.g. "assms(1)[of x]") *) |
|
81746 | 61 |
val instantiate = Attrib.setup_config_bool @{binding "metis_instantiate"} (K false) |
81254 | 62 |
(* Config option to allow instantiation of variables with "undefined" *) |
81746 | 63 |
val instantiate_undefined = Attrib.setup_config_bool @{binding "metis_instantiate_undefined"} (K true) |
81254 | 64 |
|
65 |
fun metis_call type_enc lam_trans = |
|
66 |
let |
|
67 |
val type_enc = |
|
68 |
case AList.find (fn (enc, encs) => enc = hd encs) type_enc_aliases type_enc of |
|
69 |
[alias] => alias |
|
70 |
| _ => type_enc |
|
71 |
val opts = |
|
72 |
[] |> lam_trans <> default_metis_lam_trans ? cons lam_trans |
|
73 |
|> type_enc <> partial_typesN ? cons type_enc |
|
74 |
in |
|
75 |
metisN ^ (if null opts then "" else " (" ^ commas opts ^ ")") |
|
76 |
end |
|
77 |
||
78 |
(* Variables of a theorem ordered the same way as in of-instantiations |
|
79 |
* (cf. Rule_Insts.of_rule) *) |
|
80 |
fun of_vars_of_thm th = Vars.build (Vars.add_vars (Thm.full_prop_of th)) |> Vars.list_set |
|
81 |
||
82 |
(* "_" terms are used for variables without instantiation *) |
|
83 |
val is_dummy_cterm = Term.is_dummy_pattern o Thm.term_of |
|
84 |
val thm_insts_trivial = forall (forall is_dummy_cterm o snd) |
|
85 |
||
86 |
fun table_of_thm_inst (th, cts) = |
|
87 |
of_vars_of_thm th ~~ cts |
|
88 |
|> filter_out (is_dummy_cterm o snd) |
|
89 |
|> Vars.make |
|
90 |
||
91 |
fun open_attrib name = |
|
92 |
case try (unsuffix "]") name of |
|
93 |
NONE => name ^ "[" |
|
94 |
| SOME name' => name' ^ ", " |
|
95 |
||
96 |
fun pretty_inst_cterm ctxt ct = |
|
97 |
let |
|
98 |
val pretty = Syntax.pretty_term ctxt (Thm.term_of ct) |
|
99 |
val is_dummy = ATP_Util.content_of_pretty pretty = "_" |
|
100 |
in |
|
101 |
(* Do not quote single "_" *) |
|
102 |
pretty |> not is_dummy ? ATP_Util.pretty_maybe_quote ctxt |
|
103 |
end |
|
104 |
||
105 |
(* Pretty of-instantiation *) |
|
106 |
fun pretty_name_inst ctxt (name, inst) = |
|
107 |
case drop_suffix is_dummy_cterm inst of |
|
108 |
[] => Pretty.str name |
|
109 |
| cts => |
|
110 |
map (pretty_inst_cterm ctxt) cts |
|
111 |
|> Pretty.breaks o cons (Pretty.str "of") |
|
112 |
|> Pretty.enclose (open_attrib name) "]" |
|
113 |
||
114 |
val string_of_name_inst = Pretty.string_of oo pretty_name_inst |
|
115 |
||
116 |
(* Infer Metis axioms with corresponding subtitutions *) |
|
117 |
fun infer_metis_axiom_substs mth = |
|
118 |
let |
|
119 |
fun add_metis_axiom_substs msubst mth = |
|
120 |
case Metis_Thm.inference mth of |
|
121 |
(Metis_Thm.Axiom, []) => cons (msubst, mth) |
|
122 |
| (Metis_Thm.Metis_Subst, _) => |
|
123 |
(case Metis_Proof.thmToInference mth of |
|
124 |
Metis_Proof.Metis_Subst (msubst', par) => |
|
125 |
add_metis_axiom_substs (Metis_Subst.compose msubst' msubst) par |
|
126 |
| _ => raise Fail "expected Metis_Subst") |
|
127 |
| (_, pars) => fold (add_metis_axiom_substs msubst) pars |
|
128 |
in |
|
129 |
add_metis_axiom_substs Metis_Subst.empty mth [] |
|
130 |
end |
|
131 |
||
132 |
(* Variables are renamed during clausification by importing and exporting |
|
133 |
* theorems. Do the same here to find the right variable. *) |
|
134 |
fun import_export_thm ctxt th = |
|
135 |
let |
|
136 |
(* cf. Meson_Clausify.nnf_axiom *) |
|
137 |
val (import_th, import_ctxt) = |
|
138 |
Drule.zero_var_indexes th |
|
139 |
|> (fn th' => Variable.import true [th'] ctxt) |
|
140 |
|>> the_single o snd |
|
141 |
(* cf. Meson_Clausify.cnf_axiom *) |
|
142 |
val export_th = import_th |
|
143 |
|> singleton (Variable.export import_ctxt ctxt) |
|
144 |
in |
|
145 |
(* cf. Meson.finish_cnf *) |
|
146 |
Drule.zero_var_indexes export_th |
|
147 |
end |
|
148 |
||
149 |
(* Find the theorem corresponding to a Metis axiom if it has a name. |
|
150 |
* "Theorems" are Isabelle theorems given to the metis tactic. |
|
151 |
* "Axioms" are clauses given to the Metis prover. *) |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
152 |
fun metis_axiom_to_thm ctxt {th_name, th_cls_pairs, axioms, ...} (msubst, maxiom) = |
81254 | 153 |
let |
154 |
val axiom = lookth axioms maxiom |
|
155 |
in |
|
156 |
List.find (have_common_thm ctxt [axiom] o snd) th_cls_pairs |
|
157 |
|> Option.mapPartial (fn (th, _) => th_name th |> Option.map (pair th)) |
|
158 |
|> Option.map (pair (msubst, maxiom, axiom)) |
|
159 |
end |
|
160 |
||
161 |
(* Build a table : term Vartab.table list Thmtab.table that maps a theorem to a |
|
162 |
* list of variable substitutions (theorems can be instantiated multiple times) |
|
163 |
* from Metis substitutions *) |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
164 |
fun metis_substs_to_table ctxt {infer_ctxt, type_enc, lifted, new_skolem, sym_tab, ...} |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
165 |
th_of_vars th_msubsts = |
81254 | 166 |
let |
167 |
val _ = trace_msg ctxt (K "AXIOM SUBSTITUTIONS") |
|
168 |
val type_enc = type_enc_of_string Strict type_enc |
|
169 |
val thy = Proof_Context.theory_of ctxt |
|
170 |
||
171 |
(* Replace Skolem terms with "_" because they are unknown to the user |
|
172 |
* (cf. Metis_Generate.reveal_old_skolem_terms) *) |
|
173 |
val dummy_old_skolem_terms = map_aterms |
|
174 |
(fn t as Const (s, T) => |
|
175 |
if String.isPrefix old_skolem_const_prefix s |
|
176 |
then Term.dummy_pattern T |
|
177 |
else t |
|
178 |
| t => t) |
|
179 |
||
180 |
(* Infer types and replace "_" terms/types with schematic variables *) |
|
181 |
val infer_types_pattern = |
|
182 |
singleton (Type_Infer_Context.infer_types_finished |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
183 |
(Proof_Context.set_mode Proof_Context.mode_pattern infer_ctxt)) |
81254 | 184 |
|
185 |
(* "undefined" if allowed and not using new_skolem, "_" otherwise. *) |
|
186 |
val undefined_pattern = |
|
187 |
(* new_skolem uses schematic variables which should not be instantiated with "undefined" *) |
|
81746 | 188 |
if not new_skolem andalso Config.get ctxt instantiate_undefined then |
81254 | 189 |
Const o (pair @{const_name "undefined"}) |
190 |
else |
|
191 |
Term.dummy_pattern |
|
192 |
||
193 |
(* Instantiate schematic and free variables *) |
|
194 |
val inst_vars_frees = map_aterms |
|
195 |
(fn t as Free (x, T) => |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
196 |
(* an undeclared/internal free variable is unknown/inaccessible to the user, |
81254 | 197 |
* so we replace it with "_" *) |
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
198 |
let |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
199 |
val x' = Variable.revert_fixed ctxt x |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
200 |
in |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
201 |
if not (Variable.is_declared ctxt x') orelse Name.is_internal x' then |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
202 |
Term.dummy_pattern T |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
203 |
else |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
204 |
t |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
205 |
end |
81254 | 206 |
| Var (_, T) => |
207 |
(* a schematic variable can be replaced with any term, |
|
208 |
* so we replace it with "undefined" *) |
|
209 |
undefined_pattern T |
|
210 |
| t => t) |
|
211 |
||
212 |
(* Remove arguments of "_" unknown functions because the functions could |
|
213 |
* change (e.g. instantiations can change Skolem functions). |
|
214 |
* Type inference also introduces arguments (cf. Term.replace_dummy_patterns). *) |
|
215 |
fun eliminate_dummy_args (t $ u) = |
|
216 |
let |
|
217 |
val t' = eliminate_dummy_args t |
|
218 |
in |
|
219 |
if Term.is_dummy_pattern t' then |
|
220 |
Term.dummy_pattern (Term.fastype_of t' |> Term.range_type) |
|
221 |
else |
|
222 |
t' $ eliminate_dummy_args u |
|
223 |
end |
|
224 |
| eliminate_dummy_args (Abs (s, T, t)) = Abs (s, T, eliminate_dummy_args t) |
|
225 |
| eliminate_dummy_args t = t |
|
226 |
||
227 |
(* Polish Isabelle term, s.t. it can be displayed |
|
228 |
* (cf. Metis_Reconstruct.polish_hol_terms and ATP_Proof_Reconstruct.termify_line) *) |
|
229 |
val polish_term = |
|
230 |
reveal_lam_lifted lifted (* reveal lifted lambdas *) |
|
231 |
#> dummy_old_skolem_terms (* eliminate Skolem terms *) |
|
232 |
#> infer_types_pattern (* type inference *) |
|
233 |
#> eliminate_lam_wrappers (* eliminate Metis.lambda wrappers *) |
|
234 |
#> uncombine_term thy (* eliminate Meson.COMB* terms *) |
|
235 |
#> Envir.beta_eta_contract (* simplify lambda terms *) |
|
236 |
#> simplify_bool (* simplify boolean terms *) |
|
237 |
#> Term.show_dummy_patterns (* reveal "_" that have been replaced by type inference *) |
|
238 |
#> inst_vars_frees (* eliminate schematic and free variables *) |
|
239 |
#> eliminate_dummy_args (* eliminate arguments of "_" (unknown functions) *) |
|
240 |
||
241 |
(* Translate a Metis substitution to Isabelle *) |
|
242 |
fun add_subst_to_table ((msubst, maxiom, axiom), (th, name)) th_substs_table = |
|
243 |
let |
|
244 |
val _ = trace_msg ctxt (fn () => "Metis axiom: " ^ Metis_Thm.toString maxiom) |
|
245 |
val _ = trace_msg ctxt (fn () => "Metis substitution: " ^ Metis_Subst.toString msubst) |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
246 |
val _ = trace_msg ctxt (fn () => "Isabelle axiom: " ^ Thm.string_of_thm infer_ctxt axiom) |
81254 | 247 |
val _ = trace_msg ctxt (fn () => "Isabelle theorem: " ^ |
248 |
Thm.string_of_thm ctxt th ^ " (" ^ name ^ ")") |
|
249 |
||
250 |
(* Only indexnames of variables without types |
|
251 |
* because types can change during clausification *) |
|
252 |
val vars = |
|
253 |
inter (op =) (map fst (of_vars_of_thm axiom)) |
|
254 |
(map fst (Thmtab.lookup th_of_vars th |> the)) |
|
255 |
||
256 |
(* Translate Metis variable and term to Isabelle *) |
|
257 |
fun translate_var_term (mvar, mt) = |
|
258 |
Metis_Name.toString mvar |
|
259 |
(* cf. remove_typeinst in Metis_Reconstruct.inst_inference *) |
|
260 |
|> unprefix_and_unascii schematic_var_prefix |
|
261 |
(* cf. find_var in Metis_Reconstruct.inst_inference *) |
|
262 |
|> Option.mapPartial (fn name => List.find (fn (a, _) => a = name) vars) |
|
263 |
(* cf. subst_translation in Metis_Reconstruct.inst_inference *) |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
264 |
|> Option.map (fn var => (var, hol_term_of_metis infer_ctxt type_enc sym_tab mt)) |
81254 | 265 |
|
266 |
(* Build the substitution table and instantiate the remaining schematic variables *) |
|
267 |
fun build_subst_table substs = |
|
268 |
let |
|
269 |
(* Variable of the axiom that didn't get instantiated by Metis, |
|
270 |
* the type is later set via type unification *) |
|
271 |
val undefined_term = undefined_pattern (TVar ((Name.aT, 0), [])) |
|
272 |
in |
|
273 |
Vartab.build (vars |> fold (fn var => Vartab.default |
|
274 |
(var, AList.lookup (op =) substs var |> the_default undefined_term))) |
|
275 |
end |
|
276 |
||
277 |
val raw_substs = |
|
278 |
Metis_Subst.toList msubst |
|
279 |
|> map_filter translate_var_term |
|
280 |
val _ = if null raw_substs then () else |
|
281 |
trace_msg ctxt (fn () => cat_lines ("Raw Isabelle substitution:" :: |
|
282 |
(raw_substs |> map (fn (var, t) => |
|
283 |
Term.string_of_vname var ^ " |-> " ^ |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
284 |
Syntax.string_of_term infer_ctxt t)))) |
81254 | 285 |
|
286 |
val subst = raw_substs |
|
287 |
|> map (apsnd polish_term) |
|
288 |
|> build_subst_table |
|
289 |
val _ = trace_msg ctxt (fn () => |
|
290 |
if Vartab.is_empty subst then |
|
291 |
"No Isabelle substitution" |
|
292 |
else |
|
293 |
cat_lines ("Final Isabelle substitution:" :: |
|
294 |
(Vartab.dest subst |> map (fn (var ,t) => |
|
295 |
Term.string_of_vname var ^ " |-> " ^ |
|
296 |
Syntax.string_of_term ctxt t)))) |
|
297 |
||
298 |
(* Try to merge the substitution with another substitution of the theorem *) |
|
299 |
fun insert_subst [] = [subst] |
|
300 |
| insert_subst (subst' :: substs') = |
|
301 |
Vartab.merge Term.aconv_untyped (subst', subst) :: substs' |
|
302 |
handle Vartab.DUP _ => subst' :: insert_subst substs' |
|
303 |
in |
|
304 |
Thmtab.lookup th_substs_table th |
|
305 |
|> insert_subst o these |
|
306 |
|> (fn substs => Thmtab.update (th, substs) th_substs_table) |
|
307 |
end |
|
308 |
in |
|
309 |
fold add_subst_to_table th_msubsts Thmtab.empty |
|
310 |
end |
|
311 |
||
312 |
(* Build variable instantiations : thm * inst list from the table *) |
|
313 |
fun table_to_thm_insts ctxt th_of_vars th_substs_table = |
|
314 |
let |
|
315 |
val thy = Proof_Context.theory_of ctxt |
|
316 |
||
317 |
(* Unify types of a variable with the term for instantiation |
|
318 |
* (cf. Drule.infer_instantiate_types, Metis_Reconstruct.cterm_incr_types) *) |
|
319 |
fun unify (tyenv, maxidx) T t = |
|
320 |
let |
|
321 |
val t' = Term.map_types (Logic.incr_tvar (maxidx + 1)) t |
|
322 |
val U = Term.fastype_of t' |
|
323 |
val maxidx' = Term.maxidx_term t' maxidx |
|
324 |
in |
|
325 |
(t', Sign.typ_unify thy (T, U) (tyenv, maxidx')) |
|
326 |
end |
|
327 |
||
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
328 |
(* Instantiate type variables with a type unifier and import remaining ones |
81254 | 329 |
* (cf. Drule.infer_instantiate_types) *) |
330 |
fun inst_unifier (ts, tyenv) = |
|
331 |
let |
|
332 |
val instT = |
|
333 |
TVars.build (tyenv |> Vartab.fold (fn (xi, (S, T)) => |
|
334 |
TVars.add ((xi, S), Envir.norm_type tyenv T))) |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
335 |
val ts' = map (Term_Subst.instantiate (instT, Vars.empty)) ts |
81254 | 336 |
in |
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
337 |
Variable.importT_terms ts' ctxt |> fst |
81254 | 338 |
end |
339 |
||
340 |
(* Build variable instantiations from a substitution table and instantiate |
|
341 |
* the remaining schematic variables with "_" *) |
|
342 |
fun add_subst th subst = |
|
343 |
let |
|
344 |
val of_vars = Thmtab.lookup th_of_vars th |> the |
|
345 |
||
346 |
fun unify_or_dummy (var, T) unify_params = |
|
347 |
Vartab.lookup subst var |
|
348 |
|> Option.map (unify unify_params T) |
|
349 |
|> the_default (Term.dummy_pattern T, unify_params) |
|
350 |
in |
|
351 |
fold_map unify_or_dummy of_vars (Vartab.empty, Thm.maxidx_of th) |
|
352 |
|> inst_unifier o apsnd fst |
|
353 |
|> map (Thm.cterm_of ctxt) |
|
354 |
|> cons o pair th |
|
355 |
end |
|
356 |
in |
|
357 |
Thmtab.fold_rev (fn (th, substs) => fold (add_subst th) substs) th_substs_table [] |
|
358 |
end |
|
359 |
||
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
360 |
fun really_infer_thm_insts ctxt (infer_params as {th_name, th_cls_pairs, mth, ...}) = |
81254 | 361 |
let |
362 |
(* Compute the theorem variables ordered as in of-instantiations. |
|
363 |
* import_export_thm is used to get the same variables as in axioms. *) |
|
364 |
val th_of_vars = |
|
365 |
Thmtab.build (th_cls_pairs |> fold (fn (th, _) => Thmtab.default |
|
366 |
(th, of_vars_of_thm (import_export_thm ctxt th)))) |
|
367 |
||
368 |
val th_insts = |
|
369 |
infer_metis_axiom_substs mth |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
370 |
|> map_filter (metis_axiom_to_thm ctxt infer_params) |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
371 |
|> metis_substs_to_table ctxt infer_params th_of_vars |
81254 | 372 |
|> table_to_thm_insts ctxt th_of_vars |
373 |
||
374 |
val _ = trace_msg ctxt (fn () => cat_lines ("THEOREM INSTANTIATIONS" :: |
|
375 |
(if thm_insts_trivial th_insts then |
|
376 |
["No instantiation available"] |
|
377 |
else |
|
378 |
th_insts |> maps (fn th_inst as (th, _) => |
|
379 |
let |
|
380 |
val table = table_of_thm_inst th_inst |
|
381 |
in |
|
382 |
if Vars.is_empty table then |
|
383 |
[] |
|
384 |
else |
|
385 |
"Theorem " ^ Thm.string_of_thm ctxt th ^ " (" ^ the (th_name th) ^ "):" :: |
|
386 |
(Vars.dest table |> map (fn (var, ct) => |
|
387 |
Syntax.string_of_term ctxt (Var var) ^ " |-> " ^ |
|
388 |
Syntax.string_of_term ctxt (Thm.term_of ct))) |
|
389 |
end)))) |
|
390 |
in |
|
391 |
th_insts |
|
392 |
end |
|
393 |
||
394 |
(* Infer variable instantiations *) |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
395 |
fun infer_thm_insts ctxt infer_params = |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
396 |
\<^try>\<open>SOME (really_infer_thm_insts ctxt infer_params) |
81254 | 397 |
catch exn => (trace_msg ctxt (fn () => Runtime.exn_message exn); NONE)\<close> |
398 |
||
399 |
(* Write suggested command with of-instantiations *) |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
400 |
fun write_command ctxt {type_enc, lam_trans, th_name, ...} st th_insts = |
81254 | 401 |
let |
402 |
val _ = trace_msg ctxt (K "SUGGEST OF-INSTANTIATIONS") |
|
403 |
val apply_str = if Thm.nprems_of st = 0 then "by" else "apply" |
|
404 |
val command_str = th_insts |
|
405 |
|> map (pretty_name_inst ctxt o apfst (the o th_name)) |
|
406 |
|> cons (Pretty.str (metis_call type_enc lam_trans)) |
|
407 |
|> Pretty.enclose (apply_str ^ " (") ")" o Pretty.breaks |
|
408 |
|> Pretty.symbolic_string_of (* markup string *) |
|
409 |
in |
|
410 |
writeln ("Found variable instantiations, try this:" ^ |
|
411 |
(* Add optional markup break (command may need multiple lines) *) |
|
412 |
Markup.markup (Markup.break {width = 1, indent = 2}) " " ^ |
|
413 |
Active.sendback_markup_command command_str) |
|
414 |
end |
|
415 |
||
416 |
(* Infer variable instantiations and suggest of-instantiations *) |
|
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
417 |
fun instantiate_call ctxt infer_params st = |
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
418 |
infer_thm_insts ctxt infer_params |
81254 | 419 |
|> Option.mapPartial (Option.filter (not o thm_insts_trivial)) |
82574
318526b74e6f
tuned metis_instantiate (import types, support unfixed variables)
Lukas Bartl <lukas.bartl@uni-a.de>
parents:
81757
diff
changeset
|
420 |
|> Option.app (write_command ctxt infer_params st) |
81254 | 421 |
|
422 |
end; |