author | paulson |
Thu, 06 Sep 2007 17:03:32 +0200 | |
changeset 24546 | c90cee3163b7 |
parent 23531 | 38a304b3fe1e |
child 24570 | 621b60b1df00 |
permissions | -rw-r--r-- |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
1 |
(* Title: HOL/Nominal/nominal_inductive.ML |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
2 |
ID: $Id$ |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
3 |
Author: Stefan Berghofer, TU Muenchen |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
4 |
|
22530 | 5 |
Infrastructure for proving equivariance and strong induction theorems |
6 |
for inductive predicates involving nominal datatypes. |
|
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
7 |
*) |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
8 |
|
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
9 |
signature NOMINAL_INDUCTIVE = |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
10 |
sig |
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
11 |
val prove_strong_ind: string -> (string * string list) list -> theory -> Proof.state |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
12 |
val prove_eqvt: string -> string list -> theory -> theory |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
13 |
end |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
14 |
|
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
15 |
structure NominalInductive : NOMINAL_INDUCTIVE = |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
16 |
struct |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
17 |
|
22530 | 18 |
val finite_Un = thm "finite_Un"; |
19 |
val supp_prod = thm "supp_prod"; |
|
20 |
val fresh_prod = thm "fresh_prod"; |
|
21 |
||
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
22 |
val perm_boolI = thm "perm_boolI"; |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
23 |
val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
24 |
(Drule.strip_imp_concl (cprop_of perm_boolI)))); |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
25 |
|
22530 | 26 |
val allE_Nil = read_instantiate_sg (the_context()) [("x", "[]")] allE; |
27 |
||
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
28 |
fun transp ([] :: _) = [] |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
29 |
| transp xs = map hd xs :: transp (map tl xs); |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
30 |
|
22530 | 31 |
fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of |
32 |
(Const (s, T), ts) => (case strip_type T of |
|
33 |
(Ts, Type (tname, _)) => |
|
34 |
(case NominalPackage.get_nominal_datatype thy tname of |
|
35 |
NONE => fold (add_binders thy i) ts bs |
|
36 |
| SOME {descr, index, ...} => (case AList.lookup op = |
|
37 |
(#3 (the (AList.lookup op = descr index))) s of |
|
38 |
NONE => fold (add_binders thy i) ts bs |
|
39 |
| SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') => |
|
40 |
let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs' |
|
41 |
in (add_binders thy i u |
|
42 |
(fold (fn (u, T) => |
|
43 |
if exists (fn j => j < i) (loose_bnos u) then I |
|
44 |
else insert (op aconv o pairself fst) |
|
45 |
(incr_boundvars (~i) u, T)) cargs1 bs'), cargs2) |
|
46 |
end) cargs (bs, ts ~~ Ts)))) |
|
47 |
| _ => fold (add_binders thy i) ts bs) |
|
48 |
| (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs)) |
|
49 |
| add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs |
|
50 |
| add_binders thy i _ bs = bs; |
|
51 |
||
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
52 |
fun prove_strong_ind s avoids thy = |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
53 |
let |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
54 |
val ctxt = ProofContext.init thy; |
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
55 |
val ({names, ...}, {raw_induct, ...}) = |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
56 |
InductivePackage.the_inductive ctxt (Sign.intern_const thy s); |
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
57 |
val eqvt_thms = NominalThmDecls.get_eqvt_thms thy; |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
58 |
val _ = (case names \\ foldl (apfst prop_of #> add_term_consts) [] eqvt_thms of |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
59 |
[] => () |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
60 |
| xs => error ("Missing equivariance theorem for predicate(s): " ^ |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
61 |
commas_quote xs)); |
22530 | 62 |
val induct_cases = map fst (fst (RuleCases.get (the |
63 |
(InductAttrib.lookup_inductS ctxt (hd names))))); |
|
64 |
val raw_induct' = Logic.unvarify (prop_of raw_induct); |
|
65 |
val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> |
|
66 |
HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb); |
|
67 |
val ps = map (fst o snd) concls; |
|
68 |
||
69 |
val _ = (case duplicates (op = o pairself fst) avoids of |
|
70 |
[] => () |
|
71 |
| xs => error ("Duplicate case names: " ^ commas_quote (map fst xs))); |
|
72 |
val _ = assert_all (null o duplicates op = o snd) avoids |
|
73 |
(fn (a, _) => error ("Duplicate variable names for case " ^ quote a)); |
|
74 |
val _ = (case map fst avoids \\ induct_cases of |
|
75 |
[] => () |
|
76 |
| xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)); |
|
77 |
val avoids' = map (fn name => |
|
78 |
(name, the_default [] (AList.lookup op = avoids name))) induct_cases; |
|
79 |
fun mk_avoids params (name, ps) = |
|
80 |
let val k = length params - 1 |
|
81 |
in map (fn x => case find_index (equal x o fst) params of |
|
82 |
~1 => error ("No such variable in case " ^ quote name ^ |
|
83 |
" of inductive definition: " ^ quote x) |
|
84 |
| i => (Bound (k - i), snd (nth params i))) ps |
|
85 |
end; |
|
86 |
||
87 |
val prems = map (fn (prem, avoid) => |
|
88 |
let |
|
89 |
val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem); |
|
90 |
val concl = incr_boundvars 1 (Logic.strip_assums_concl prem); |
|
91 |
val params = Logic.strip_params prem |
|
92 |
in |
|
93 |
(params, |
|
94 |
fold (add_binders thy 0) (prems @ [concl]) [] @ |
|
95 |
map (apfst (incr_boundvars 1)) (mk_avoids params avoid), |
|
96 |
prems, strip_comb (HOLogic.dest_Trueprop concl)) |
|
97 |
end) (Logic.strip_imp_prems raw_induct' ~~ avoids'); |
|
98 |
||
99 |
val atomTs = distinct op = (maps (map snd o #2) prems); |
|
100 |
val ind_sort = if null atomTs then HOLogic.typeS |
|
101 |
else Sign.certify_sort thy (map (fn T => Sign.intern_class thy |
|
102 |
("fs_" ^ Sign.base_name (fst (dest_Type T)))) atomTs); |
|
103 |
val fs_ctxt_tyname = Name.variant (map fst (term_tfrees raw_induct')) "'n"; |
|
104 |
val fs_ctxt_name = Name.variant (add_term_names (raw_induct', [])) "z"; |
|
105 |
val fsT = TFree (fs_ctxt_tyname, ind_sort); |
|
106 |
||
107 |
fun lift_pred' t (Free (s, T)) ts = |
|
108 |
list_comb (Free (s, fsT --> T), t :: ts); |
|
109 |
val lift_pred = lift_pred' (Bound 0); |
|
110 |
||
111 |
fun lift_prem (Const ("Trueprop", _) $ t) = |
|
112 |
let val (u, ts) = strip_comb t |
|
113 |
in |
|
114 |
if u mem ps then |
|
115 |
all fsT $ Abs ("z", fsT, HOLogic.mk_Trueprop |
|
116 |
(lift_pred u (map (incr_boundvars 1) ts))) |
|
117 |
else HOLogic.mk_Trueprop (lift_prem t) |
|
118 |
end |
|
119 |
| lift_prem (t as (f $ u)) = |
|
120 |
let val (p, ts) = strip_comb t |
|
121 |
in |
|
122 |
if p mem ps then |
|
123 |
HOLogic.all_const fsT $ Abs ("z", fsT, |
|
124 |
lift_pred p (map (incr_boundvars 1) ts)) |
|
125 |
else lift_prem f $ lift_prem u |
|
126 |
end |
|
127 |
| lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t) |
|
128 |
| lift_prem t = t; |
|
129 |
||
130 |
fun mk_distinct [] = [] |
|
131 |
| mk_distinct ((x, T) :: xs) = List.mapPartial (fn (y, U) => |
|
132 |
if T = U then SOME (HOLogic.mk_Trueprop |
|
133 |
(HOLogic.mk_not (HOLogic.eq_const T $ x $ y))) |
|
134 |
else NONE) xs @ mk_distinct xs; |
|
135 |
||
136 |
fun mk_fresh (x, T) = HOLogic.mk_Trueprop |
|
137 |
(Const ("Nominal.fresh", T --> fsT --> HOLogic.boolT) $ x $ Bound 0); |
|
138 |
||
139 |
val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) => |
|
140 |
let |
|
141 |
val params' = params @ [("y", fsT)]; |
|
142 |
val prem = Logic.list_implies |
|
143 |
(map mk_fresh bvars @ mk_distinct bvars @ |
|
144 |
map (fn prem => |
|
145 |
if null (term_frees prem inter ps) then prem |
|
146 |
else lift_prem prem) prems, |
|
147 |
HOLogic.mk_Trueprop (lift_pred p ts)); |
|
148 |
val vs = map (Var o apfst (rpair 0)) (rename_wrt_term prem params') |
|
149 |
in |
|
150 |
(list_all (params', prem), (rev vs, subst_bounds (vs, prem))) |
|
151 |
end) prems); |
|
152 |
||
153 |
val ind_vars = |
|
154 |
(DatatypeProp.indexify_names (replicate (length atomTs) "pi") ~~ |
|
155 |
map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; |
|
156 |
val ind_Ts = rev (map snd ind_vars); |
|
157 |
||
158 |
val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
|
159 |
(map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, |
|
160 |
HOLogic.list_all (ind_vars, lift_pred p |
|
161 |
(map (fold_rev (NominalPackage.mk_perm ind_Ts) |
|
162 |
(map Bound (length atomTs downto 1))) ts)))) concls)); |
|
163 |
||
164 |
val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
|
165 |
(map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, |
|
166 |
lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls)); |
|
167 |
||
168 |
val vc_compat = map (fn (params, bvars, prems, (p, ts)) => |
|
169 |
map (fn q => list_all (params, incr_boundvars ~1 (Logic.list_implies |
|
170 |
(filter (fn prem => null (ps inter term_frees prem)) prems, q)))) |
|
171 |
(mk_distinct bvars @ |
|
172 |
maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop |
|
173 |
(Const ("Nominal.fresh", U --> T --> HOLogic.boolT) $ u $ t)) bvars) |
|
174 |
(ts ~~ binder_types (fastype_of p)))) prems; |
|
175 |
||
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
176 |
val eqvt_ss = HOL_basic_ss addsimps eqvt_thms; |
22530 | 177 |
val fresh_bij = PureThy.get_thms thy (Name "fresh_bij"); |
178 |
val perm_bij = PureThy.get_thms thy (Name "perm_bij"); |
|
179 |
val fs_atoms = map (fn aT => PureThy.get_thm thy |
|
180 |
(Name ("fs_" ^ Sign.base_name (fst (dest_Type aT)) ^ "1"))) atomTs; |
|
181 |
val exists_fresh' = PureThy.get_thms thy (Name "exists_fresh'"); |
|
182 |
val fresh_atm = PureThy.get_thms thy (Name "fresh_atm"); |
|
183 |
val calc_atm = PureThy.get_thms thy (Name "calc_atm"); |
|
184 |
val perm_fresh_fresh = PureThy.get_thms thy (Name "perm_fresh_fresh"); |
|
185 |
val pt2_atoms = map (fn aT => PureThy.get_thm thy |
|
186 |
(Name ("pt_" ^ Sign.base_name (fst (dest_Type aT)) ^ "2")) RS sym) atomTs; |
|
187 |
||
188 |
fun obtain_fresh_name ts T (freshs1, freshs2, ctxt) = |
|
189 |
let |
|
190 |
(** protect terms to avoid that supp_prod interferes with **) |
|
191 |
(** pairs used in introduction rules of inductive predicate **) |
|
192 |
fun protect t = |
|
193 |
let val T = fastype_of t in Const ("Fun.id", T --> T) $ t end; |
|
194 |
val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1); |
|
195 |
val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop |
|
196 |
(HOLogic.exists_const T $ Abs ("x", T, |
|
197 |
Const ("Nominal.fresh", T --> fastype_of p --> HOLogic.boolT) $ |
|
198 |
Bound 0 $ p))) |
|
199 |
(fn _ => EVERY |
|
200 |
[resolve_tac exists_fresh' 1, |
|
201 |
simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]); |
|
202 |
val (([cx], ths), ctxt') = Obtain.result |
|
203 |
(fn _ => EVERY |
|
204 |
[etac exE 1, |
|
205 |
full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1, |
|
206 |
full_simp_tac (HOL_basic_ss addsimps [id_apply]) 1, |
|
207 |
REPEAT (etac conjE 1)]) |
|
208 |
[ex] ctxt |
|
209 |
in (freshs1 @ [term_of cx], freshs2 @ ths, ctxt') end; |
|
210 |
||
211 |
fun mk_proof thy thss = |
|
212 |
let val ctxt = ProofContext.init thy |
|
213 |
in Goal.prove_global thy [] prems' concl' (fn ihyps => |
|
214 |
let val th = Goal.prove ctxt [] [] concl (fn {context, ...} => |
|
215 |
rtac raw_induct 1 THEN |
|
216 |
EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) => |
|
217 |
[REPEAT (rtac allI 1), simp_tac eqvt_ss 1, |
|
218 |
SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} => |
|
219 |
let |
|
220 |
val (params', (pis, z)) = |
|
221 |
chop (length params - length atomTs - 1) (map term_of params) ||> |
|
222 |
split_last; |
|
223 |
val bvars' = map |
|
224 |
(fn (Bound i, T) => (nth params' (length params' - i), T) |
|
225 |
| (t, T) => (t, T)) bvars; |
|
226 |
val pi_bvars = map (fn (t, _) => |
|
227 |
fold_rev (NominalPackage.mk_perm []) pis t) bvars'; |
|
228 |
val (P, ts) = strip_comb (HOLogic.dest_Trueprop (term_of concl)); |
|
229 |
val (freshs1, freshs2, ctxt'') = fold |
|
230 |
(obtain_fresh_name (ts @ pi_bvars)) |
|
231 |
(map snd bvars') ([], [], ctxt'); |
|
232 |
val freshs2' = NominalPackage.mk_not_sym freshs2; |
|
233 |
val pis' = map NominalPackage.perm_of_pair (pi_bvars ~~ freshs1); |
|
234 |
val env = Pattern.first_order_match thy (ihypt, prop_of ihyp) |
|
235 |
(Vartab.empty, Vartab.empty); |
|
236 |
val ihyp' = Thm.instantiate ([], map (pairself (cterm_of thy)) |
|
237 |
(map (Envir.subst_vars env) vs ~~ |
|
238 |
map (fold_rev (NominalPackage.mk_perm []) |
|
239 |
(rev pis' @ pis)) params' @ [z])) ihyp; |
|
240 |
val (gprems1, gprems2) = pairself (map fst) (List.partition |
|
241 |
(fn (th, t) => null (term_frees t inter ps)) (gprems ~~ oprems)); |
|
242 |
val vc_compat_ths' = map (fn th => |
|
243 |
let |
|
244 |
val th' = gprems1 MRS |
|
22901 | 245 |
Thm.instantiate (Thm.first_order_match |
23531 | 246 |
(Conjunction.mk_conjunction_balanced (cprems_of th), |
247 |
Conjunction.mk_conjunction_balanced (map cprop_of gprems1))) th; |
|
22530 | 248 |
val (bop, lhs, rhs) = (case concl_of th' of |
249 |
_ $ (fresh $ lhs $ rhs) => |
|
250 |
(fn t => fn u => fresh $ t $ u, lhs, rhs) |
|
251 |
| _ $ (_ $ (_ $ lhs $ rhs)) => |
|
252 |
(curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs)); |
|
253 |
val th'' = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop |
|
254 |
(bop (fold_rev (NominalPackage.mk_perm []) pis lhs) |
|
255 |
(fold_rev (NominalPackage.mk_perm []) pis rhs))) |
|
256 |
(fn _ => simp_tac (HOL_basic_ss addsimps |
|
257 |
(fresh_bij @ perm_bij)) 1 THEN rtac th' 1) |
|
258 |
in Simplifier.simplify (eqvt_ss addsimps fresh_atm) th'' end) |
|
259 |
vc_compat_ths; |
|
260 |
val vc_compat_ths'' = NominalPackage.mk_not_sym vc_compat_ths'; |
|
261 |
val gprems1' = map (fn th => fold_rev (fn pi => fn th' => |
|
262 |
Simplifier.simplify eqvt_ss (th' RS Drule.cterm_instantiate |
|
263 |
[(perm_boolI_pi, cterm_of thy pi)] perm_boolI)) |
|
264 |
(rev pis' @ pis) th) gprems1; |
|
265 |
val gprems2' = map (Simplifier.simplify eqvt_ss) gprems2; |
|
266 |
(** Since calc_atm simplifies (pi :: 'a prm) o (x :: 'b) to x **) |
|
267 |
(** we have to pre-simplify the rewrite rules **) |
|
268 |
val calc_atm_ss = HOL_ss addsimps calc_atm @ |
|
269 |
map (Simplifier.simplify (HOL_ss addsimps calc_atm)) |
|
270 |
(vc_compat_ths'' @ freshs2'); |
|
271 |
val th = Goal.prove ctxt'' [] [] |
|
272 |
(HOLogic.mk_Trueprop (list_comb (P $ hd ts, |
|
273 |
map (fold (NominalPackage.mk_perm []) pis') (tl ts)))) |
|
274 |
(fn _ => EVERY ([simp_tac eqvt_ss 1, rtac ihyp' 1, |
|
275 |
REPEAT_DETERM_N (nprems_of ihyp - length gprems) |
|
276 |
(simp_tac calc_atm_ss 1), |
|
277 |
REPEAT_DETERM_N (length gprems) |
|
278 |
(resolve_tac gprems1' 1 ORELSE |
|
279 |
simp_tac (HOL_basic_ss addsimps pt2_atoms @ gprems2' |
|
280 |
addsimprocs [NominalPackage.perm_simproc]) 1)])); |
|
281 |
val final = Goal.prove ctxt'' [] [] (term_of concl) |
|
282 |
(fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (HOL_ss |
|
283 |
addsimps vc_compat_ths'' @ freshs2' @ |
|
284 |
perm_fresh_fresh @ fresh_atm) 1); |
|
285 |
val final' = ProofContext.export ctxt'' ctxt' [final]; |
|
286 |
in resolve_tac final' 1 end) context 1]) |
|
287 |
(prems ~~ thss ~~ ihyps ~~ prems''))) |
|
288 |
in |
|
289 |
cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN |
|
290 |
REPEAT (REPEAT (resolve_tac [conjI, impI] 1) THEN |
|
291 |
etac impE 1 THEN atac 1 THEN REPEAT (etac allE_Nil 1) THEN |
|
292 |
asm_full_simp_tac (simpset_of thy) 1) |
|
293 |
end) |
|
294 |
end; |
|
295 |
||
296 |
in |
|
297 |
thy |> |
|
298 |
ProofContext.init |> |
|
299 |
Proof.theorem_i NONE (fn thss => ProofContext.theory (fn thy => |
|
300 |
let |
|
301 |
val ctxt = ProofContext.init thy; |
|
302 |
val rec_name = space_implode "_" (map Sign.base_name names); |
|
303 |
val ind_case_names = RuleCases.case_names induct_cases; |
|
304 |
val strong_raw_induct = mk_proof thy thss; |
|
305 |
val strong_induct = |
|
306 |
if length names > 1 then |
|
307 |
(strong_raw_induct, [ind_case_names, RuleCases.consumes 0]) |
|
308 |
else (strong_raw_induct RSN (2, rev_mp), |
|
309 |
[ind_case_names, RuleCases.consumes 1]); |
|
310 |
val ([strong_induct'], thy') = thy |> |
|
311 |
Theory.add_path rec_name |> |
|
312 |
PureThy.add_thms [(("strong_induct", #1 strong_induct), #2 strong_induct)]; |
|
313 |
val strong_inducts = |
|
314 |
ProjectRule.projects ctxt (1 upto length names) strong_induct' |
|
315 |
in |
|
316 |
thy' |> |
|
317 |
PureThy.add_thmss [(("strong_inducts", strong_inducts), |
|
318 |
[ind_case_names, RuleCases.consumes 1])] |> snd |> |
|
319 |
Theory.parent_path |
|
320 |
end)) |
|
321 |
(map (map (rpair [])) vc_compat) |
|
322 |
end; |
|
323 |
||
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
324 |
fun prove_eqvt s xatoms thy = |
22530 | 325 |
let |
326 |
val ctxt = ProofContext.init thy; |
|
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
327 |
val ({names, ...}, {raw_induct, intrs, elims, ...}) = |
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
328 |
InductivePackage.the_inductive ctxt (Sign.intern_const thy s); |
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
329 |
val intrs' = InductivePackage.unpartition_rules intrs |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
330 |
(map (fn (((s, ths), (_, k)), th) => |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
331 |
(s, ths ~~ InductivePackage.infer_intro_vars th k ths)) |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
332 |
(InductivePackage.partition_rules raw_induct intrs ~~ |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
333 |
InductivePackage.arities_of raw_induct ~~ elims)); |
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
334 |
val atoms' = NominalAtoms.atoms_of thy; |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
335 |
val atoms = |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
336 |
if null xatoms then atoms' else |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
337 |
let val atoms = map (Sign.intern_type thy) xatoms |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
338 |
in |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
339 |
(case duplicates op = atoms of |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
340 |
[] => () |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
341 |
| xs => error ("Duplicate atoms: " ^ commas xs); |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
342 |
case atoms \\ atoms' of |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
343 |
[] => () |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
344 |
| xs => error ("No such atoms: " ^ commas xs); |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
345 |
atoms) |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
346 |
end; |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
347 |
val eqvt_ss = HOL_basic_ss addsimps NominalThmDecls.get_eqvt_thms thy; |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
348 |
val t = Logic.unvarify (concl_of raw_induct); |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
349 |
val pi = Name.variant (add_term_names (t, [])) "pi"; |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
350 |
val ps = map (fst o HOLogic.dest_imp) |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
351 |
(HOLogic.dest_conj (HOLogic.dest_Trueprop t)); |
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
352 |
fun eqvt_tac th pi (intr, vs) st = |
22544 | 353 |
let |
354 |
fun eqvt_err s = error |
|
355 |
("Could not prove equivariance for introduction rule\n" ^ |
|
356 |
Sign.string_of_term (theory_of_thm intr) |
|
357 |
(Logic.unvarify (prop_of intr)) ^ "\n" ^ s); |
|
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
358 |
val res = SUBPROOF (fn {prems, params, ...} => |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
359 |
let |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
360 |
val prems' = map (fn th' => Simplifier.simplify eqvt_ss |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
361 |
(if null (names inter term_consts (prop_of th')) then th' RS th |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
362 |
else th')) prems; |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
363 |
val intr' = Drule.cterm_instantiate (map (cterm_of thy) vs ~~ |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
364 |
map (cterm_of thy o NominalPackage.mk_perm [] pi o term_of) params) |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
365 |
intr |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
366 |
in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac prems')) 1 |
22544 | 367 |
end) ctxt 1 st |
368 |
in |
|
369 |
case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of |
|
370 |
NONE => eqvt_err ("Rule does not match goal\n" ^ |
|
371 |
Sign.string_of_term (theory_of_thm st) (hd (prems_of st))) |
|
372 |
| SOME (th, _) => Seq.single th |
|
373 |
end; |
|
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
374 |
val thss = map (fn atom => |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
375 |
let |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
376 |
val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, []))); |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
377 |
val perm_boolI' = Drule.cterm_instantiate |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
378 |
[(perm_boolI_pi, cterm_of thy pi')] perm_boolI |
22530 | 379 |
in map (fn th => zero_var_indexes (th RS mp)) |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
380 |
(DatatypeAux.split_conj_thm (Goal.prove_global thy [] [] |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
381 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p => |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
382 |
HOLogic.mk_imp (p, list_comb |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
383 |
(apsnd (map (NominalPackage.mk_perm [] pi')) (strip_comb p)))) ps))) |
22788
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
384 |
(fn _ => EVERY (rtac raw_induct 1 :: map (fn intr_vs => |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
385 |
full_simp_tac eqvt_ss 1 THEN |
3038bd211582
eqvt_tac now instantiates introduction rules before applying them.
berghofe
parents:
22755
diff
changeset
|
386 |
eqvt_tac perm_boolI' pi' intr_vs) intrs')))) |
22544 | 387 |
end) atoms |
388 |
in |
|
389 |
fold (fn (name, ths) => |
|
390 |
Theory.add_path (Sign.base_name name) #> |
|
391 |
PureThy.add_thmss [(("eqvt", ths), [NominalThmDecls.eqvt_add])] #> snd #> |
|
392 |
Theory.parent_path) (names ~~ transp thss) thy |
|
393 |
end; |
|
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
394 |
|
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
395 |
|
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
396 |
(* outer syntax *) |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
397 |
|
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
398 |
local structure P = OuterParse and K = OuterKeyword in |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
399 |
|
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
400 |
val nominal_inductiveP = |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
401 |
OuterSyntax.command "nominal_inductive" |
22530 | 402 |
"prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal |
403 |
(P.name -- Scan.optional (P.$$$ "avoids" |-- P.and_list1 (P.name -- |
|
404 |
(P.$$$ ":" |-- Scan.repeat1 P.name))) [] >> (fn (name, avoids) => |
|
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
405 |
Toplevel.print o Toplevel.theory_to_proof (prove_strong_ind name avoids))); |
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
406 |
|
22530 | 407 |
val equivarianceP = |
408 |
OuterSyntax.command "equivariance" |
|
409 |
"prove equivariance for inductive predicate involving nominal datatypes" K.thy_decl |
|
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
410 |
(P.name -- Scan.optional (P.$$$ "[" |-- P.list1 P.name --| P.$$$ "]") [] >> |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22544
diff
changeset
|
411 |
(fn (name, atoms) => Toplevel.theory (prove_eqvt name atoms))); |
22530 | 412 |
|
413 |
val _ = OuterSyntax.add_keywords ["avoids"]; |
|
414 |
val _ = OuterSyntax.add_parsers [nominal_inductiveP, equivarianceP]; |
|
22313
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
415 |
|
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
416 |
end; |
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
417 |
|
1a507b463f50
First steps towards strengthening of induction rules for
berghofe
parents:
diff
changeset
|
418 |
end |