author | wenzelm |
Tue, 21 Jan 2025 19:26:09 +0100 | |
changeset 81945 | 23957ebffaf7 |
parent 74525 | c960bfcb91db |
permissions | -rw-r--r-- |
52286 | 1 |
(* Title: HOL/Tools/reification.ML |
51726 | 2 |
Author: Amine Chaieb, TU Muenchen |
3 |
||
4 |
A trial for automatical reification. |
|
5 |
*) |
|
6 |
||
52286 | 7 |
signature REIFICATION = |
51726 | 8 |
sig |
52286 | 9 |
val conv: Proof.context -> thm list -> conv |
10 |
val tac: Proof.context -> thm list -> term option -> int -> tactic |
|
11 |
val lift_conv: Proof.context -> conv -> term option -> int -> tactic |
|
12 |
val dereify: Proof.context -> thm list -> conv |
|
51726 | 13 |
end; |
14 |
||
52286 | 15 |
structure Reification : REIFICATION = |
51726 | 16 |
struct |
17 |
||
69593 | 18 |
fun dest_listT (Type (\<^type_name>\<open>list\<close>, [T])) = T; |
52275 | 19 |
|
51726 | 20 |
val FWD = curry (op OF); |
21 |
||
52275 | 22 |
fun rewrite_with ctxt eqs = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps eqs); |
23 |
||
24 |
val pure_subst = @{lemma "x == y ==> PROP P y ==> PROP P x" by simp} |
|
51726 | 25 |
|
60696 | 26 |
fun lift_conv ctxt conv some_t = Subgoal.FOCUS (fn {context = ctxt', concl, ...} => |
52275 | 27 |
let |
60696 | 28 |
val ct = |
29 |
(case some_t of |
|
30 |
NONE => Thm.dest_arg concl |
|
31 |
| SOME t => Thm.cterm_of ctxt' t) |
|
52275 | 32 |
val thm = conv ct; |
33 |
in |
|
34 |
if Thm.is_reflexive thm then no_tac |
|
60752 | 35 |
else ALLGOALS (resolve_tac ctxt [pure_subst OF [thm]]) |
52275 | 36 |
end) ctxt; |
51726 | 37 |
|
38 |
(* Make a congruence rule out of a defining equation for the interpretation |
|
39 |
||
40 |
th is one defining equation of f, |
|
41 |
i.e. th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" |
|
42 |
Cp is a constructor pattern and P is a pattern |
|
43 |
||
44 |
The result is: |
|
45 |
[|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn) |
|
46 |
+ the a list of names of the A1 .. An, Those are fresh in the ctxt *) |
|
47 |
||
48 |
fun mk_congeq ctxt fs th = |
|
49 |
let |
|
59582 | 50 |
val Const (fN, _) = th |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |
51726 | 51 |
|> fst |> strip_comb |> fst; |
52 |
val ((_, [th']), ctxt') = Variable.import true [th] ctxt; |
|
53 |
val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th')); |
|
54 |
fun add_fterms (t as t1 $ t2) = |
|
55 |
if exists (fn f => Term.could_unify (t |> strip_comb |> fst, f)) fs |
|
56 |
then insert (op aconv) t |
|
57 |
else add_fterms t1 #> add_fterms t2 |
|
58 |
| add_fterms (t as Abs _) = |
|
59 |
if exists_Const (fn (c, _) => c = fN) t |
|
60 |
then K [t] |
|
61 |
else K [] |
|
62 |
| add_fterms _ = I; |
|
63 |
val fterms = add_fterms rhs []; |
|
64 |
val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt'; |
|
65 |
val tys = map fastype_of fterms; |
|
66 |
val vs = map Free (xs ~~ tys); |
|
67 |
val env = fterms ~~ vs; (*FIXME*) |
|
68 |
fun replace_fterms (t as t1 $ t2) = |
|
69 |
(case AList.lookup (op aconv) env t of |
|
70 |
SOME v => v |
|
71 |
| NONE => replace_fterms t1 $ replace_fterms t2) |
|
72 |
| replace_fterms t = |
|
73 |
(case AList.lookup (op aconv) env t of |
|
74 |
SOME v => v |
|
75 |
| NONE => t); |
|
76 |
fun mk_def (Abs (x, xT, t), v) = |
|
77 |
HOLogic.mk_Trueprop (HOLogic.all_const xT $ Abs (x, xT, HOLogic.mk_eq (v $ Bound 0, t))) |
|
78 |
| mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t)); |
|
79 |
fun tryext x = |
|
80 |
(x RS @{lemma "(\<forall>x. f x = g x) \<Longrightarrow> f = g" by blast} handle THM _ => x); |
|
81 |
val cong = |
|
82 |
(Goal.prove ctxt'' [] (map mk_def env) |
|
83 |
(HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs))) |
|
84 |
(fn {context, prems, ...} => |
|
63170 | 85 |
Local_Defs.unfold0_tac context (map tryext prems) THEN resolve_tac ctxt'' [th'] 1)) RS sym; |
51726 | 86 |
val (cong' :: vars') = |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59617
diff
changeset
|
87 |
Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o Thm.cterm_of ctxt'') vs); |
51726 | 88 |
val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars'; |
89 |
||
90 |
in (vs', cong') end; |
|
91 |
||
92 |
(* congs is a list of pairs (P,th) where th is a theorem for |
|
93 |
[| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *) |
|
94 |
||
95 |
fun rearrange congs = |
|
96 |
let |
|
97 |
fun P (_, th) = |
|
69593 | 98 |
let val \<^term>\<open>Trueprop\<close> $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ l $ _) = Thm.concl_of th |
51726 | 99 |
in can dest_Var l end; |
100 |
val (yes, no) = List.partition P congs; |
|
101 |
in no @ yes end; |
|
102 |
||
52275 | 103 |
fun dereify ctxt eqs = |
104 |
rewrite_with ctxt (eqs @ @{thms nth_Cons_0 nth_Cons_Suc}); |
|
105 |
||
52288 | 106 |
fun index_of t bds = |
107 |
let |
|
108 |
val tt = HOLogic.listT (fastype_of t); |
|
109 |
in |
|
110 |
(case AList.lookup Type.could_unify bds tt of |
|
111 |
NONE => error "index_of: type not found in environements!" |
|
112 |
| SOME (tbs, tats) => |
|
113 |
let |
|
114 |
val i = find_index (fn t' => t' = t) tats; |
|
115 |
val j = find_index (fn t' => t' = t) tbs; |
|
116 |
in |
|
117 |
if j = ~1 then |
|
118 |
if i = ~1 |
|
119 |
then (length tbs + length tats, AList.update Type.could_unify (tt, (tbs, tats @ [t])) bds) |
|
120 |
else (i, bds) |
|
121 |
else (j, bds) |
|
122 |
end) |
|
123 |
end; |
|
124 |
||
125 |
(* Generic decomp for reification : matches the actual term with the |
|
126 |
rhs of one cong rule. The result of the matching guides the |
|
127 |
proof synthesis: The matches of the introduced Variables A1 .. An are |
|
128 |
processed recursively |
|
129 |
The rest is instantiated in the cong rule,i.e. no reification is needed *) |
|
130 |
||
131 |
(* da is the decomposition for atoms, ie. it returns ([],g) where g |
|
132 |
returns the right instance f (AtC n) = t , where AtC is the Atoms |
|
133 |
constructor and n is the number of the atom corresponding to t *) |
|
134 |
fun decomp_reify da cgns (ct, ctxt) bds = |
|
51726 | 135 |
let |
52288 | 136 |
val thy = Proof_Context.theory_of ctxt; |
137 |
fun tryabsdecomp (ct, ctxt) bds = |
|
138 |
(case Thm.term_of ct of |
|
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74408
diff
changeset
|
139 |
Abs (_, xT, _) => |
52288 | 140 |
let |
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74408
diff
changeset
|
141 |
val ((cx, cta), ctxt') = Variable.dest_abs_cterm ct ctxt; |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74408
diff
changeset
|
142 |
val x = Thm.term_of cx; |
52288 | 143 |
val bds = (case AList.lookup Type.could_unify bds (HOLogic.listT xT) of |
144 |
NONE => error "tryabsdecomp: Type not found in the Environement" |
|
145 |
| SOME (bsT, atsT) => AList.update Type.could_unify (HOLogic.listT xT, |
|
146 |
(x :: bsT, atsT)) bds); |
|
147 |
in (([(cta, ctxt')], |
|
148 |
fn ([th], bds) => |
|
149 |
(hd (Variable.export ctxt' ctxt [(Thm.forall_intr cx th) COMP allI]), |
|
150 |
let |
|
151 |
val (bsT, asT) = the (AList.lookup Type.could_unify bds (HOLogic.listT xT)); |
|
152 |
in |
|
153 |
AList.update Type.could_unify (HOLogic.listT xT, (tl bsT, asT)) bds |
|
154 |
end)), |
|
155 |
bds) |
|
156 |
end |
|
157 |
| _ => da (ct, ctxt) bds) |
|
158 |
in |
|
159 |
(case cgns of |
|
160 |
[] => tryabsdecomp (ct, ctxt) bds |
|
161 |
| ((vns, cong) :: congs) => |
|
162 |
(let |
|
163 |
val (tyenv, tmenv) = |
|
164 |
Pattern.match thy |
|
59582 | 165 |
((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (Thm.concl_of cong), Thm.term_of ct) |
52288 | 166 |
(Vartab.empty, Vartab.empty); |
167 |
val (fnvs, invs) = List.partition (fn ((vn, _),_) => member (op =) vns vn) (Vartab.dest tmenv); |
|
168 |
val (fts, its) = |
|
169 |
(map (snd o snd) fnvs, |
|
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59642
diff
changeset
|
170 |
map (fn ((vn, vi), (tT, t)) => (((vn, vi), tT), Thm.cterm_of ctxt t)) invs); |
59617 | 171 |
val ctyenv = |
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59642
diff
changeset
|
172 |
map (fn ((vn, vi), (s, ty)) => (((vn, vi), s), Thm.ctyp_of ctxt ty)) |
59617 | 173 |
(Vartab.dest tyenv); |
52288 | 174 |
in |
59642 | 175 |
((map (Thm.cterm_of ctxt) fts ~~ replicate (length fts) ctxt, |
74282 | 176 |
apfst (FWD (Drule.instantiate_normalize (TVars.make ctyenv, Vars.make its) cong))), bds) |
52288 | 177 |
end handle Pattern.MATCH => decomp_reify da congs (ct, ctxt) bds)) |
178 |
end; |
|
179 |
||
69593 | 180 |
fun get_nths (t as (Const (\<^const_name>\<open>List.nth\<close>, _) $ vs $ n)) = |
52288 | 181 |
AList.update (op aconv) (t, (vs, n)) |
182 |
| get_nths (t1 $ t2) = get_nths t1 #> get_nths t2 |
|
183 |
| get_nths (Abs (_, _, t')) = get_nths t' |
|
184 |
| get_nths _ = I; |
|
185 |
||
186 |
fun tryeqs [] (ct, ctxt) bds = error "Cannot find the atoms equation" |
|
187 |
| tryeqs (eq :: eqs) (ct, ctxt) bds = (( |
|
51726 | 188 |
let |
59582 | 189 |
val rhs = eq |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd; |
52288 | 190 |
val nths = get_nths rhs []; |
191 |
val (vss, _) = fold_rev (fn (_, (vs, n)) => fn (vss, ns) => |
|
192 |
(insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([], []); |
|
193 |
val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt; |
|
194 |
val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt'; |
|
195 |
val thy = Proof_Context.theory_of ctxt''; |
|
196 |
val vsns_map = vss ~~ vsns; |
|
197 |
val xns_map = fst (split_list nths) ~~ xns; |
|
198 |
val subst = map (fn (nt, xn) => (nt, Var ((xn, 0), fastype_of nt))) xns_map; |
|
199 |
val rhs_P = subst_free subst rhs; |
|
200 |
val (tyenv, tmenv) = Pattern.match thy (rhs_P, Thm.term_of ct) (Vartab.empty, Vartab.empty); |
|
201 |
val sbst = Envir.subst_term (tyenv, tmenv); |
|
202 |
val sbsT = Envir.subst_type tyenv; |
|
59617 | 203 |
val subst_ty = |
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59642
diff
changeset
|
204 |
map (fn (n, (s, t)) => ((n, s), Thm.ctyp_of ctxt'' t)) (Vartab.dest tyenv) |
52288 | 205 |
val tml = Vartab.dest tmenv; |
206 |
val (subst_ns, bds) = fold_map |
|
207 |
(fn (Const _ $ _ $ n, Var (xn0, _)) => fn bds => |
|
208 |
let |
|
209 |
val name = snd (the (AList.lookup (op =) tml xn0)); |
|
210 |
val (idx, bds) = index_of name bds; |
|
59642 | 211 |
in (apply2 (Thm.cterm_of ctxt'') (n, idx |> HOLogic.mk_nat), bds) end) subst bds; |
52288 | 212 |
val subst_vs = |
51726 | 213 |
let |
52288 | 214 |
fun h (Const _ $ (vs as Var (_, lT)) $ _, Var (_, T)) = |
51726 | 215 |
let |
69593 | 216 |
val cns = sbst (Const (\<^const_name>\<open>List.Cons\<close>, T --> lT --> lT)); |
52288 | 217 |
val lT' = sbsT lT; |
218 |
val (bsT, _) = the (AList.lookup Type.could_unify bds lT); |
|
219 |
val vsn = the (AList.lookup (op =) vsns_map vs); |
|
59617 | 220 |
val vs' = fold_rev (fn x => fn xs => cns $ x $xs) bsT (Free (vsn, lT')); |
59642 | 221 |
in apply2 (Thm.cterm_of ctxt'') (vs, vs') end; |
52288 | 222 |
in map h subst end; |
59617 | 223 |
val cts = |
59642 | 224 |
map (fn ((vn, vi), (tT, t)) => apply2 (Thm.cterm_of ctxt'') (Var ((vn, vi), tT), t)) |
59617 | 225 |
(fold (AList.delete (fn (((a : string), _), (b, _)) => a = b)) |
226 |
(map (fn n => (n, 0)) xns) tml); |
|
52288 | 227 |
val substt = |
228 |
let |
|
74282 | 229 |
val ih = Drule.cterm_rule (Thm.instantiate (TVars.make subst_ty, Vars.empty)); |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
52288
diff
changeset
|
230 |
in map (apply2 ih) (subst_ns @ subst_vs @ cts) end; |
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59642
diff
changeset
|
231 |
val th = |
74282 | 232 |
(Drule.instantiate_normalize |
233 |
(TVars.make subst_ty, Vars.make (map (apfst (dest_Var o Thm.term_of)) substt)) eq) |
|
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59642
diff
changeset
|
234 |
RS sym; |
52288 | 235 |
in (hd (Variable.export ctxt'' ctxt [th]), bds) end) |
236 |
handle Pattern.MATCH => tryeqs eqs (ct, ctxt) bds); |
|
237 |
||
238 |
(* looks for the atoms equation and instantiates it with the right number *) |
|
51726 | 239 |
|
52288 | 240 |
fun mk_decompatom eqs (ct, ctxt) bds = (([], fn (_, bds) => |
241 |
let |
|
242 |
val tT = fastype_of (Thm.term_of ct); |
|
243 |
fun isat eq = |
|
51726 | 244 |
let |
59582 | 245 |
val rhs = eq |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd; |
52288 | 246 |
in exists_Const |
69593 | 247 |
(fn (n, ty) => n = \<^const_name>\<open>List.nth\<close> |
52288 | 248 |
andalso AList.defined Type.could_unify bds (domain_type ty)) rhs |
249 |
andalso Type.could_unify (fastype_of rhs, tT) |
|
250 |
end; |
|
251 |
in tryeqs (filter isat eqs) (ct, ctxt) bds end), bds); |
|
252 |
||
253 |
(* Generic reification procedure: *) |
|
254 |
(* creates all needed cong rules and then just uses the theorem synthesis *) |
|
51726 | 255 |
|
52288 | 256 |
fun mk_congs ctxt eqs = |
257 |
let |
|
59582 | 258 |
val fs = fold_rev (fn eq => insert (op =) (eq |> Thm.prop_of |> HOLogic.dest_Trueprop |
52288 | 259 |
|> HOLogic.dest_eq |> fst |> strip_comb |
260 |
|> fst)) eqs []; |
|
261 |
val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl)) fs []; |
|
262 |
val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt; |
|
263 |
val subst = |
|
59617 | 264 |
the o AList.lookup (op =) |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59617
diff
changeset
|
265 |
(map2 (fn T => fn v => (T, Thm.cterm_of ctxt' (Free (v, T)))) tys vs); |
52288 | 266 |
fun prep_eq eq = |
267 |
let |
|
59582 | 268 |
val (_, _ :: vs) = eq |> Thm.prop_of |> HOLogic.dest_Trueprop |
52288 | 269 |
|> HOLogic.dest_eq |> fst |> strip_comb; |
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59642
diff
changeset
|
270 |
val subst = map_filter (fn Var v => SOME (v, subst (#2 v)) | _ => NONE) vs; |
74282 | 271 |
in Thm.instantiate (TVars.empty, Vars.make subst) eq end; |
52288 | 272 |
val (ps, congs) = map_split (mk_congeq ctxt' fs o prep_eq) eqs; |
273 |
val bds = AList.make (K ([], [])) tys; |
|
274 |
in (ps ~~ Variable.export ctxt' ctxt congs, bds) end |
|
275 |
||
276 |
fun conv ctxt eqs ct = |
|
277 |
let |
|
51726 | 278 |
val (congs, bds) = mk_congs ctxt eqs; |
279 |
val congs = rearrange congs; |
|
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
52288
diff
changeset
|
280 |
val (th, bds') = |
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
52288
diff
changeset
|
281 |
apfst mk_eq (divide_and_conquer' (decomp_reify (mk_decompatom eqs) congs) (ct, ctxt) bds); |
52273 | 282 |
fun is_list_var (Var (_, t)) = can dest_listT t |
283 |
| is_list_var _ = false; |
|
59582 | 284 |
val vars = th |> Thm.prop_of |> Logic.dest_equals |> snd |
52273 | 285 |
|> strip_comb |> snd |> filter is_list_var; |
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59642
diff
changeset
|
286 |
val vs = map (fn Var (v as (_, T)) => |
52275 | 287 |
(v, the (AList.lookup Type.could_unify bds' T) |> snd |> HOLogic.mk_list (dest_listT T))) vars; |
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59642
diff
changeset
|
288 |
val th' = |
74282 | 289 |
Drule.instantiate_normalize (TVars.empty, Vars.make (map (apsnd (Thm.cterm_of ctxt)) vs)) th; |
52275 | 290 |
val th'' = Thm.symmetric (dereify ctxt [] (Thm.lhs_of th')); |
291 |
in Thm.transitive th'' th' end; |
|
51726 | 292 |
|
52286 | 293 |
fun tac ctxt eqs = |
294 |
lift_conv ctxt (conv ctxt eqs); |
|
51726 | 295 |
|
52276 | 296 |
end; |