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 |
|
52275
|
18 |
fun dest_listT (Type (@{type_name "list"}, [T])) = T;
|
|
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 |
|
52276
|
26 |
fun lift_conv ctxt conv some_t = Subgoal.FOCUS (fn { concl, ... } =>
|
52275
|
27 |
let
|
|
28 |
val ct = case some_t
|
|
29 |
of NONE => Thm.dest_arg concl
|
|
30 |
| SOME t => Thm.cterm_of (Proof_Context.theory_of ctxt) t
|
|
31 |
val thm = conv ct;
|
|
32 |
in
|
|
33 |
if Thm.is_reflexive thm then no_tac
|
|
34 |
else ALLGOALS (rtac (pure_subst OF [thm]))
|
|
35 |
end) ctxt;
|
51726
|
36 |
|
|
37 |
(* Make a congruence rule out of a defining equation for the interpretation
|
|
38 |
|
|
39 |
th is one defining equation of f,
|
|
40 |
i.e. th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)"
|
|
41 |
Cp is a constructor pattern and P is a pattern
|
|
42 |
|
|
43 |
The result is:
|
|
44 |
[|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn)
|
|
45 |
+ the a list of names of the A1 .. An, Those are fresh in the ctxt *)
|
|
46 |
|
|
47 |
fun mk_congeq ctxt fs th =
|
|
48 |
let
|
|
49 |
val Const (fN, _) = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq
|
|
50 |
|> fst |> strip_comb |> fst;
|
52273
|
51 |
val cert = Thm.cterm_of (Proof_Context.theory_of ctxt);
|
51726
|
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, ...} =>
|
|
85 |
Local_Defs.unfold_tac context (map tryext prems) THEN rtac th' 1)) RS sym;
|
|
86 |
val (cong' :: vars') =
|
|
87 |
Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs);
|
|
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) =
|
|
98 |
let val @{term "Trueprop"} $ (Const (@{const_name HOL.eq}, _) $ l $ _) = concl_of th
|
|
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 |
val cert = cterm_of thy;
|
|
138 |
val certT = ctyp_of thy;
|
|
139 |
fun tryabsdecomp (ct, ctxt) bds =
|
|
140 |
(case Thm.term_of ct of
|
|
141 |
Abs (_, xT, ta) =>
|
|
142 |
let
|
|
143 |
val ([raw_xn], ctxt') = Variable.variant_fixes ["x"] ctxt;
|
|
144 |
val (xn, ta) = Syntax_Trans.variant_abs (raw_xn, xT, ta); (* FIXME !? *)
|
|
145 |
val x = Free (xn, xT);
|
|
146 |
val cx = cert x;
|
|
147 |
val cta = cert ta;
|
|
148 |
val bds = (case AList.lookup Type.could_unify bds (HOLogic.listT xT) of
|
|
149 |
NONE => error "tryabsdecomp: Type not found in the Environement"
|
|
150 |
| SOME (bsT, atsT) => AList.update Type.could_unify (HOLogic.listT xT,
|
|
151 |
(x :: bsT, atsT)) bds);
|
|
152 |
in (([(cta, ctxt')],
|
|
153 |
fn ([th], bds) =>
|
|
154 |
(hd (Variable.export ctxt' ctxt [(Thm.forall_intr cx th) COMP allI]),
|
|
155 |
let
|
|
156 |
val (bsT, asT) = the (AList.lookup Type.could_unify bds (HOLogic.listT xT));
|
|
157 |
in
|
|
158 |
AList.update Type.could_unify (HOLogic.listT xT, (tl bsT, asT)) bds
|
|
159 |
end)),
|
|
160 |
bds)
|
|
161 |
end
|
|
162 |
| _ => da (ct, ctxt) bds)
|
|
163 |
in
|
|
164 |
(case cgns of
|
|
165 |
[] => tryabsdecomp (ct, ctxt) bds
|
|
166 |
| ((vns, cong) :: congs) =>
|
|
167 |
(let
|
|
168 |
val (tyenv, tmenv) =
|
|
169 |
Pattern.match thy
|
|
170 |
((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), Thm.term_of ct)
|
|
171 |
(Vartab.empty, Vartab.empty);
|
|
172 |
val (fnvs, invs) = List.partition (fn ((vn, _),_) => member (op =) vns vn) (Vartab.dest tmenv);
|
|
173 |
val (fts, its) =
|
|
174 |
(map (snd o snd) fnvs,
|
|
175 |
map (fn ((vn, vi), (tT, t)) => (cert (Var ((vn, vi), tT)), cert t)) invs);
|
|
176 |
val ctyenv = map (fn ((vn, vi), (s, ty)) => (certT (TVar((vn, vi), s)), certT ty)) (Vartab.dest tyenv);
|
|
177 |
in
|
|
178 |
((map cert fts ~~ replicate (length fts) ctxt,
|
|
179 |
apfst (FWD (Drule.instantiate_normalize (ctyenv, its) cong))), bds)
|
|
180 |
end handle Pattern.MATCH => decomp_reify da congs (ct, ctxt) bds))
|
|
181 |
end;
|
|
182 |
|
|
183 |
fun get_nths (t as (Const (@{const_name "List.nth"}, _) $ vs $ n)) =
|
|
184 |
AList.update (op aconv) (t, (vs, n))
|
|
185 |
| get_nths (t1 $ t2) = get_nths t1 #> get_nths t2
|
|
186 |
| get_nths (Abs (_, _, t')) = get_nths t'
|
|
187 |
| get_nths _ = I;
|
|
188 |
|
|
189 |
fun tryeqs [] (ct, ctxt) bds = error "Cannot find the atoms equation"
|
|
190 |
| tryeqs (eq :: eqs) (ct, ctxt) bds = ((
|
51726
|
191 |
let
|
52288
|
192 |
val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd;
|
|
193 |
val nths = get_nths rhs [];
|
|
194 |
val (vss, _) = fold_rev (fn (_, (vs, n)) => fn (vss, ns) =>
|
|
195 |
(insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([], []);
|
|
196 |
val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt;
|
|
197 |
val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt';
|
|
198 |
val thy = Proof_Context.theory_of ctxt'';
|
51726
|
199 |
val cert = cterm_of thy;
|
52273
|
200 |
val certT = ctyp_of thy;
|
52288
|
201 |
val vsns_map = vss ~~ vsns;
|
|
202 |
val xns_map = fst (split_list nths) ~~ xns;
|
|
203 |
val subst = map (fn (nt, xn) => (nt, Var ((xn, 0), fastype_of nt))) xns_map;
|
|
204 |
val rhs_P = subst_free subst rhs;
|
|
205 |
val (tyenv, tmenv) = Pattern.match thy (rhs_P, Thm.term_of ct) (Vartab.empty, Vartab.empty);
|
|
206 |
val sbst = Envir.subst_term (tyenv, tmenv);
|
|
207 |
val sbsT = Envir.subst_type tyenv;
|
|
208 |
val subst_ty = map (fn (n, (s, t)) =>
|
|
209 |
(certT (TVar (n, s)), certT t)) (Vartab.dest tyenv)
|
|
210 |
val tml = Vartab.dest tmenv;
|
|
211 |
val (subst_ns, bds) = fold_map
|
|
212 |
(fn (Const _ $ _ $ n, Var (xn0, _)) => fn bds =>
|
|
213 |
let
|
|
214 |
val name = snd (the (AList.lookup (op =) tml xn0));
|
|
215 |
val (idx, bds) = index_of name bds;
|
|
216 |
in ((cert n, idx |> (HOLogic.mk_nat #> cert)), bds) end) subst bds;
|
|
217 |
val subst_vs =
|
51726
|
218 |
let
|
52288
|
219 |
fun h (Const _ $ (vs as Var (_, lT)) $ _, Var (_, T)) =
|
51726
|
220 |
let
|
52288
|
221 |
val cns = sbst (Const (@{const_name "List.Cons"}, T --> lT --> lT));
|
|
222 |
val lT' = sbsT lT;
|
|
223 |
val (bsT, _) = the (AList.lookup Type.could_unify bds lT);
|
|
224 |
val vsn = the (AList.lookup (op =) vsns_map vs);
|
|
225 |
val cvs = cert (fold_rev (fn x => fn xs => cns $ x $xs) bsT (Free (vsn, lT')));
|
|
226 |
in (cert vs, cvs) end;
|
|
227 |
in map h subst end;
|
|
228 |
val cts = map (fn ((vn, vi), (tT, t)) => (cert (Var ((vn, vi), tT)), cert t))
|
|
229 |
(fold (AList.delete (fn (((a : string), _), (b, _)) => a = b))
|
|
230 |
(map (fn n => (n, 0)) xns) tml);
|
|
231 |
val substt =
|
|
232 |
let
|
|
233 |
val ih = Drule.cterm_rule (Thm.instantiate (subst_ty, []));
|
|
234 |
in map (pairself ih) (subst_ns @ subst_vs @ cts) end;
|
|
235 |
val th = (Drule.instantiate_normalize (subst_ty, substt) eq) RS sym;
|
|
236 |
in (hd (Variable.export ctxt'' ctxt [th]), bds) end)
|
|
237 |
handle Pattern.MATCH => tryeqs eqs (ct, ctxt) bds);
|
|
238 |
|
|
239 |
(* looks for the atoms equation and instantiates it with the right number *)
|
51726
|
240 |
|
52288
|
241 |
fun mk_decompatom eqs (ct, ctxt) bds = (([], fn (_, bds) =>
|
|
242 |
let
|
|
243 |
val tT = fastype_of (Thm.term_of ct);
|
|
244 |
fun isat eq =
|
51726
|
245 |
let
|
52288
|
246 |
val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd;
|
|
247 |
in exists_Const
|
|
248 |
(fn (n, ty) => n = @{const_name "List.nth"}
|
|
249 |
andalso AList.defined Type.could_unify bds (domain_type ty)) rhs
|
|
250 |
andalso Type.could_unify (fastype_of rhs, tT)
|
|
251 |
end;
|
|
252 |
in tryeqs (filter isat eqs) (ct, ctxt) bds end), bds);
|
|
253 |
|
|
254 |
(* Generic reification procedure: *)
|
|
255 |
(* creates all needed cong rules and then just uses the theorem synthesis *)
|
51726
|
256 |
|
52288
|
257 |
fun mk_congs ctxt eqs =
|
|
258 |
let
|
|
259 |
val fs = fold_rev (fn eq => insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop
|
|
260 |
|> HOLogic.dest_eq |> fst |> strip_comb
|
|
261 |
|> fst)) eqs [];
|
|
262 |
val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl)) fs [];
|
|
263 |
val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt;
|
|
264 |
val cert = cterm_of (Proof_Context.theory_of ctxt');
|
|
265 |
val subst =
|
|
266 |
the o AList.lookup (op =) (map2 (fn T => fn v => (T, cert (Free (v, T)))) tys vs);
|
|
267 |
fun prep_eq eq =
|
|
268 |
let
|
|
269 |
val (_, _ :: vs) = eq |> prop_of |> HOLogic.dest_Trueprop
|
|
270 |
|> HOLogic.dest_eq |> fst |> strip_comb;
|
|
271 |
val subst = map_filter (fn (v as Var (_, T)) => SOME (cert v, subst T)
|
|
272 |
| _ => NONE) vs;
|
|
273 |
in Thm.instantiate ([], subst) eq end;
|
|
274 |
val (ps, congs) = map_split (mk_congeq ctxt' fs o prep_eq) eqs;
|
|
275 |
val bds = AList.make (K ([], [])) tys;
|
|
276 |
in (ps ~~ Variable.export ctxt' ctxt congs, bds) end
|
|
277 |
|
|
278 |
fun conv ctxt eqs ct =
|
|
279 |
let
|
51726
|
280 |
val (congs, bds) = mk_congs ctxt eqs;
|
|
281 |
val congs = rearrange congs;
|
52275
|
282 |
val (th, bds') = apfst mk_eq (divide_and_conquer' (decomp_reify (mk_decompatom eqs) congs) (ct, ctxt) bds);
|
52273
|
283 |
fun is_list_var (Var (_, t)) = can dest_listT t
|
|
284 |
| is_list_var _ = false;
|
52275
|
285 |
val vars = th |> prop_of |> Logic.dest_equals |> snd
|
52273
|
286 |
|> strip_comb |> snd |> filter is_list_var;
|
51726
|
287 |
val cert = cterm_of (Proof_Context.theory_of ctxt);
|
52275
|
288 |
val vs = map (fn v as Var (_, T) =>
|
|
289 |
(v, the (AList.lookup Type.could_unify bds' T) |> snd |> HOLogic.mk_list (dest_listT T))) vars;
|
|
290 |
val th' = Drule.instantiate_normalize ([], (map o pairself) cert vs) th;
|
|
291 |
val th'' = Thm.symmetric (dereify ctxt [] (Thm.lhs_of th'));
|
|
292 |
in Thm.transitive th'' th' end;
|
51726
|
293 |
|
52286
|
294 |
fun tac ctxt eqs =
|
|
295 |
lift_conv ctxt (conv ctxt eqs);
|
51726
|
296 |
|
52276
|
297 |
end;
|