author | urbanc |
Fri, 25 Nov 2005 14:51:39 +0100 | |
changeset 18246 | 676d2e625d98 |
parent 18245 | 65e60434b3c2 |
child 18261 | 1318955d57ac |
permissions | -rw-r--r-- |
17870 | 1 |
(* $Id$ *) |
2 |
||
3 |
signature NOMINAL_PACKAGE = |
|
4 |
sig |
|
5 |
val add_nominal_datatype : bool -> string list -> (string list * bstring * mixfix * |
|
18068 | 6 |
(bstring * string list * mixfix) list) list -> theory -> theory |
17870 | 7 |
end |
8 |
||
18068 | 9 |
structure NominalPackage : NOMINAL_PACKAGE = |
17870 | 10 |
struct |
11 |
||
12 |
open DatatypeAux; |
|
18068 | 13 |
open NominalAtoms; |
17870 | 14 |
|
18016 | 15 |
(** FIXME: DatatypePackage should export this function **) |
16 |
||
17 |
local |
|
18 |
||
19 |
fun dt_recs (DtTFree _) = [] |
|
20 |
| dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts) |
|
21 |
| dt_recs (DtRec i) = [i]; |
|
22 |
||
23 |
fun dt_cases (descr: descr) (_, args, constrs) = |
|
24 |
let |
|
25 |
fun the_bname i = Sign.base_name (#1 (valOf (AList.lookup (op =) descr i))); |
|
26 |
val bnames = map the_bname (distinct (List.concat (map dt_recs args))); |
|
27 |
in map (fn (c, _) => space_implode "_" (Sign.base_name c :: bnames)) constrs end; |
|
28 |
||
29 |
||
30 |
fun induct_cases descr = |
|
31 |
DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr))); |
|
32 |
||
33 |
fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i)); |
|
34 |
||
35 |
in |
|
36 |
||
37 |
fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr); |
|
38 |
||
39 |
fun mk_case_names_exhausts descr new = |
|
40 |
map (RuleCases.case_names o exhaust_cases descr o #1) |
|
41 |
(List.filter (fn ((_, (name, _, _))) => name mem_string new) descr); |
|
42 |
||
43 |
end; |
|
44 |
||
45 |
(*******************************) |
|
46 |
||
17870 | 47 |
val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma); |
48 |
||
49 |
fun read_typ sign ((Ts, sorts), str) = |
|
50 |
let |
|
51 |
val T = Type.no_tvars (Sign.read_typ (sign, (AList.lookup op =) |
|
52 |
(map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg |
|
53 |
in (Ts @ [T], add_typ_tfrees (T, sorts)) end; |
|
54 |
||
55 |
(** taken from HOL/Tools/datatype_aux.ML **) |
|
56 |
||
57 |
fun indtac indrule indnames i st = |
|
58 |
let |
|
59 |
val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule)); |
|
60 |
val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop |
|
61 |
(Logic.strip_imp_concl (List.nth (prems_of st, i - 1)))); |
|
62 |
val getP = if can HOLogic.dest_imp (hd ts) then |
|
63 |
(apfst SOME) o HOLogic.dest_imp else pair NONE; |
|
64 |
fun abstr (t1, t2) = (case t1 of |
|
65 |
NONE => (case filter (fn Free (s, _) => s mem indnames | _ => false) |
|
66 |
(term_frees t2) of |
|
67 |
[Free (s, T)] => absfree (s, T, t2) |
|
68 |
| _ => sys_error "indtac") |
|
69 |
| SOME (_ $ t' $ _) => Abs ("x", fastype_of t', abstract_over (t', t2))) |
|
70 |
val cert = cterm_of (Thm.sign_of_thm st); |
|
71 |
val Ps = map (cert o head_of o snd o getP) ts; |
|
72 |
val indrule' = cterm_instantiate (Ps ~~ |
|
73 |
(map (cert o abstr o getP) ts')) indrule |
|
74 |
in |
|
75 |
rtac indrule' i st |
|
76 |
end; |
|
77 |
||
78 |
fun gen_add_nominal_datatype prep_typ err flat_names new_type_names dts thy = |
|
79 |
let |
|
80 |
(* this theory is used just for parsing *) |
|
81 |
||
82 |
val tmp_thy = thy |> |
|
83 |
Theory.copy |> |
|
84 |
Theory.add_types (map (fn (tvs, tname, mx, _) => |
|
85 |
(tname, length tvs, mx)) dts); |
|
86 |
||
87 |
val sign = Theory.sign_of tmp_thy; |
|
88 |
||
89 |
val atoms = atoms_of thy; |
|
90 |
val classes = map (NameSpace.map_base (fn s => "pt_" ^ s)) atoms; |
|
91 |
val cp_classes = List.concat (map (fn atom1 => map (fn atom2 => |
|
92 |
Sign.intern_class thy ("cp_" ^ Sign.base_name atom1 ^ "_" ^ |
|
93 |
Sign.base_name atom2)) atoms) atoms); |
|
94 |
fun augment_sort S = S union classes; |
|
95 |
val augment_sort_typ = map_type_tfree (fn (s, S) => TFree (s, augment_sort S)); |
|
96 |
||
97 |
fun prep_constr ((constrs, sorts), (cname, cargs, mx)) = |
|
98 |
let val (cargs', sorts') = Library.foldl (prep_typ sign) (([], sorts), cargs) |
|
99 |
in (constrs @ [(cname, cargs', mx)], sorts') end |
|
100 |
||
101 |
fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) = |
|
102 |
let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs) |
|
103 |
in (dts @ [(tvs, tname, mx, constrs')], sorts') end |
|
104 |
||
105 |
val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts); |
|
106 |
val sorts' = map (apsnd augment_sort) sorts; |
|
107 |
val tyvars = map #1 dts'; |
|
108 |
||
109 |
val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts'; |
|
110 |
val constr_syntax = map (fn (tvs, tname, mx, constrs) => |
|
111 |
map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts'; |
|
112 |
||
113 |
val ps = map (fn (_, n, _, _) => |
|
114 |
(Sign.full_name sign n, Sign.full_name sign (n ^ "_Rep"))) dts; |
|
115 |
val rps = map Library.swap ps; |
|
116 |
||
117 |
fun replace_types (Type ("nominal.ABS", [T, U])) = |
|
118 |
Type ("fun", [T, Type ("nominal.nOption", [replace_types U])]) |
|
119 |
| replace_types (Type (s, Ts)) = |
|
120 |
Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts) |
|
121 |
| replace_types T = T; |
|
122 |
||
123 |
fun replace_types' (Type (s, Ts)) = |
|
124 |
Type (getOpt (AList.lookup op = rps s, s), map replace_types' Ts) |
|
125 |
| replace_types' T = T; |
|
126 |
||
127 |
val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, tname ^ "_Rep", NoSyn, |
|
128 |
map (fn (cname, cargs, mx) => (cname, |
|
129 |
map (augment_sort_typ o replace_types) cargs, NoSyn)) constrs)) dts'; |
|
130 |
||
131 |
val new_type_names' = map (fn n => n ^ "_Rep") new_type_names; |
|
132 |
val full_new_type_names' = map (Sign.full_name (sign_of thy)) new_type_names'; |
|
133 |
||
18045 | 134 |
val ({induction, ...},thy1) = |
17870 | 135 |
DatatypePackage.add_datatype_i err flat_names new_type_names' dts'' thy; |
136 |
||
137 |
val SOME {descr, ...} = Symtab.lookup |
|
138 |
(DatatypePackage.get_datatypes thy1) (hd full_new_type_names'); |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
139 |
fun nth_dtyp i = typ_of_dtyp descr sorts' (DtRec i); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
140 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
141 |
val dt_atomTs = filter (fn Type (s, []) => s mem atoms | _ => false) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
142 |
(get_nonrec_types descr sorts); |
17870 | 143 |
|
144 |
(**** define permutation functions ****) |
|
145 |
||
146 |
val permT = mk_permT (TFree ("'x", HOLogic.typeS)); |
|
147 |
val pi = Free ("pi", permT); |
|
148 |
val perm_types = map (fn (i, _) => |
|
149 |
let val T = nth_dtyp i |
|
150 |
in permT --> T --> T end) descr; |
|
151 |
val perm_names = replicate (length new_type_names) "nominal.perm" @ |
|
152 |
DatatypeProp.indexify_names (map (fn i => Sign.full_name (sign_of thy1) |
|
153 |
("perm_" ^ name_of_typ (nth_dtyp i))) |
|
154 |
(length new_type_names upto length descr - 1)); |
|
155 |
val perm_names_types = perm_names ~~ perm_types; |
|
156 |
||
157 |
val perm_eqs = List.concat (map (fn (i, (_, _, constrs)) => |
|
158 |
let val T = nth_dtyp i |
|
159 |
in map (fn (cname, dts) => |
|
160 |
let |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
161 |
val Ts = map (typ_of_dtyp descr sorts') dts; |
17870 | 162 |
val names = DatatypeProp.make_tnames Ts; |
163 |
val args = map Free (names ~~ Ts); |
|
164 |
val c = Const (cname, Ts ---> T); |
|
165 |
fun perm_arg (dt, x) = |
|
166 |
let val T = type_of x |
|
167 |
in if is_rec_type dt then |
|
168 |
let val (Us, _) = strip_type T |
|
169 |
in list_abs (map (pair "x") Us, |
|
170 |
Const (List.nth (perm_names_types, body_index dt)) $ pi $ |
|
171 |
list_comb (x, map (fn (i, U) => |
|
172 |
Const ("nominal.perm", permT --> U --> U) $ |
|
173 |
(Const ("List.rev", permT --> permT) $ pi) $ |
|
174 |
Bound i) ((length Us - 1 downto 0) ~~ Us))) |
|
175 |
end |
|
176 |
else Const ("nominal.perm", permT --> T --> T) $ pi $ x |
|
177 |
end; |
|
178 |
in |
|
179 |
(("", HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
180 |
(Const (List.nth (perm_names_types, i)) $ |
|
181 |
Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $ |
|
182 |
list_comb (c, args), |
|
183 |
list_comb (c, map perm_arg (dts ~~ args))))), []) |
|
184 |
end) constrs |
|
185 |
end) descr); |
|
186 |
||
187 |
val (thy2, perm_simps) = thy1 |> |
|
188 |
Theory.add_consts_i (map (fn (s, T) => (Sign.base_name s, T, NoSyn)) |
|
189 |
(List.drop (perm_names_types, length new_type_names))) |> |
|
190 |
PrimrecPackage.add_primrec_i "" perm_eqs; |
|
191 |
||
192 |
(**** prove that permutation functions introduced by unfolding are ****) |
|
193 |
(**** equivalent to already existing permutation functions ****) |
|
194 |
||
195 |
val _ = warning ("length descr: " ^ string_of_int (length descr)); |
|
196 |
val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names)); |
|
197 |
||
198 |
val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types); |
|
199 |
val perm_fun_def = PureThy.get_thm thy2 (Name "perm_fun_def"); |
|
200 |
||
201 |
val unfolded_perm_eq_thms = |
|
202 |
if length descr = length new_type_names then [] |
|
203 |
else map standard (List.drop (split_conj_thm |
|
18010 | 204 |
(Goal.prove thy2 [] [] |
17870 | 205 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
206 |
(map (fn (c as (s, T), x) => |
|
207 |
let val [T1, T2] = binder_types T |
|
208 |
in HOLogic.mk_eq (Const c $ pi $ Free (x, T2), |
|
209 |
Const ("nominal.perm", T) $ pi $ Free (x, T2)) |
|
210 |
end) |
|
18010 | 211 |
(perm_names_types ~~ perm_indnames)))) |
212 |
(fn _ => EVERY [indtac induction perm_indnames 1, |
|
17870 | 213 |
ALLGOALS (asm_full_simp_tac |
214 |
(simpset_of thy2 addsimps [perm_fun_def]))])), |
|
215 |
length new_type_names)); |
|
216 |
||
217 |
(**** prove [] \<bullet> t = t ****) |
|
218 |
||
219 |
val _ = warning "perm_empty_thms"; |
|
220 |
||
221 |
val perm_empty_thms = List.concat (map (fn a => |
|
222 |
let val permT = mk_permT (Type (a, [])) |
|
223 |
in map standard (List.take (split_conj_thm |
|
18010 | 224 |
(Goal.prove thy2 [] [] |
17870 | 225 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
226 |
(map (fn ((s, T), x) => HOLogic.mk_eq |
|
227 |
(Const (s, permT --> T --> T) $ |
|
228 |
Const ("List.list.Nil", permT) $ Free (x, T), |
|
229 |
Free (x, T))) |
|
230 |
(perm_names ~~ |
|
18010 | 231 |
map body_type perm_types ~~ perm_indnames)))) |
232 |
(fn _ => EVERY [indtac induction perm_indnames 1, |
|
17870 | 233 |
ALLGOALS (asm_full_simp_tac (simpset_of thy2))])), |
234 |
length new_type_names)) |
|
235 |
end) |
|
236 |
atoms); |
|
237 |
||
238 |
(**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****) |
|
239 |
||
240 |
val _ = warning "perm_append_thms"; |
|
241 |
||
242 |
(*FIXME: these should be looked up statically*) |
|
243 |
val at_pt_inst = PureThy.get_thm thy2 (Name "at_pt_inst"); |
|
244 |
val pt2 = PureThy.get_thm thy2 (Name "pt2"); |
|
245 |
||
246 |
val perm_append_thms = List.concat (map (fn a => |
|
247 |
let |
|
248 |
val permT = mk_permT (Type (a, [])); |
|
249 |
val pi1 = Free ("pi1", permT); |
|
250 |
val pi2 = Free ("pi2", permT); |
|
251 |
val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst")); |
|
252 |
val pt2' = pt_inst RS pt2; |
|
253 |
val pt2_ax = PureThy.get_thm thy2 |
|
254 |
(Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "2") a)); |
|
255 |
in List.take (map standard (split_conj_thm |
|
18010 | 256 |
(Goal.prove thy2 [] [] |
17870 | 257 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
258 |
(map (fn ((s, T), x) => |
|
259 |
let val perm = Const (s, permT --> T --> T) |
|
260 |
in HOLogic.mk_eq |
|
261 |
(perm $ (Const ("List.op @", permT --> permT --> permT) $ |
|
262 |
pi1 $ pi2) $ Free (x, T), |
|
263 |
perm $ pi1 $ (perm $ pi2 $ Free (x, T))) |
|
264 |
end) |
|
265 |
(perm_names ~~ |
|
18010 | 266 |
map body_type perm_types ~~ perm_indnames)))) |
267 |
(fn _ => EVERY [indtac induction perm_indnames 1, |
|
17870 | 268 |
ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))), |
269 |
length new_type_names) |
|
270 |
end) atoms); |
|
271 |
||
272 |
(**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****) |
|
273 |
||
274 |
val _ = warning "perm_eq_thms"; |
|
275 |
||
276 |
val pt3 = PureThy.get_thm thy2 (Name "pt3"); |
|
277 |
val pt3_rev = PureThy.get_thm thy2 (Name "pt3_rev"); |
|
278 |
||
279 |
val perm_eq_thms = List.concat (map (fn a => |
|
280 |
let |
|
281 |
val permT = mk_permT (Type (a, [])); |
|
282 |
val pi1 = Free ("pi1", permT); |
|
283 |
val pi2 = Free ("pi2", permT); |
|
284 |
(*FIXME: not robust - better access these theorems using NominalData?*) |
|
285 |
val at_inst = PureThy.get_thm thy2 (Name ("at_" ^ Sign.base_name a ^ "_inst")); |
|
286 |
val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst")); |
|
287 |
val pt3' = pt_inst RS pt3; |
|
288 |
val pt3_rev' = at_inst RS (pt_inst RS pt3_rev); |
|
289 |
val pt3_ax = PureThy.get_thm thy2 |
|
290 |
(Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "3") a)); |
|
291 |
in List.take (map standard (split_conj_thm |
|
18010 | 292 |
(Goal.prove thy2 [] [] (Logic.mk_implies |
17870 | 293 |
(HOLogic.mk_Trueprop (Const ("nominal.prm_eq", |
294 |
permT --> permT --> HOLogic.boolT) $ pi1 $ pi2), |
|
295 |
HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
|
296 |
(map (fn ((s, T), x) => |
|
297 |
let val perm = Const (s, permT --> T --> T) |
|
298 |
in HOLogic.mk_eq |
|
299 |
(perm $ pi1 $ Free (x, T), |
|
300 |
perm $ pi2 $ Free (x, T)) |
|
301 |
end) |
|
302 |
(perm_names ~~ |
|
18010 | 303 |
map body_type perm_types ~~ perm_indnames))))) |
304 |
(fn _ => EVERY [indtac induction perm_indnames 1, |
|
17870 | 305 |
ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))), |
306 |
length new_type_names) |
|
307 |
end) atoms); |
|
308 |
||
309 |
(**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****) |
|
310 |
||
311 |
val cp1 = PureThy.get_thm thy2 (Name "cp1"); |
|
312 |
val dj_cp = PureThy.get_thm thy2 (Name "dj_cp"); |
|
313 |
val pt_perm_compose = PureThy.get_thm thy2 (Name "pt_perm_compose"); |
|
314 |
val pt_perm_compose_rev = PureThy.get_thm thy2 (Name "pt_perm_compose_rev"); |
|
315 |
val dj_perm_perm_forget = PureThy.get_thm thy2 (Name "dj_perm_perm_forget"); |
|
316 |
||
317 |
fun composition_instance name1 name2 thy = |
|
318 |
let |
|
319 |
val name1' = Sign.base_name name1; |
|
320 |
val name2' = Sign.base_name name2; |
|
321 |
val cp_class = Sign.intern_class thy ("cp_" ^ name1' ^ "_" ^ name2'); |
|
322 |
val permT1 = mk_permT (Type (name1, [])); |
|
323 |
val permT2 = mk_permT (Type (name2, [])); |
|
324 |
val augment = map_type_tfree |
|
325 |
(fn (x, S) => TFree (x, cp_class :: S)); |
|
326 |
val Ts = map (augment o body_type) perm_types; |
|
327 |
val cp_inst = PureThy.get_thm thy |
|
328 |
(Name ("cp_" ^ name1' ^ "_" ^ name2' ^ "_inst")); |
|
329 |
val simps = simpset_of thy addsimps (perm_fun_def :: |
|
330 |
(if name1 <> name2 then |
|
331 |
let val dj = PureThy.get_thm thy (Name ("dj_" ^ name2' ^ "_" ^ name1')) |
|
332 |
in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end |
|
333 |
else |
|
334 |
let |
|
335 |
val at_inst = PureThy.get_thm thy (Name ("at_" ^ name1' ^ "_inst")); |
|
336 |
val pt_inst = PureThy.get_thm thy (Name ("pt_" ^ name1' ^ "_inst")) |
|
337 |
in |
|
338 |
[cp_inst RS cp1 RS sym, |
|
339 |
at_inst RS (pt_inst RS pt_perm_compose) RS sym, |
|
340 |
at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym] |
|
341 |
end)) |
|
18010 | 342 |
val thms = split_conj_thm (standard (Goal.prove thy [] [] |
17870 | 343 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
344 |
(map (fn ((s, T), x) => |
|
345 |
let |
|
346 |
val pi1 = Free ("pi1", permT1); |
|
347 |
val pi2 = Free ("pi2", permT2); |
|
348 |
val perm1 = Const (s, permT1 --> T --> T); |
|
349 |
val perm2 = Const (s, permT2 --> T --> T); |
|
350 |
val perm3 = Const ("nominal.perm", permT1 --> permT2 --> permT2) |
|
351 |
in HOLogic.mk_eq |
|
352 |
(perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)), |
|
353 |
perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T))) |
|
354 |
end) |
|
18010 | 355 |
(perm_names ~~ Ts ~~ perm_indnames)))) |
356 |
(fn _ => EVERY [indtac induction perm_indnames 1, |
|
357 |
ALLGOALS (asm_full_simp_tac simps)]))) |
|
17870 | 358 |
in |
359 |
foldl (fn ((s, tvs), thy) => AxClass.add_inst_arity_i |
|
360 |
(s, replicate (length tvs) (cp_class :: classes), [cp_class]) |
|
361 |
(AxClass.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy) |
|
362 |
thy (full_new_type_names' ~~ tyvars) |
|
363 |
end; |
|
364 |
||
365 |
val (thy3, perm_thmss) = thy2 |> |
|
366 |
fold (fn name1 => fold (composition_instance name1) atoms) atoms |> |
|
367 |
curry (Library.foldr (fn ((i, (tyname, args, _)), thy) => |
|
368 |
AxClass.add_inst_arity_i (tyname, replicate (length args) classes, classes) |
|
369 |
(AxClass.intro_classes_tac [] THEN REPEAT (EVERY |
|
370 |
[resolve_tac perm_empty_thms 1, |
|
371 |
resolve_tac perm_append_thms 1, |
|
372 |
resolve_tac perm_eq_thms 1, assume_tac 1])) thy)) |
|
373 |
(List.take (descr, length new_type_names)) |> |
|
374 |
PureThy.add_thmss |
|
375 |
[((space_implode "_" new_type_names ^ "_unfolded_perm_eq", |
|
376 |
unfolded_perm_eq_thms), [Simplifier.simp_add_global]), |
|
377 |
((space_implode "_" new_type_names ^ "_perm_empty", |
|
378 |
perm_empty_thms), [Simplifier.simp_add_global]), |
|
379 |
((space_implode "_" new_type_names ^ "_perm_append", |
|
380 |
perm_append_thms), [Simplifier.simp_add_global]), |
|
381 |
((space_implode "_" new_type_names ^ "_perm_eq", |
|
382 |
perm_eq_thms), [Simplifier.simp_add_global])]; |
|
383 |
||
384 |
(**** Define representing sets ****) |
|
385 |
||
386 |
val _ = warning "representing sets"; |
|
387 |
||
388 |
val rep_set_names = map (Sign.full_name thy3) (DatatypeProp.indexify_names |
|
389 |
(map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr)); |
|
390 |
val big_rep_name = |
|
391 |
space_implode "_" (DatatypeProp.indexify_names (List.mapPartial |
|
392 |
(fn (i, ("nominal.nOption", _, _)) => NONE |
|
393 |
| (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set"; |
|
394 |
val _ = warning ("big_rep_name: " ^ big_rep_name); |
|
395 |
||
396 |
fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) = |
|
397 |
(case AList.lookup op = descr i of |
|
398 |
SOME ("nominal.nOption", _, [(_, [dt']), _]) => |
|
399 |
apfst (cons dt) (strip_option dt') |
|
400 |
| _ => ([], dtf)) |
|
401 |
| strip_option dt = ([], dt); |
|
402 |
||
403 |
fun make_intr s T (cname, cargs) = |
|
404 |
let |
|
405 |
fun mk_prem (dt, (j, j', prems, ts)) = |
|
406 |
let |
|
407 |
val (dts, dt') = strip_option dt; |
|
408 |
val (dts', dt'') = strip_dtyp dt'; |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
409 |
val Ts = map (typ_of_dtyp descr sorts') dts; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
410 |
val Us = map (typ_of_dtyp descr sorts') dts'; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
411 |
val T = typ_of_dtyp descr sorts' dt''; |
17870 | 412 |
val free = mk_Free "x" (Us ---> T) j; |
413 |
val free' = app_bnds free (length Us); |
|
414 |
fun mk_abs_fun (T, (i, t)) = |
|
415 |
let val U = fastype_of t |
|
416 |
in (i + 1, Const ("nominal.abs_fun", [T, U, T] ---> |
|
417 |
Type ("nominal.nOption", [U])) $ mk_Free "y" T i $ t) |
|
418 |
end |
|
419 |
in (j + 1, j' + length Ts, |
|
420 |
case dt'' of |
|
421 |
DtRec k => list_all (map (pair "x") Us, |
|
422 |
HOLogic.mk_Trueprop (HOLogic.mk_mem (free', |
|
423 |
Const (List.nth (rep_set_names, k), |
|
424 |
HOLogic.mk_setT T)))) :: prems |
|
425 |
| _ => prems, |
|
426 |
snd (foldr mk_abs_fun (j', free) Ts) :: ts) |
|
427 |
end; |
|
428 |
||
429 |
val (_, _, prems, ts) = foldr mk_prem (1, 1, [], []) cargs; |
|
430 |
val concl = HOLogic.mk_Trueprop (HOLogic.mk_mem |
|
431 |
(list_comb (Const (cname, map fastype_of ts ---> T), ts), |
|
432 |
Const (s, HOLogic.mk_setT T))) |
|
433 |
in Logic.list_implies (prems, concl) |
|
434 |
end; |
|
435 |
||
436 |
val (intr_ts, ind_consts) = |
|
437 |
apfst List.concat (ListPair.unzip (List.mapPartial |
|
438 |
(fn ((_, ("nominal.nOption", _, _)), _) => NONE |
|
439 |
| ((i, (_, _, constrs)), rep_set_name) => |
|
440 |
let val T = nth_dtyp i |
|
441 |
in SOME (map (make_intr rep_set_name T) constrs, |
|
442 |
Const (rep_set_name, HOLogic.mk_setT T)) |
|
443 |
end) |
|
444 |
(descr ~~ rep_set_names))); |
|
445 |
||
446 |
val (thy4, {raw_induct = rep_induct, intrs = rep_intrs, ...}) = |
|
447 |
setmp InductivePackage.quiet_mode false |
|
448 |
(InductivePackage.add_inductive_i false true big_rep_name false true false |
|
449 |
ind_consts (map (fn x => (("", x), [])) intr_ts) []) thy3; |
|
450 |
||
451 |
(**** Prove that representing set is closed under permutation ****) |
|
452 |
||
453 |
val _ = warning "proving closure under permutation..."; |
|
454 |
||
455 |
val perm_indnames' = List.mapPartial |
|
456 |
(fn (x, (_, ("nominal.nOption", _, _))) => NONE | (x, _) => SOME x) |
|
457 |
(perm_indnames ~~ descr); |
|
458 |
||
459 |
fun mk_perm_closed name = map (fn th => standard (th RS mp)) |
|
18010 | 460 |
(List.take (split_conj_thm (Goal.prove thy4 [] [] |
17870 | 461 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map |
462 |
(fn (S, x) => |
|
463 |
let |
|
464 |
val S = map_term_types (map_type_tfree |
|
465 |
(fn (s, cs) => TFree (s, cs union cp_classes))) S; |
|
466 |
val T = HOLogic.dest_setT (fastype_of S); |
|
467 |
val permT = mk_permT (Type (name, [])) |
|
468 |
in HOLogic.mk_imp (HOLogic.mk_mem (Free (x, T), S), |
|
469 |
HOLogic.mk_mem (Const ("nominal.perm", permT --> T --> T) $ |
|
470 |
Free ("pi", permT) $ Free (x, T), S)) |
|
18010 | 471 |
end) (ind_consts ~~ perm_indnames')))) |
472 |
(fn _ => EVERY (* CU: added perm_fun_def in the final tactic in order to deal with funs *) |
|
17870 | 473 |
[indtac rep_induct [] 1, |
474 |
ALLGOALS (simp_tac (simpset_of thy4 addsimps |
|
475 |
(symmetric perm_fun_def :: PureThy.get_thms thy4 (Name ("abs_perm"))))), |
|
476 |
ALLGOALS (resolve_tac rep_intrs |
|
477 |
THEN_ALL_NEW (asm_full_simp_tac (simpset_of thy4 addsimps [perm_fun_def])))])), |
|
478 |
length new_type_names)); |
|
479 |
||
480 |
(* FIXME: theorems are stored in database for testing only *) |
|
481 |
val perm_closed_thmss = map mk_perm_closed atoms; |
|
482 |
val (thy5, _) = PureThy.add_thmss [(("perm_closed", |
|
483 |
List.concat perm_closed_thmss), [])] thy4; |
|
484 |
||
485 |
(**** typedef ****) |
|
486 |
||
487 |
val _ = warning "defining type..."; |
|
488 |
||
489 |
val (thy6, typedefs) = |
|
490 |
foldl_map (fn (thy, ((((name, mx), tvs), c), name')) => |
|
491 |
setmp TypedefPackage.quiet_mode true |
|
492 |
(TypedefPackage.add_typedef_i false (SOME name') (name, tvs, mx) c NONE |
|
493 |
(rtac exI 1 THEN |
|
494 |
QUIET_BREADTH_FIRST (has_fewer_prems 1) |
|
495 |
(resolve_tac rep_intrs 1))) thy |> (fn (thy, r) => |
|
496 |
let |
|
497 |
val permT = mk_permT (TFree (variant tvs "'a", HOLogic.typeS)); |
|
498 |
val pi = Free ("pi", permT); |
|
499 |
val tvs' = map (fn s => TFree (s, the (AList.lookup op = sorts' s))) tvs; |
|
500 |
val T = Type (Sign.intern_type thy name, tvs'); |
|
501 |
val Const (_, Type (_, [U])) = c |
|
502 |
in apsnd (pair r o hd) |
|
503 |
(PureThy.add_defs_i true [(("prm_" ^ name ^ "_def", Logic.mk_equals |
|
504 |
(Const ("nominal.perm", permT --> T --> T) $ pi $ Free ("x", T), |
|
505 |
Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $ |
|
506 |
(Const ("nominal.perm", permT --> U --> U) $ pi $ |
|
507 |
(Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $ |
|
508 |
Free ("x", T))))), [])] thy) |
|
509 |
end)) |
|
510 |
(thy5, types_syntax ~~ tyvars ~~ |
|
511 |
(List.take (ind_consts, length new_type_names)) ~~ new_type_names); |
|
512 |
||
513 |
val perm_defs = map snd typedefs; |
|
514 |
val Abs_inverse_thms = map (#Abs_inverse o fst) typedefs; |
|
18016 | 515 |
val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs; |
17870 | 516 |
val Rep_thms = map (#Rep o fst) typedefs; |
517 |
||
18016 | 518 |
val big_name = space_implode "_" new_type_names; |
519 |
||
520 |
||
17870 | 521 |
(** prove that new types are in class pt_<name> **) |
522 |
||
523 |
val _ = warning "prove that new types are in class pt_<name> ..."; |
|
524 |
||
525 |
fun pt_instance ((class, atom), perm_closed_thms) = |
|
526 |
fold (fn (((({Abs_inverse, Rep_inverse, Rep, ...}, |
|
527 |
perm_def), name), tvs), perm_closed) => fn thy => |
|
528 |
AxClass.add_inst_arity_i |
|
529 |
(Sign.intern_type thy name, |
|
530 |
replicate (length tvs) (classes @ cp_classes), [class]) |
|
531 |
(EVERY [AxClass.intro_classes_tac [], |
|
532 |
rewrite_goals_tac [perm_def], |
|
533 |
asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1, |
|
534 |
asm_full_simp_tac (simpset_of thy addsimps |
|
535 |
[Rep RS perm_closed RS Abs_inverse]) 1, |
|
536 |
asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy |
|
537 |
(Name ("pt_" ^ Sign.base_name atom ^ "3"))]) 1]) thy) |
|
538 |
(typedefs ~~ new_type_names ~~ tyvars ~~ perm_closed_thms); |
|
539 |
||
540 |
||
541 |
(** prove that new types are in class cp_<name1>_<name2> **) |
|
542 |
||
543 |
val _ = warning "prove that new types are in class cp_<name1>_<name2> ..."; |
|
544 |
||
545 |
fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy = |
|
546 |
let |
|
547 |
val name = "cp_" ^ Sign.base_name atom1 ^ "_" ^ Sign.base_name atom2; |
|
548 |
val class = Sign.intern_class thy name; |
|
549 |
val cp1' = PureThy.get_thm thy (Name (name ^ "_inst")) RS cp1 |
|
550 |
in fold (fn ((((({Abs_inverse, Rep_inverse, Rep, ...}, |
|
551 |
perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy => |
|
552 |
AxClass.add_inst_arity_i |
|
553 |
(Sign.intern_type thy name, |
|
554 |
replicate (length tvs) (classes @ cp_classes), [class]) |
|
555 |
(EVERY [AxClass.intro_classes_tac [], |
|
556 |
rewrite_goals_tac [perm_def], |
|
557 |
asm_full_simp_tac (simpset_of thy addsimps |
|
558 |
((Rep RS perm_closed1 RS Abs_inverse) :: |
|
559 |
(if atom1 = atom2 then [] |
|
560 |
else [Rep RS perm_closed2 RS Abs_inverse]))) 1, |
|
18016 | 561 |
cong_tac 1, |
17870 | 562 |
rtac refl 1, |
563 |
rtac cp1' 1]) thy) |
|
564 |
(typedefs ~~ new_type_names ~~ tyvars ~~ perm_closed_thms1 ~~ |
|
565 |
perm_closed_thms2) thy |
|
566 |
end; |
|
567 |
||
568 |
val thy7 = fold (fn x => fn thy => thy |> |
|
569 |
pt_instance x |> |
|
570 |
fold (cp_instance (apfst snd x)) (atoms ~~ perm_closed_thmss)) |
|
571 |
(classes ~~ atoms ~~ perm_closed_thmss) thy6; |
|
572 |
||
573 |
(**** constructors ****) |
|
574 |
||
575 |
fun mk_abs_fun (x, t) = |
|
576 |
let |
|
577 |
val T = fastype_of x; |
|
578 |
val U = fastype_of t |
|
579 |
in |
|
580 |
Const ("nominal.abs_fun", T --> U --> T --> |
|
581 |
Type ("nominal.nOption", [U])) $ x $ t |
|
582 |
end; |
|
583 |
||
18016 | 584 |
val (ty_idxs, _) = foldl |
585 |
(fn ((i, ("nominal.nOption", _, _)), p) => p |
|
586 |
| ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr; |
|
587 |
||
588 |
fun reindex (DtType (s, dts)) = DtType (s, map reindex dts) |
|
589 |
| reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i)) |
|
590 |
| reindex dt = dt; |
|
591 |
||
592 |
fun strip_suffix i s = implode (List.take (explode s, size s - i)); |
|
593 |
||
594 |
(** strips the "_Rep" in type names *) |
|
18045 | 595 |
fun strip_nth_name i s = |
596 |
let val xs = NameSpace.unpack s; |
|
597 |
in NameSpace.pack (Library.nth_map (length xs - i) (strip_suffix 4) xs) end; |
|
18016 | 598 |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
599 |
val (descr'', ndescr) = ListPair.unzip (List.mapPartial |
18016 | 600 |
(fn (i, ("nominal.nOption", _, _)) => NONE |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
601 |
| (i, (s, dts, constrs)) => |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
602 |
let |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
603 |
val SOME index = AList.lookup op = ty_idxs i; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
604 |
val (constrs1, constrs2) = ListPair.unzip |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
605 |
(map (fn (cname, cargs) => apfst (pair (strip_nth_name 2 cname)) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
606 |
(foldl_map (fn (dts, dt) => |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
607 |
let val (dts', dt') = strip_option dt |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
608 |
in (dts @ dts' @ [reindex dt'], (length dts, length dts')) end) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
609 |
([], cargs))) constrs) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
610 |
in SOME ((index, (strip_nth_name 1 s, map reindex dts, constrs1)), |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
611 |
(index, constrs2)) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
612 |
end) descr); |
18045 | 613 |
|
18016 | 614 |
val (descr1, descr2) = splitAt (length new_type_names, descr''); |
615 |
val descr' = [descr1, descr2]; |
|
616 |
||
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
617 |
val typ_of_dtyp' = replace_types' o typ_of_dtyp descr sorts'; |
17870 | 618 |
|
619 |
val rep_names = map (fn s => |
|
620 |
Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names; |
|
621 |
val abs_names = map (fn s => |
|
622 |
Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names; |
|
623 |
||
18016 | 624 |
val recTs' = List.mapPartial |
625 |
(fn ((_, ("nominal.nOption", _, _)), T) => NONE |
|
626 |
| (_, T) => SOME T) (descr ~~ get_rec_types descr sorts'); |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
627 |
val recTs = get_rec_types descr'' sorts'; |
18016 | 628 |
val newTs' = Library.take (length new_type_names, recTs'); |
629 |
val newTs = Library.take (length new_type_names, recTs); |
|
17870 | 630 |
|
631 |
val full_new_type_names = map (Sign.full_name (sign_of thy)) new_type_names; |
|
632 |
||
633 |
fun make_constr_def tname T T' ((thy, defs, eqns), ((cname, cargs), (cname', mx))) = |
|
634 |
let |
|
635 |
fun constr_arg (dt, (j, l_args, r_args)) = |
|
636 |
let |
|
637 |
val x' = mk_Free "x" (typ_of_dtyp' dt) j; |
|
638 |
val (dts, dt') = strip_option dt; |
|
639 |
val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp' dt) i) |
|
640 |
(dts ~~ (j upto j + length dts - 1)) |
|
641 |
val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts) |
|
642 |
val (dts', dt'') = strip_dtyp dt' |
|
643 |
in case dt'' of |
|
644 |
DtRec k => if k < length new_type_names then |
|
645 |
(j + length dts + 1, |
|
646 |
xs @ x :: l_args, |
|
647 |
foldr mk_abs_fun |
|
648 |
(list_abs (map (pair "z" o typ_of_dtyp') dts', |
|
649 |
Const (List.nth (rep_names, k), typ_of_dtyp' dt'' --> |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
650 |
typ_of_dtyp descr sorts' dt'') $ app_bnds x (length dts'))) |
17870 | 651 |
xs :: r_args) |
652 |
else error "nested recursion not (yet) supported" |
|
653 |
| _ => (j + 1, x' :: l_args, x' :: r_args) |
|
654 |
end |
|
655 |
||
656 |
val (_, l_args, r_args) = foldr constr_arg (1, [], []) cargs; |
|
657 |
val abs_name = Sign.intern_const (Theory.sign_of thy) ("Abs_" ^ tname); |
|
658 |
val rep_name = Sign.intern_const (Theory.sign_of thy) ("Rep_" ^ tname); |
|
659 |
val constrT = map fastype_of l_args ---> T; |
|
660 |
val lhs = list_comb (Const (Sign.full_name thy (Sign.base_name cname), |
|
661 |
constrT), l_args); |
|
662 |
val rhs = list_comb (Const (cname, map fastype_of r_args ---> T'), r_args); |
|
663 |
val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs); |
|
664 |
val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
665 |
(Const (rep_name, T --> T') $ lhs, rhs)); |
|
666 |
val def_name = (Sign.base_name cname) ^ "_def"; |
|
667 |
val (thy', [def_thm]) = thy |> |
|
668 |
Theory.add_consts_i [(cname', constrT, mx)] |> |
|
669 |
(PureThy.add_defs_i false o map Thm.no_attributes) [(def_name, def)] |
|
670 |
in (thy', defs @ [def_thm], eqns @ [eqn]) end; |
|
671 |
||
672 |
fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), |
|
673 |
(((((_, (_, _, constrs)), tname), T), T'), constr_syntax)) = |
|
674 |
let |
|
675 |
val rep_const = cterm_of thy |
|
676 |
(Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T')); |
|
677 |
val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma); |
|
678 |
val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T') |
|
679 |
((Theory.add_path tname thy, defs, []), constrs ~~ constr_syntax) |
|
680 |
in |
|
681 |
(parent_path flat_names thy', defs', eqns @ [eqns'], dist_lemmas @ [dist]) |
|
682 |
end; |
|
683 |
||
684 |
val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs |
|
685 |
((thy7, [], [], []), List.take (descr, length new_type_names) ~~ |
|
686 |
new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax); |
|
687 |
||
688 |
val abs_inject_thms = map (fn tname => |
|
689 |
PureThy.get_thm thy8 (Name ("Abs_" ^ tname ^ "_inject"))) new_type_names; |
|
690 |
||
691 |
val rep_inject_thms = map (fn tname => |
|
692 |
PureThy.get_thm thy8 (Name ("Rep_" ^ tname ^ "_inject"))) new_type_names; |
|
693 |
||
694 |
val rep_thms = map (fn tname => |
|
695 |
PureThy.get_thm thy8 (Name ("Rep_" ^ tname))) new_type_names; |
|
696 |
||
697 |
val rep_inverse_thms = map (fn tname => |
|
698 |
PureThy.get_thm thy8 (Name ("Rep_" ^ tname ^ "_inverse"))) new_type_names; |
|
699 |
||
700 |
(* prove theorem Rep_i (Constr_j ...) = Constr'_j ... *) |
|
701 |
||
702 |
fun prove_constr_rep_thm eqn = |
|
703 |
let |
|
704 |
val inj_thms = map (fn r => r RS iffD1) abs_inject_thms; |
|
705 |
val rewrites = constr_defs @ map mk_meta_eq rep_inverse_thms |
|
18010 | 706 |
in standard (Goal.prove thy8 [] [] eqn (fn _ => EVERY |
17870 | 707 |
[resolve_tac inj_thms 1, |
708 |
rewrite_goals_tac rewrites, |
|
709 |
rtac refl 3, |
|
710 |
resolve_tac rep_intrs 2, |
|
18010 | 711 |
REPEAT (resolve_tac rep_thms 1)])) |
17870 | 712 |
end; |
713 |
||
714 |
val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns; |
|
715 |
||
716 |
(* prove theorem pi \<bullet> Rep_i x = Rep_i (pi \<bullet> x) *) |
|
717 |
||
718 |
fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th => |
|
719 |
let |
|
720 |
val _ $ (_ $ (Rep $ x) $ _) = Logic.unvarify (prop_of th); |
|
721 |
val Type ("fun", [T, U]) = fastype_of Rep; |
|
722 |
val permT = mk_permT (Type (atom, [])); |
|
723 |
val pi = Free ("pi", permT); |
|
724 |
in |
|
18010 | 725 |
standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq |
17870 | 726 |
(Const ("nominal.perm", permT --> U --> U) $ pi $ (Rep $ x), |
18010 | 727 |
Rep $ (Const ("nominal.perm", permT --> T --> T) $ pi $ x)))) |
728 |
(fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @ |
|
729 |
perm_closed_thms @ Rep_thms)) 1)) |
|
17870 | 730 |
end) Rep_thms; |
731 |
||
732 |
val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm |
|
733 |
(atoms ~~ perm_closed_thmss)); |
|
734 |
||
735 |
(* prove distinctness theorems *) |
|
736 |
||
18016 | 737 |
val distinct_props = setmp DatatypeProp.dtK 1000 |
738 |
(DatatypeProp.make_distincts new_type_names descr' sorts') thy8; |
|
17870 | 739 |
|
740 |
val dist_rewrites = map (fn (rep_thms, dist_lemma) => |
|
741 |
dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])) |
|
742 |
(constr_rep_thmss ~~ dist_lemmas); |
|
743 |
||
744 |
fun prove_distinct_thms (_, []) = [] |
|
745 |
| prove_distinct_thms (p as (rep_thms, dist_lemma), t::ts) = |
|
746 |
let |
|
18010 | 747 |
val dist_thm = standard (Goal.prove thy8 [] [] t (fn _ => |
748 |
simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1)) |
|
17870 | 749 |
in dist_thm::(standard (dist_thm RS not_sym)):: |
750 |
(prove_distinct_thms (p, ts)) |
|
751 |
end; |
|
752 |
||
753 |
val distinct_thms = map prove_distinct_thms |
|
754 |
(constr_rep_thmss ~~ dist_lemmas ~~ distinct_props); |
|
755 |
||
756 |
(** prove equations for permutation functions **) |
|
757 |
||
758 |
val abs_perm = PureThy.get_thms thy8 (Name "abs_perm"); (* FIXME *) |
|
759 |
||
760 |
val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => |
|
761 |
let val T = replace_types' (nth_dtyp i) |
|
762 |
in List.concat (map (fn (atom, perm_closed_thms) => |
|
763 |
map (fn ((cname, dts), constr_rep_thm) => |
|
764 |
let |
|
765 |
val cname = Sign.intern_const thy8 |
|
766 |
(NameSpace.append tname (Sign.base_name cname)); |
|
767 |
val permT = mk_permT (Type (atom, [])); |
|
768 |
val pi = Free ("pi", permT); |
|
769 |
||
770 |
fun perm t = |
|
771 |
let val T = fastype_of t |
|
772 |
in Const ("nominal.perm", permT --> T --> T) $ pi $ t end; |
|
773 |
||
774 |
fun constr_arg (dt, (j, l_args, r_args)) = |
|
775 |
let |
|
776 |
val x' = mk_Free "x" (typ_of_dtyp' dt) j; |
|
777 |
val (dts, dt') = strip_option dt; |
|
778 |
val Ts = map typ_of_dtyp' dts; |
|
779 |
val xs = map (fn (T, i) => mk_Free "x" T i) |
|
780 |
(Ts ~~ (j upto j + length dts - 1)) |
|
781 |
val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts); |
|
782 |
val (dts', dt'') = strip_dtyp dt'; |
|
783 |
in case dt'' of |
|
784 |
DtRec k => if k < length new_type_names then |
|
785 |
(j + length dts + 1, |
|
786 |
xs @ x :: l_args, |
|
787 |
map perm (xs @ [x]) @ r_args) |
|
788 |
else error "nested recursion not (yet) supported" |
|
789 |
| _ => (j + 1, x' :: l_args, perm x' :: r_args) |
|
790 |
end |
|
791 |
||
792 |
val (_, l_args, r_args) = foldr constr_arg (1, [], []) dts; |
|
793 |
val c = Const (cname, map fastype_of l_args ---> T) |
|
794 |
in |
|
18010 | 795 |
standard (Goal.prove thy8 [] [] |
17870 | 796 |
(HOLogic.mk_Trueprop (HOLogic.mk_eq |
18010 | 797 |
(perm (list_comb (c, l_args)), list_comb (c, r_args)))) |
798 |
(fn _ => EVERY |
|
17870 | 799 |
[simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1, |
800 |
simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @ |
|
801 |
constr_defs @ perm_closed_thms)) 1, |
|
802 |
TRY (simp_tac (HOL_basic_ss addsimps |
|
803 |
(symmetric perm_fun_def :: abs_perm)) 1), |
|
804 |
TRY (simp_tac (HOL_basic_ss addsimps |
|
805 |
(perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @ |
|
18010 | 806 |
perm_closed_thms)) 1)])) |
17870 | 807 |
end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss)) |
808 |
end) (List.take (descr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); |
|
809 |
||
810 |
(** prove injectivity of constructors **) |
|
811 |
||
812 |
val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms; |
|
813 |
val alpha = PureThy.get_thms thy8 (Name "alpha"); |
|
814 |
val abs_fresh = PureThy.get_thms thy8 (Name "abs_fresh"); |
|
815 |
val fresh_def = PureThy.get_thm thy8 (Name "fresh_def"); |
|
816 |
val supp_def = PureThy.get_thm thy8 (Name "supp_def"); |
|
817 |
||
818 |
val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => |
|
819 |
let val T = replace_types' (nth_dtyp i) |
|
820 |
in List.mapPartial (fn ((cname, dts), constr_rep_thm) => |
|
821 |
if null dts then NONE else SOME |
|
822 |
let |
|
823 |
val cname = Sign.intern_const thy8 |
|
824 |
(NameSpace.append tname (Sign.base_name cname)); |
|
825 |
||
826 |
fun make_inj (dt, (j, args1, args2, eqs)) = |
|
827 |
let |
|
828 |
val x' = mk_Free "x" (typ_of_dtyp' dt) j; |
|
829 |
val y' = mk_Free "y" (typ_of_dtyp' dt) j; |
|
830 |
val (dts, dt') = strip_option dt; |
|
831 |
val Ts_idx = map typ_of_dtyp' dts ~~ (j upto j + length dts - 1); |
|
832 |
val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx; |
|
833 |
val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx; |
|
834 |
val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts); |
|
835 |
val y = mk_Free "y" (typ_of_dtyp' dt') (j + length dts); |
|
836 |
val (dts', dt'') = strip_dtyp dt'; |
|
837 |
in case dt'' of |
|
838 |
DtRec k => if k < length new_type_names then |
|
839 |
(j + length dts + 1, |
|
840 |
xs @ (x :: args1), ys @ (y :: args2), |
|
841 |
HOLogic.mk_eq |
|
842 |
(foldr mk_abs_fun x xs, foldr mk_abs_fun y ys) :: eqs) |
|
843 |
else error "nested recursion not (yet) supported" |
|
844 |
| _ => (j + 1, x' :: args1, y' :: args2, HOLogic.mk_eq (x', y') :: eqs) |
|
845 |
end; |
|
846 |
||
847 |
val (_, args1, args2, eqs) = foldr make_inj (1, [], [], []) dts; |
|
848 |
val Ts = map fastype_of args1; |
|
849 |
val c = Const (cname, Ts ---> T) |
|
850 |
in |
|
18010 | 851 |
standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq |
17870 | 852 |
(HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)), |
18010 | 853 |
foldr1 HOLogic.mk_conj eqs))) |
854 |
(fn _ => EVERY |
|
17870 | 855 |
[asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: |
856 |
rep_inject_thms')) 1, |
|
857 |
TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def :: |
|
858 |
alpha @ abs_perm @ abs_fresh @ rep_inject_thms @ |
|
17874
8be65cf94d2e
Improved proof of injectivity theorems to make it work on
berghofe
parents:
17873
diff
changeset
|
859 |
perm_rep_perm_thms)) 1), |
8be65cf94d2e
Improved proof of injectivity theorems to make it work on
berghofe
parents:
17873
diff
changeset
|
860 |
TRY (asm_full_simp_tac (HOL_basic_ss addsimps (perm_fun_def :: |
18010 | 861 |
expand_fun_eq :: rep_inject_thms @ perm_rep_perm_thms)) 1)])) |
17870 | 862 |
end) (constrs ~~ constr_rep_thms) |
863 |
end) (List.take (descr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); |
|
864 |
||
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
865 |
(** equations for support and freshness **) |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
866 |
|
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
867 |
val Un_assoc = PureThy.get_thm thy8 (Name "Un_assoc"); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
868 |
val de_Morgan_conj = PureThy.get_thm thy8 (Name "de_Morgan_conj"); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
869 |
val Collect_disj_eq = PureThy.get_thm thy8 (Name "Collect_disj_eq"); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
870 |
val finite_Un = PureThy.get_thm thy8 (Name "finite_Un"); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
871 |
|
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
872 |
val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
873 |
(map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') => |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
874 |
let val T = replace_types' (nth_dtyp i) |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
875 |
in List.concat (map (fn (cname, dts) => map (fn atom => |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
876 |
let |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
877 |
val cname = Sign.intern_const thy8 |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
878 |
(NameSpace.append tname (Sign.base_name cname)); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
879 |
val atomT = Type (atom, []); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
880 |
|
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
881 |
fun process_constr (dt, (j, args1, args2)) = |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
882 |
let |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
883 |
val x' = mk_Free "x" (typ_of_dtyp' dt) j; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
884 |
val (dts, dt') = strip_option dt; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
885 |
val Ts_idx = map typ_of_dtyp' dts ~~ (j upto j + length dts - 1); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
886 |
val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
887 |
val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
888 |
val (dts', dt'') = strip_dtyp dt'; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
889 |
in case dt'' of |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
890 |
DtRec k => if k < length new_type_names then |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
891 |
(j + length dts + 1, |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
892 |
xs @ (x :: args1), foldr mk_abs_fun x xs :: args2) |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
893 |
else error "nested recursion not (yet) supported" |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
894 |
| _ => (j + 1, x' :: args1, x' :: args2) |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
895 |
end; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
896 |
|
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
897 |
val (_, args1, args2) = foldr process_constr (1, [], []) dts; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
898 |
val Ts = map fastype_of args1; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
899 |
val c = list_comb (Const (cname, Ts ---> T), args1); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
900 |
fun supp t = |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
901 |
Const ("nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
902 |
fun fresh t = |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
903 |
Const ("nominal.fresh", atomT --> fastype_of t --> HOLogic.boolT) $ |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
904 |
Free ("a", atomT) $ t; |
18010 | 905 |
val supp_thm = standard (Goal.prove thy8 [] [] |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
906 |
(HOLogic.mk_Trueprop (HOLogic.mk_eq |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
907 |
(supp c, |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
908 |
if null dts then Const ("{}", HOLogic.mk_setT atomT) |
18010 | 909 |
else foldr1 (HOLogic.mk_binop "op Un") (map supp args2)))) |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
910 |
(fn _ => |
18010 | 911 |
simp_tac (HOL_basic_ss addsimps (supp_def :: |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
912 |
Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un :: |
17874
8be65cf94d2e
Improved proof of injectivity theorems to make it work on
berghofe
parents:
17873
diff
changeset
|
913 |
symmetric empty_def :: Finites.emptyI :: simp_thms @ |
18010 | 914 |
abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1)) |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
915 |
in |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
916 |
(supp_thm, |
18010 | 917 |
standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
918 |
(fresh c, |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
919 |
if null dts then HOLogic.true_const |
18010 | 920 |
else foldr1 HOLogic.mk_conj (map fresh args2)))) |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
921 |
(fn _ => |
18010 | 922 |
simp_tac (simpset_of thy8 addsimps [fresh_def, supp_thm]) 1))) |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
923 |
end) atoms) constrs) |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
924 |
end) (List.take (descr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps'))); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
925 |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
926 |
(**** weak induction theorem ****) |
18016 | 927 |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
928 |
val arities = get_arities descr''; |
18016 | 929 |
|
930 |
fun mk_funs_inv thm = |
|
931 |
let |
|
932 |
val {sign, prop, ...} = rep_thm thm; |
|
933 |
val _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ S)) $ |
|
934 |
(_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop; |
|
935 |
val used = add_term_tfree_names (a, []); |
|
936 |
||
937 |
fun mk_thm i = |
|
938 |
let |
|
939 |
val Ts = map (TFree o rpair HOLogic.typeS) |
|
940 |
(variantlist (replicate i "'t", used)); |
|
941 |
val f = Free ("f", Ts ---> U) |
|
942 |
in standard (Goal.prove sign [] [] (Logic.mk_implies |
|
943 |
(HOLogic.mk_Trueprop (HOLogic.list_all |
|
944 |
(map (pair "x") Ts, HOLogic.mk_mem (app_bnds f i, S))), |
|
945 |
HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts, |
|
946 |
r $ (a $ app_bnds f i)), f)))) |
|
947 |
(fn _ => EVERY [REPEAT (rtac ext 1), REPEAT (etac allE 1), rtac thm 1, atac 1])) |
|
948 |
end |
|
949 |
in map (fn r => r RS subst) (thm :: map mk_thm arities) end; |
|
950 |
||
951 |
fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) = |
|
952 |
let |
|
953 |
val Rep_t = Const (List.nth (rep_names, i), T --> U) $ |
|
954 |
mk_Free "x" T i; |
|
955 |
||
956 |
val Abs_t = Const (List.nth (abs_names, i), U --> T) |
|
957 |
||
958 |
in (prems @ [HOLogic.imp $ HOLogic.mk_mem (Rep_t, |
|
959 |
Const (List.nth (rep_set_names, i), HOLogic.mk_setT U)) $ |
|
960 |
(mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))], |
|
961 |
concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i]) |
|
962 |
end; |
|
963 |
||
964 |
val (indrule_lemma_prems, indrule_lemma_concls) = |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
965 |
Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs')); |
18016 | 966 |
|
967 |
val indrule_lemma = standard (Goal.prove thy8 [] [] |
|
968 |
(Logic.mk_implies |
|
969 |
(HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems), |
|
970 |
HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY |
|
971 |
[REPEAT (etac conjE 1), |
|
972 |
REPEAT (EVERY |
|
973 |
[TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1, |
|
974 |
etac mp 1, resolve_tac Rep_thms 1])])); |
|
975 |
||
976 |
val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma))); |
|
977 |
val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else |
|
978 |
map (Free o apfst fst o dest_Var) Ps; |
|
979 |
val indrule_lemma' = cterm_instantiate |
|
980 |
(map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma; |
|
981 |
||
982 |
val Abs_inverse_thms' = List.concat (map mk_funs_inv Abs_inverse_thms); |
|
983 |
||
984 |
val dt_induct_prop = DatatypeProp.make_ind descr' sorts'; |
|
985 |
val dt_induct = standard (Goal.prove thy8 [] |
|
986 |
(Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) |
|
987 |
(fn prems => EVERY |
|
988 |
[rtac indrule_lemma' 1, |
|
989 |
(DatatypeAux.indtac rep_induct THEN_ALL_NEW ObjectLogic.atomize_tac) 1, |
|
990 |
EVERY (map (fn (prem, r) => (EVERY |
|
991 |
[REPEAT (eresolve_tac Abs_inverse_thms' 1), |
|
992 |
simp_tac (HOL_basic_ss addsimps [symmetric r]) 1, |
|
993 |
DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)])) |
|
994 |
(prems ~~ constr_defs))])); |
|
995 |
||
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
996 |
val case_names_induct = mk_case_names_induct descr''; |
18016 | 997 |
|
18066
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
998 |
(**** prove that new datatypes have finite support ****) |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
999 |
|
18246
676d2e625d98
added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
18245
diff
changeset
|
1000 |
val _ = warning "proving finite support for the new datatype"; |
676d2e625d98
added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
18245
diff
changeset
|
1001 |
|
18066
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1002 |
val indnames = DatatypeProp.make_tnames recTs; |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1003 |
|
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1004 |
val abs_supp = PureThy.get_thms thy8 (Name "abs_supp"); |
18067 | 1005 |
val supp_atm = PureThy.get_thms thy8 (Name "supp_atm"); |
18066
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1006 |
|
18246
676d2e625d98
added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
18245
diff
changeset
|
1007 |
|
676d2e625d98
added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
18245
diff
changeset
|
1008 |
(* |
18066
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1009 |
val finite_supp_thms = map (fn atom => |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1010 |
let val atomT = Type (atom, []) |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1011 |
in map standard (List.take |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1012 |
(split_conj_thm (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1013 |
(foldr1 HOLogic.mk_conj (map (fn (s, T) => HOLogic.mk_mem |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1014 |
(Const ("nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T), |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1015 |
Const ("Finite_Set.Finites", HOLogic.mk_setT (HOLogic.mk_setT atomT)))) |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1016 |
(indnames ~~ recTs)))) |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1017 |
(fn _ => indtac dt_induct indnames 1 THEN |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1018 |
ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps |
18067 | 1019 |
(abs_supp @ supp_atm @ |
18066
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1020 |
PureThy.get_thms thy8 (Name ("fs_" ^ Sign.base_name atom ^ "1")) @ |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1021 |
List.concat supp_thms))))), |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1022 |
length new_type_names)) |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1023 |
end) atoms; |
18246
676d2e625d98
added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
18245
diff
changeset
|
1024 |
*) |
18066
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1025 |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1026 |
(**** strong induction theorem ****) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1027 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1028 |
val pnames = if length descr'' = 1 then ["P"] |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1029 |
else map (fn i => "P" ^ string_of_int i) (1 upto length descr''); |
18245
65e60434b3c2
Fixed problem with strong induction theorem for datatypes containing
berghofe
parents:
18142
diff
changeset
|
1030 |
val ind_sort = if null dt_atomTs then HOLogic.typeS |
65e60434b3c2
Fixed problem with strong induction theorem for datatypes containing
berghofe
parents:
18142
diff
changeset
|
1031 |
else map (fn T => Sign.intern_class thy8 ("fs_" ^ |
65e60434b3c2
Fixed problem with strong induction theorem for datatypes containing
berghofe
parents:
18142
diff
changeset
|
1032 |
Sign.base_name (fst (dest_Type T)))) dt_atomTs; |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1033 |
val fsT = TFree ("'n", ind_sort); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1034 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1035 |
fun make_pred i T = |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1036 |
Free (List.nth (pnames, i), T --> fsT --> HOLogic.boolT); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1037 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1038 |
fun make_ind_prem k T ((cname, cargs), idxs) = |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1039 |
let |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1040 |
val recs = List.filter is_rec_type cargs; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1041 |
val Ts = map (typ_of_dtyp descr'' sorts') cargs; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1042 |
val recTs' = map (typ_of_dtyp descr'' sorts') recs; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1043 |
val tnames = variantlist (DatatypeProp.make_tnames Ts, pnames); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1044 |
val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs)); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1045 |
val frees = tnames ~~ Ts; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1046 |
val z = (variant tnames "z", fsT); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1047 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1048 |
fun mk_prem ((dt, s), T) = |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1049 |
let |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1050 |
val (Us, U) = strip_type T; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1051 |
val l = length Us |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1052 |
in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1053 |
(make_pred (body_index dt) U $ app_bnds (Free (s, T)) l $ Bound l)) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1054 |
end; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1055 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1056 |
val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs'); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1057 |
val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1058 |
(Const ("nominal.fresh", T --> fsT --> HOLogic.boolT) $ |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1059 |
Free p $ Free z)) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1060 |
(map (curry List.nth frees) (List.concat (map (fn (m, n) => |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1061 |
m upto m + n - 1) idxs))) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1062 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1063 |
in list_all_free (z :: frees, Logic.list_implies (prems' @ prems, |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1064 |
HOLogic.mk_Trueprop (make_pred k T $ |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1065 |
list_comb (Const (cname, Ts ---> T), map Free frees) $ Free z))) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1066 |
end; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1067 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1068 |
val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) => |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1069 |
map (make_ind_prem i T) (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs)); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1070 |
val tnames = DatatypeProp.make_tnames recTs; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1071 |
val z = (variant tnames "z", fsT); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1072 |
val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1073 |
(map (fn (((i, _), T), tname) => make_pred i T $ Free (tname, T) $ Free z) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1074 |
(descr'' ~~ recTs ~~ tnames))); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1075 |
val induct = Logic.list_implies (ind_prems, ind_concl); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1076 |
|
18104
dbe58b104cb9
added thms perm, distinct and fresh to the simplifier.
urbanc
parents:
18068
diff
changeset
|
1077 |
val simp_atts = replicate (length new_type_names) [Simplifier.simp_add_global]; |
dbe58b104cb9
added thms perm, distinct and fresh to the simplifier.
urbanc
parents:
18068
diff
changeset
|
1078 |
|
17870 | 1079 |
val (thy9, _) = thy8 |> |
18016 | 1080 |
Theory.add_path big_name |> |
18017
f6abeac6dcb5
fixed case names in the weak induction principle and
urbanc
parents:
18016
diff
changeset
|
1081 |
PureThy.add_thms [(("induct_weak", dt_induct), [case_names_induct])] |>> |
18016 | 1082 |
Theory.parent_path |>>> |
18104
dbe58b104cb9
added thms perm, distinct and fresh to the simplifier.
urbanc
parents:
18068
diff
changeset
|
1083 |
DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms |>>> |
17870 | 1084 |
DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss |>>> |
18104
dbe58b104cb9
added thms perm, distinct and fresh to the simplifier.
urbanc
parents:
18068
diff
changeset
|
1085 |
DatatypeAux.store_thmss_atts "perm" new_type_names simp_atts perm_simps' |>>> |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
1086 |
DatatypeAux.store_thmss "inject" new_type_names inject_thms |>>> |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
1087 |
DatatypeAux.store_thmss "supp" new_type_names supp_thms |>>> |
18104
dbe58b104cb9
added thms perm, distinct and fresh to the simplifier.
urbanc
parents:
18068
diff
changeset
|
1088 |
DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms |>> |
18246
676d2e625d98
added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
18245
diff
changeset
|
1089 |
(* fold (fn (atom, ths) => fn thy => |
18066
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1090 |
let val class = Sign.intern_class thy ("fs_" ^ Sign.base_name atom) |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1091 |
in fold (fn T => AxClass.add_inst_arity_i |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1092 |
(fst (dest_Type T), |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1093 |
replicate (length sorts) [class], [class]) |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1094 |
(AxClass.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy |
18246
676d2e625d98
added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
18245
diff
changeset
|
1095 |
end) (atoms ~~ finite_supp_thms) |>> *) |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1096 |
Theory.add_path big_name |>>> |
18142
a51ab4152097
called the induction principle "unsafe" instead of "test".
urbanc
parents:
18107
diff
changeset
|
1097 |
PureThy.add_axioms_i [(("induct_unsafe", induct), [case_names_induct])] |>> |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1098 |
Theory.parent_path; |
17870 | 1099 |
|
1100 |
in |
|
18068 | 1101 |
thy9 |
17870 | 1102 |
end; |
1103 |
||
1104 |
val add_nominal_datatype = gen_add_nominal_datatype read_typ true; |
|
1105 |
||
1106 |
||
1107 |
(* FIXME: The following stuff should be exported by DatatypePackage *) |
|
1108 |
||
1109 |
local structure P = OuterParse and K = OuterKeyword in |
|
1110 |
||
1111 |
val datatype_decl = |
|
1112 |
Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix -- |
|
1113 |
(P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix)); |
|
1114 |
||
1115 |
fun mk_datatype args = |
|
1116 |
let |
|
1117 |
val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args; |
|
1118 |
val specs = map (fn ((((_, vs), t), mx), cons) => |
|
1119 |
(vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args; |
|
18068 | 1120 |
in add_nominal_datatype false names specs end; |
17870 | 1121 |
|
1122 |
val nominal_datatypeP = |
|
1123 |
OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl |
|
1124 |
(P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype)); |
|
1125 |
||
1126 |
val _ = OuterSyntax.add_parsers [nominal_datatypeP]; |
|
1127 |
||
1128 |
end; |
|
1129 |
||
1130 |
end |