author | wenzelm |
Sat, 13 May 2006 02:51:48 +0200 | |
changeset 19635 | f7aa7d174343 |
parent 19494 | 2e909d5309f4 |
child 19649 | c887656778bc |
permissions | -rw-r--r-- |
19494 | 1 |
(* Title: HOL/Nominal/nominal_package.ML |
2 |
ID: $Id$ |
|
3 |
Author: Stefan Berghofer and Christian Urban, TU Muenchen |
|
4 |
||
5 |
Nominal datatype package for Isabelle/HOL. |
|
6 |
*) |
|
17870 | 7 |
|
8 |
signature NOMINAL_PACKAGE = |
|
9 |
sig |
|
10 |
val add_nominal_datatype : bool -> string list -> (string list * bstring * mixfix * |
|
18068 | 11 |
(bstring * string list * mixfix) list) list -> theory -> theory |
17870 | 12 |
end |
13 |
||
18068 | 14 |
structure NominalPackage : NOMINAL_PACKAGE = |
17870 | 15 |
struct |
16 |
||
17 |
open DatatypeAux; |
|
18068 | 18 |
open NominalAtoms; |
17870 | 19 |
|
18016 | 20 |
(** FIXME: DatatypePackage should export this function **) |
21 |
||
22 |
local |
|
23 |
||
24 |
fun dt_recs (DtTFree _) = [] |
|
25 |
| dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts) |
|
26 |
| dt_recs (DtRec i) = [i]; |
|
27 |
||
28 |
fun dt_cases (descr: descr) (_, args, constrs) = |
|
29 |
let |
|
30 |
fun the_bname i = Sign.base_name (#1 (valOf (AList.lookup (op =) descr i))); |
|
19133
7e84a1a3741c
Adapted to Florian's recent changes to the AxClass package.
berghofe
parents:
18759
diff
changeset
|
31 |
val bnames = map the_bname (distinct op = (List.concat (map dt_recs args))); |
18016 | 32 |
in map (fn (c, _) => space_implode "_" (Sign.base_name c :: bnames)) constrs end; |
33 |
||
34 |
||
35 |
fun induct_cases descr = |
|
36 |
DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr))); |
|
37 |
||
38 |
fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i)); |
|
39 |
||
40 |
in |
|
41 |
||
42 |
fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr); |
|
43 |
||
44 |
fun mk_case_names_exhausts descr new = |
|
45 |
map (RuleCases.case_names o exhaust_cases descr o #1) |
|
46 |
(List.filter (fn ((_, (name, _, _))) => name mem_string new) descr); |
|
47 |
||
48 |
end; |
|
49 |
||
50 |
(*******************************) |
|
51 |
||
17870 | 52 |
val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma); |
53 |
||
54 |
fun read_typ sign ((Ts, sorts), str) = |
|
55 |
let |
|
56 |
val T = Type.no_tvars (Sign.read_typ (sign, (AList.lookup op =) |
|
57 |
(map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg |
|
58 |
in (Ts @ [T], add_typ_tfrees (T, sorts)) end; |
|
59 |
||
60 |
(** taken from HOL/Tools/datatype_aux.ML **) |
|
61 |
||
62 |
fun indtac indrule indnames i st = |
|
63 |
let |
|
64 |
val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule)); |
|
65 |
val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop |
|
66 |
(Logic.strip_imp_concl (List.nth (prems_of st, i - 1)))); |
|
67 |
val getP = if can HOLogic.dest_imp (hd ts) then |
|
68 |
(apfst SOME) o HOLogic.dest_imp else pair NONE; |
|
69 |
fun abstr (t1, t2) = (case t1 of |
|
70 |
NONE => (case filter (fn Free (s, _) => s mem indnames | _ => false) |
|
71 |
(term_frees t2) of |
|
72 |
[Free (s, T)] => absfree (s, T, t2) |
|
73 |
| _ => sys_error "indtac") |
|
74 |
| SOME (_ $ t' $ _) => Abs ("x", fastype_of t', abstract_over (t', t2))) |
|
75 |
val cert = cterm_of (Thm.sign_of_thm st); |
|
76 |
val Ps = map (cert o head_of o snd o getP) ts; |
|
77 |
val indrule' = cterm_instantiate (Ps ~~ |
|
78 |
(map (cert o abstr o getP) ts')) indrule |
|
79 |
in |
|
80 |
rtac indrule' i st |
|
81 |
end; |
|
82 |
||
18658 | 83 |
fun mk_subgoal i f st = |
84 |
let |
|
85 |
val cg = List.nth (cprems_of st, i - 1); |
|
86 |
val g = term_of cg; |
|
87 |
val revcut_rl' = Thm.lift_rule cg revcut_rl; |
|
88 |
val v = head_of (Logic.strip_assums_concl (hd (prems_of revcut_rl'))); |
|
89 |
val ps = Logic.strip_params g; |
|
90 |
val cert = cterm_of (sign_of_thm st); |
|
91 |
in |
|
92 |
compose_tac (false, |
|
93 |
Thm.instantiate ([], [(cert v, cert (list_abs (ps, |
|
94 |
f (rev ps) (Logic.strip_assums_hyp g) (Logic.strip_assums_concl g))))]) |
|
95 |
revcut_rl', 2) i st |
|
96 |
end; |
|
97 |
||
98 |
(** simplification procedure for sorting permutations **) |
|
99 |
||
100 |
val dj_cp = thm "dj_cp"; |
|
101 |
||
102 |
fun dest_permT (Type ("fun", [Type ("List.list", [Type ("*", [T, _])]), |
|
103 |
Type ("fun", [_, U])])) = (T, U); |
|
104 |
||
19494 | 105 |
fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u |
18658 | 106 |
| permTs_of _ = []; |
107 |
||
19494 | 108 |
fun perm_simproc' thy ss (Const ("Nominal.perm", T) $ t $ (u as Const ("Nominal.perm", U) $ r $ s)) = |
18658 | 109 |
let |
110 |
val (aT as Type (a, []), S) = dest_permT T; |
|
111 |
val (bT as Type (b, []), _) = dest_permT U |
|
112 |
in if aT mem permTs_of u andalso aT <> bT then |
|
113 |
let |
|
114 |
val a' = Sign.base_name a; |
|
115 |
val b' = Sign.base_name b; |
|
116 |
val cp = PureThy.get_thm thy (Name ("cp_" ^ a' ^ "_" ^ b' ^ "_inst")); |
|
117 |
val dj = PureThy.get_thm thy (Name ("dj_" ^ b' ^ "_" ^ a')); |
|
118 |
val dj_cp' = [cp, dj] MRS dj_cp; |
|
119 |
val cert = SOME o cterm_of thy |
|
120 |
in |
|
121 |
SOME (mk_meta_eq (Drule.instantiate' [SOME (ctyp_of thy S)] |
|
122 |
[cert t, cert r, cert s] dj_cp')) |
|
123 |
end |
|
124 |
else NONE |
|
125 |
end |
|
126 |
| perm_simproc' thy ss _ = NONE; |
|
127 |
||
128 |
val perm_simproc = |
|
129 |
Simplifier.simproc (the_context ()) "perm_simp" ["pi1 \\<bullet> (pi2 \\<bullet> x)"] perm_simproc'; |
|
130 |
||
131 |
val allE_Nil = read_instantiate_sg (the_context()) [("x", "[]")] allE; |
|
132 |
||
133 |
val meta_spec = thm "meta_spec"; |
|
134 |
||
18582
4f4cc426b440
provide projections of induct_weak, induct_unsafe;
wenzelm
parents:
18579
diff
changeset
|
135 |
fun projections rule = |
4f4cc426b440
provide projections of induct_weak, induct_unsafe;
wenzelm
parents:
18579
diff
changeset
|
136 |
ProjectRule.projections rule |
4f4cc426b440
provide projections of induct_weak, induct_unsafe;
wenzelm
parents:
18579
diff
changeset
|
137 |
|> map (standard #> RuleCases.save rule); |
4f4cc426b440
provide projections of induct_weak, induct_unsafe;
wenzelm
parents:
18579
diff
changeset
|
138 |
|
18658 | 139 |
fun norm_sort thy = Sorts.norm_sort (snd (#classes (Type.rep_tsig (Sign.tsig_of thy)))); |
140 |
||
17870 | 141 |
fun gen_add_nominal_datatype prep_typ err flat_names new_type_names dts thy = |
142 |
let |
|
143 |
(* this theory is used just for parsing *) |
|
144 |
||
145 |
val tmp_thy = thy |> |
|
146 |
Theory.copy |> |
|
147 |
Theory.add_types (map (fn (tvs, tname, mx, _) => |
|
148 |
(tname, length tvs, mx)) dts); |
|
149 |
||
150 |
val sign = Theory.sign_of tmp_thy; |
|
151 |
||
152 |
val atoms = atoms_of thy; |
|
153 |
val classes = map (NameSpace.map_base (fn s => "pt_" ^ s)) atoms; |
|
154 |
val cp_classes = List.concat (map (fn atom1 => map (fn atom2 => |
|
155 |
Sign.intern_class thy ("cp_" ^ Sign.base_name atom1 ^ "_" ^ |
|
156 |
Sign.base_name atom2)) atoms) atoms); |
|
157 |
fun augment_sort S = S union classes; |
|
158 |
val augment_sort_typ = map_type_tfree (fn (s, S) => TFree (s, augment_sort S)); |
|
159 |
||
160 |
fun prep_constr ((constrs, sorts), (cname, cargs, mx)) = |
|
161 |
let val (cargs', sorts') = Library.foldl (prep_typ sign) (([], sorts), cargs) |
|
162 |
in (constrs @ [(cname, cargs', mx)], sorts') end |
|
163 |
||
164 |
fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) = |
|
165 |
let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs) |
|
166 |
in (dts @ [(tvs, tname, mx, constrs')], sorts') end |
|
167 |
||
168 |
val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts); |
|
169 |
val sorts' = map (apsnd augment_sort) sorts; |
|
170 |
val tyvars = map #1 dts'; |
|
171 |
||
172 |
val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts'; |
|
173 |
val constr_syntax = map (fn (tvs, tname, mx, constrs) => |
|
174 |
map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts'; |
|
175 |
||
176 |
val ps = map (fn (_, n, _, _) => |
|
177 |
(Sign.full_name sign n, Sign.full_name sign (n ^ "_Rep"))) dts; |
|
178 |
val rps = map Library.swap ps; |
|
179 |
||
19494 | 180 |
fun replace_types (Type ("Nominal.ABS", [T, U])) = |
181 |
Type ("fun", [T, Type ("Nominal.noption", [replace_types U])]) |
|
17870 | 182 |
| replace_types (Type (s, Ts)) = |
183 |
Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts) |
|
184 |
| replace_types T = T; |
|
185 |
||
186 |
fun replace_types' (Type (s, Ts)) = |
|
187 |
Type (getOpt (AList.lookup op = rps s, s), map replace_types' Ts) |
|
188 |
| replace_types' T = T; |
|
189 |
||
190 |
val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, tname ^ "_Rep", NoSyn, |
|
191 |
map (fn (cname, cargs, mx) => (cname, |
|
192 |
map (augment_sort_typ o replace_types) cargs, NoSyn)) constrs)) dts'; |
|
193 |
||
194 |
val new_type_names' = map (fn n => n ^ "_Rep") new_type_names; |
|
195 |
val full_new_type_names' = map (Sign.full_name (sign_of thy)) new_type_names'; |
|
196 |
||
18045 | 197 |
val ({induction, ...},thy1) = |
17870 | 198 |
DatatypePackage.add_datatype_i err flat_names new_type_names' dts'' thy; |
199 |
||
200 |
val SOME {descr, ...} = Symtab.lookup |
|
201 |
(DatatypePackage.get_datatypes thy1) (hd full_new_type_names'); |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
202 |
fun nth_dtyp i = typ_of_dtyp descr sorts' (DtRec i); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
203 |
|
17870 | 204 |
(**** define permutation functions ****) |
205 |
||
206 |
val permT = mk_permT (TFree ("'x", HOLogic.typeS)); |
|
207 |
val pi = Free ("pi", permT); |
|
208 |
val perm_types = map (fn (i, _) => |
|
209 |
let val T = nth_dtyp i |
|
210 |
in permT --> T --> T end) descr; |
|
19494 | 211 |
val perm_names = replicate (length new_type_names) "Nominal.perm" @ |
17870 | 212 |
DatatypeProp.indexify_names (map (fn i => Sign.full_name (sign_of thy1) |
213 |
("perm_" ^ name_of_typ (nth_dtyp i))) |
|
214 |
(length new_type_names upto length descr - 1)); |
|
215 |
val perm_names_types = perm_names ~~ perm_types; |
|
216 |
||
217 |
val perm_eqs = List.concat (map (fn (i, (_, _, constrs)) => |
|
218 |
let val T = nth_dtyp i |
|
219 |
in map (fn (cname, dts) => |
|
220 |
let |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
221 |
val Ts = map (typ_of_dtyp descr sorts') dts; |
17870 | 222 |
val names = DatatypeProp.make_tnames Ts; |
223 |
val args = map Free (names ~~ Ts); |
|
224 |
val c = Const (cname, Ts ---> T); |
|
225 |
fun perm_arg (dt, x) = |
|
226 |
let val T = type_of x |
|
227 |
in if is_rec_type dt then |
|
228 |
let val (Us, _) = strip_type T |
|
229 |
in list_abs (map (pair "x") Us, |
|
230 |
Const (List.nth (perm_names_types, body_index dt)) $ pi $ |
|
231 |
list_comb (x, map (fn (i, U) => |
|
19494 | 232 |
Const ("Nominal.perm", permT --> U --> U) $ |
17870 | 233 |
(Const ("List.rev", permT --> permT) $ pi) $ |
234 |
Bound i) ((length Us - 1 downto 0) ~~ Us))) |
|
235 |
end |
|
19494 | 236 |
else Const ("Nominal.perm", permT --> T --> T) $ pi $ x |
17870 | 237 |
end; |
238 |
in |
|
239 |
(("", HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
240 |
(Const (List.nth (perm_names_types, i)) $ |
|
241 |
Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $ |
|
242 |
list_comb (c, args), |
|
243 |
list_comb (c, map perm_arg (dts ~~ args))))), []) |
|
244 |
end) constrs |
|
245 |
end) descr); |
|
246 |
||
247 |
val (thy2, perm_simps) = thy1 |> |
|
248 |
Theory.add_consts_i (map (fn (s, T) => (Sign.base_name s, T, NoSyn)) |
|
249 |
(List.drop (perm_names_types, length new_type_names))) |> |
|
19635 | 250 |
PrimrecPackage.add_primrec_unchecked_i "" perm_eqs; |
17870 | 251 |
|
252 |
(**** prove that permutation functions introduced by unfolding are ****) |
|
253 |
(**** equivalent to already existing permutation functions ****) |
|
254 |
||
255 |
val _ = warning ("length descr: " ^ string_of_int (length descr)); |
|
256 |
val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names)); |
|
257 |
||
258 |
val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types); |
|
259 |
val perm_fun_def = PureThy.get_thm thy2 (Name "perm_fun_def"); |
|
260 |
||
261 |
val unfolded_perm_eq_thms = |
|
262 |
if length descr = length new_type_names then [] |
|
263 |
else map standard (List.drop (split_conj_thm |
|
18010 | 264 |
(Goal.prove thy2 [] [] |
17870 | 265 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
266 |
(map (fn (c as (s, T), x) => |
|
267 |
let val [T1, T2] = binder_types T |
|
268 |
in HOLogic.mk_eq (Const c $ pi $ Free (x, T2), |
|
19494 | 269 |
Const ("Nominal.perm", T) $ pi $ Free (x, T2)) |
17870 | 270 |
end) |
18010 | 271 |
(perm_names_types ~~ perm_indnames)))) |
272 |
(fn _ => EVERY [indtac induction perm_indnames 1, |
|
17870 | 273 |
ALLGOALS (asm_full_simp_tac |
274 |
(simpset_of thy2 addsimps [perm_fun_def]))])), |
|
275 |
length new_type_names)); |
|
276 |
||
277 |
(**** prove [] \<bullet> t = t ****) |
|
278 |
||
279 |
val _ = warning "perm_empty_thms"; |
|
280 |
||
281 |
val perm_empty_thms = List.concat (map (fn a => |
|
282 |
let val permT = mk_permT (Type (a, [])) |
|
283 |
in map standard (List.take (split_conj_thm |
|
18010 | 284 |
(Goal.prove thy2 [] [] |
17870 | 285 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
286 |
(map (fn ((s, T), x) => HOLogic.mk_eq |
|
287 |
(Const (s, permT --> T --> T) $ |
|
288 |
Const ("List.list.Nil", permT) $ Free (x, T), |
|
289 |
Free (x, T))) |
|
290 |
(perm_names ~~ |
|
18010 | 291 |
map body_type perm_types ~~ perm_indnames)))) |
292 |
(fn _ => EVERY [indtac induction perm_indnames 1, |
|
17870 | 293 |
ALLGOALS (asm_full_simp_tac (simpset_of thy2))])), |
294 |
length new_type_names)) |
|
295 |
end) |
|
296 |
atoms); |
|
297 |
||
298 |
(**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****) |
|
299 |
||
300 |
val _ = warning "perm_append_thms"; |
|
301 |
||
302 |
(*FIXME: these should be looked up statically*) |
|
303 |
val at_pt_inst = PureThy.get_thm thy2 (Name "at_pt_inst"); |
|
304 |
val pt2 = PureThy.get_thm thy2 (Name "pt2"); |
|
305 |
||
306 |
val perm_append_thms = List.concat (map (fn a => |
|
307 |
let |
|
308 |
val permT = mk_permT (Type (a, [])); |
|
309 |
val pi1 = Free ("pi1", permT); |
|
310 |
val pi2 = Free ("pi2", permT); |
|
311 |
val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst")); |
|
312 |
val pt2' = pt_inst RS pt2; |
|
313 |
val pt2_ax = PureThy.get_thm thy2 |
|
314 |
(Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "2") a)); |
|
315 |
in List.take (map standard (split_conj_thm |
|
18010 | 316 |
(Goal.prove thy2 [] [] |
17870 | 317 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
318 |
(map (fn ((s, T), x) => |
|
319 |
let val perm = Const (s, permT --> T --> T) |
|
320 |
in HOLogic.mk_eq |
|
321 |
(perm $ (Const ("List.op @", permT --> permT --> permT) $ |
|
322 |
pi1 $ pi2) $ Free (x, T), |
|
323 |
perm $ pi1 $ (perm $ pi2 $ Free (x, T))) |
|
324 |
end) |
|
325 |
(perm_names ~~ |
|
18010 | 326 |
map body_type perm_types ~~ perm_indnames)))) |
327 |
(fn _ => EVERY [indtac induction perm_indnames 1, |
|
17870 | 328 |
ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))), |
329 |
length new_type_names) |
|
330 |
end) atoms); |
|
331 |
||
332 |
(**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****) |
|
333 |
||
334 |
val _ = warning "perm_eq_thms"; |
|
335 |
||
336 |
val pt3 = PureThy.get_thm thy2 (Name "pt3"); |
|
337 |
val pt3_rev = PureThy.get_thm thy2 (Name "pt3_rev"); |
|
338 |
||
339 |
val perm_eq_thms = List.concat (map (fn a => |
|
340 |
let |
|
341 |
val permT = mk_permT (Type (a, [])); |
|
342 |
val pi1 = Free ("pi1", permT); |
|
343 |
val pi2 = Free ("pi2", permT); |
|
344 |
(*FIXME: not robust - better access these theorems using NominalData?*) |
|
345 |
val at_inst = PureThy.get_thm thy2 (Name ("at_" ^ Sign.base_name a ^ "_inst")); |
|
346 |
val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst")); |
|
347 |
val pt3' = pt_inst RS pt3; |
|
348 |
val pt3_rev' = at_inst RS (pt_inst RS pt3_rev); |
|
349 |
val pt3_ax = PureThy.get_thm thy2 |
|
350 |
(Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "3") a)); |
|
351 |
in List.take (map standard (split_conj_thm |
|
18010 | 352 |
(Goal.prove thy2 [] [] (Logic.mk_implies |
19494 | 353 |
(HOLogic.mk_Trueprop (Const ("Nominal.prm_eq", |
17870 | 354 |
permT --> permT --> HOLogic.boolT) $ pi1 $ pi2), |
355 |
HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
|
356 |
(map (fn ((s, T), x) => |
|
357 |
let val perm = Const (s, permT --> T --> T) |
|
358 |
in HOLogic.mk_eq |
|
359 |
(perm $ pi1 $ Free (x, T), |
|
360 |
perm $ pi2 $ Free (x, T)) |
|
361 |
end) |
|
362 |
(perm_names ~~ |
|
18010 | 363 |
map body_type perm_types ~~ perm_indnames))))) |
364 |
(fn _ => EVERY [indtac induction perm_indnames 1, |
|
17870 | 365 |
ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))), |
366 |
length new_type_names) |
|
367 |
end) atoms); |
|
368 |
||
369 |
(**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****) |
|
370 |
||
371 |
val cp1 = PureThy.get_thm thy2 (Name "cp1"); |
|
372 |
val dj_cp = PureThy.get_thm thy2 (Name "dj_cp"); |
|
373 |
val pt_perm_compose = PureThy.get_thm thy2 (Name "pt_perm_compose"); |
|
374 |
val pt_perm_compose_rev = PureThy.get_thm thy2 (Name "pt_perm_compose_rev"); |
|
375 |
val dj_perm_perm_forget = PureThy.get_thm thy2 (Name "dj_perm_perm_forget"); |
|
376 |
||
377 |
fun composition_instance name1 name2 thy = |
|
378 |
let |
|
379 |
val name1' = Sign.base_name name1; |
|
380 |
val name2' = Sign.base_name name2; |
|
381 |
val cp_class = Sign.intern_class thy ("cp_" ^ name1' ^ "_" ^ name2'); |
|
382 |
val permT1 = mk_permT (Type (name1, [])); |
|
383 |
val permT2 = mk_permT (Type (name2, [])); |
|
384 |
val augment = map_type_tfree |
|
385 |
(fn (x, S) => TFree (x, cp_class :: S)); |
|
386 |
val Ts = map (augment o body_type) perm_types; |
|
387 |
val cp_inst = PureThy.get_thm thy |
|
388 |
(Name ("cp_" ^ name1' ^ "_" ^ name2' ^ "_inst")); |
|
389 |
val simps = simpset_of thy addsimps (perm_fun_def :: |
|
390 |
(if name1 <> name2 then |
|
391 |
let val dj = PureThy.get_thm thy (Name ("dj_" ^ name2' ^ "_" ^ name1')) |
|
392 |
in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end |
|
393 |
else |
|
394 |
let |
|
395 |
val at_inst = PureThy.get_thm thy (Name ("at_" ^ name1' ^ "_inst")); |
|
396 |
val pt_inst = PureThy.get_thm thy (Name ("pt_" ^ name1' ^ "_inst")) |
|
397 |
in |
|
398 |
[cp_inst RS cp1 RS sym, |
|
399 |
at_inst RS (pt_inst RS pt_perm_compose) RS sym, |
|
400 |
at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym] |
|
401 |
end)) |
|
18010 | 402 |
val thms = split_conj_thm (standard (Goal.prove thy [] [] |
17870 | 403 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj |
404 |
(map (fn ((s, T), x) => |
|
405 |
let |
|
406 |
val pi1 = Free ("pi1", permT1); |
|
407 |
val pi2 = Free ("pi2", permT2); |
|
408 |
val perm1 = Const (s, permT1 --> T --> T); |
|
409 |
val perm2 = Const (s, permT2 --> T --> T); |
|
19494 | 410 |
val perm3 = Const ("Nominal.perm", permT1 --> permT2 --> permT2) |
17870 | 411 |
in HOLogic.mk_eq |
412 |
(perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)), |
|
413 |
perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T))) |
|
414 |
end) |
|
18010 | 415 |
(perm_names ~~ Ts ~~ perm_indnames)))) |
416 |
(fn _ => EVERY [indtac induction perm_indnames 1, |
|
417 |
ALLGOALS (asm_full_simp_tac simps)]))) |
|
17870 | 418 |
in |
19275 | 419 |
foldl (fn ((s, tvs), thy) => AxClass.prove_arity |
17870 | 420 |
(s, replicate (length tvs) (cp_class :: classes), [cp_class]) |
19133
7e84a1a3741c
Adapted to Florian's recent changes to the AxClass package.
berghofe
parents:
18759
diff
changeset
|
421 |
(ClassPackage.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy) |
17870 | 422 |
thy (full_new_type_names' ~~ tyvars) |
423 |
end; |
|
424 |
||
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
425 |
val (perm_thmss,thy3) = thy2 |> |
17870 | 426 |
fold (fn name1 => fold (composition_instance name1) atoms) atoms |> |
427 |
curry (Library.foldr (fn ((i, (tyname, args, _)), thy) => |
|
19275 | 428 |
AxClass.prove_arity (tyname, replicate (length args) classes, classes) |
19133
7e84a1a3741c
Adapted to Florian's recent changes to the AxClass package.
berghofe
parents:
18759
diff
changeset
|
429 |
(ClassPackage.intro_classes_tac [] THEN REPEAT (EVERY |
17870 | 430 |
[resolve_tac perm_empty_thms 1, |
431 |
resolve_tac perm_append_thms 1, |
|
432 |
resolve_tac perm_eq_thms 1, assume_tac 1])) thy)) |
|
433 |
(List.take (descr, length new_type_names)) |> |
|
434 |
PureThy.add_thmss |
|
435 |
[((space_implode "_" new_type_names ^ "_unfolded_perm_eq", |
|
18759 | 436 |
unfolded_perm_eq_thms), [Simplifier.simp_add]), |
17870 | 437 |
((space_implode "_" new_type_names ^ "_perm_empty", |
18759 | 438 |
perm_empty_thms), [Simplifier.simp_add]), |
17870 | 439 |
((space_implode "_" new_type_names ^ "_perm_append", |
18759 | 440 |
perm_append_thms), [Simplifier.simp_add]), |
17870 | 441 |
((space_implode "_" new_type_names ^ "_perm_eq", |
18759 | 442 |
perm_eq_thms), [Simplifier.simp_add])]; |
17870 | 443 |
|
444 |
(**** Define representing sets ****) |
|
445 |
||
446 |
val _ = warning "representing sets"; |
|
447 |
||
448 |
val rep_set_names = map (Sign.full_name thy3) (DatatypeProp.indexify_names |
|
449 |
(map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr)); |
|
450 |
val big_rep_name = |
|
451 |
space_implode "_" (DatatypeProp.indexify_names (List.mapPartial |
|
19494 | 452 |
(fn (i, ("Nominal.noption", _, _)) => NONE |
17870 | 453 |
| (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set"; |
454 |
val _ = warning ("big_rep_name: " ^ big_rep_name); |
|
455 |
||
456 |
fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) = |
|
457 |
(case AList.lookup op = descr i of |
|
19494 | 458 |
SOME ("Nominal.noption", _, [(_, [dt']), _]) => |
17870 | 459 |
apfst (cons dt) (strip_option dt') |
460 |
| _ => ([], dtf)) |
|
19494 | 461 |
| strip_option (DtType ("fun", [dt, DtType ("Nominal.noption", [dt'])])) = |
18261
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
462 |
apfst (cons dt) (strip_option dt') |
17870 | 463 |
| strip_option dt = ([], dt); |
464 |
||
19133
7e84a1a3741c
Adapted to Florian's recent changes to the AxClass package.
berghofe
parents:
18759
diff
changeset
|
465 |
val dt_atomTs = distinct op = (map (typ_of_dtyp descr sorts') |
18280
45e139675daf
Corrected atom class constraints in strong induction rule.
berghofe
parents:
18261
diff
changeset
|
466 |
(List.concat (map (fn (_, (_, _, cs)) => List.concat |
45e139675daf
Corrected atom class constraints in strong induction rule.
berghofe
parents:
18261
diff
changeset
|
467 |
(map (List.concat o map (fst o strip_option) o snd) cs)) descr))); |
45e139675daf
Corrected atom class constraints in strong induction rule.
berghofe
parents:
18261
diff
changeset
|
468 |
|
17870 | 469 |
fun make_intr s T (cname, cargs) = |
470 |
let |
|
471 |
fun mk_prem (dt, (j, j', prems, ts)) = |
|
472 |
let |
|
473 |
val (dts, dt') = strip_option dt; |
|
474 |
val (dts', dt'') = strip_dtyp dt'; |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
475 |
val Ts = map (typ_of_dtyp descr sorts') dts; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
476 |
val Us = map (typ_of_dtyp descr sorts') dts'; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
477 |
val T = typ_of_dtyp descr sorts' dt''; |
17870 | 478 |
val free = mk_Free "x" (Us ---> T) j; |
479 |
val free' = app_bnds free (length Us); |
|
480 |
fun mk_abs_fun (T, (i, t)) = |
|
481 |
let val U = fastype_of t |
|
19494 | 482 |
in (i + 1, Const ("Nominal.abs_fun", [T, U, T] ---> |
483 |
Type ("Nominal.noption", [U])) $ mk_Free "y" T i $ t) |
|
17870 | 484 |
end |
485 |
in (j + 1, j' + length Ts, |
|
486 |
case dt'' of |
|
487 |
DtRec k => list_all (map (pair "x") Us, |
|
488 |
HOLogic.mk_Trueprop (HOLogic.mk_mem (free', |
|
489 |
Const (List.nth (rep_set_names, k), |
|
490 |
HOLogic.mk_setT T)))) :: prems |
|
491 |
| _ => prems, |
|
492 |
snd (foldr mk_abs_fun (j', free) Ts) :: ts) |
|
493 |
end; |
|
494 |
||
495 |
val (_, _, prems, ts) = foldr mk_prem (1, 1, [], []) cargs; |
|
496 |
val concl = HOLogic.mk_Trueprop (HOLogic.mk_mem |
|
497 |
(list_comb (Const (cname, map fastype_of ts ---> T), ts), |
|
498 |
Const (s, HOLogic.mk_setT T))) |
|
499 |
in Logic.list_implies (prems, concl) |
|
500 |
end; |
|
501 |
||
502 |
val (intr_ts, ind_consts) = |
|
503 |
apfst List.concat (ListPair.unzip (List.mapPartial |
|
19494 | 504 |
(fn ((_, ("Nominal.noption", _, _)), _) => NONE |
17870 | 505 |
| ((i, (_, _, constrs)), rep_set_name) => |
506 |
let val T = nth_dtyp i |
|
507 |
in SOME (map (make_intr rep_set_name T) constrs, |
|
508 |
Const (rep_set_name, HOLogic.mk_setT T)) |
|
509 |
end) |
|
510 |
(descr ~~ rep_set_names))); |
|
511 |
||
512 |
val (thy4, {raw_induct = rep_induct, intrs = rep_intrs, ...}) = |
|
513 |
setmp InductivePackage.quiet_mode false |
|
514 |
(InductivePackage.add_inductive_i false true big_rep_name false true false |
|
515 |
ind_consts (map (fn x => (("", x), [])) intr_ts) []) thy3; |
|
516 |
||
517 |
(**** Prove that representing set is closed under permutation ****) |
|
518 |
||
519 |
val _ = warning "proving closure under permutation..."; |
|
520 |
||
521 |
val perm_indnames' = List.mapPartial |
|
19494 | 522 |
(fn (x, (_, ("Nominal.noption", _, _))) => NONE | (x, _) => SOME x) |
17870 | 523 |
(perm_indnames ~~ descr); |
524 |
||
525 |
fun mk_perm_closed name = map (fn th => standard (th RS mp)) |
|
18010 | 526 |
(List.take (split_conj_thm (Goal.prove thy4 [] [] |
17870 | 527 |
(HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map |
528 |
(fn (S, x) => |
|
529 |
let |
|
530 |
val S = map_term_types (map_type_tfree |
|
531 |
(fn (s, cs) => TFree (s, cs union cp_classes))) S; |
|
532 |
val T = HOLogic.dest_setT (fastype_of S); |
|
533 |
val permT = mk_permT (Type (name, [])) |
|
534 |
in HOLogic.mk_imp (HOLogic.mk_mem (Free (x, T), S), |
|
19494 | 535 |
HOLogic.mk_mem (Const ("Nominal.perm", permT --> T --> T) $ |
17870 | 536 |
Free ("pi", permT) $ Free (x, T), S)) |
18010 | 537 |
end) (ind_consts ~~ perm_indnames')))) |
538 |
(fn _ => EVERY (* CU: added perm_fun_def in the final tactic in order to deal with funs *) |
|
17870 | 539 |
[indtac rep_induct [] 1, |
540 |
ALLGOALS (simp_tac (simpset_of thy4 addsimps |
|
541 |
(symmetric perm_fun_def :: PureThy.get_thms thy4 (Name ("abs_perm"))))), |
|
542 |
ALLGOALS (resolve_tac rep_intrs |
|
543 |
THEN_ALL_NEW (asm_full_simp_tac (simpset_of thy4 addsimps [perm_fun_def])))])), |
|
544 |
length new_type_names)); |
|
545 |
||
546 |
(* FIXME: theorems are stored in database for testing only *) |
|
547 |
val perm_closed_thmss = map mk_perm_closed atoms; |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
548 |
val (_,thy5) = PureThy.add_thmss [(("perm_closed", List.concat perm_closed_thmss), [])] thy4; |
17870 | 549 |
|
550 |
(**** typedef ****) |
|
551 |
||
552 |
val _ = warning "defining type..."; |
|
553 |
||
18366 | 554 |
val (typedefs, thy6) = |
555 |
fold_map (fn ((((name, mx), tvs), c), name') => fn thy => |
|
17870 | 556 |
setmp TypedefPackage.quiet_mode true |
557 |
(TypedefPackage.add_typedef_i false (SOME name') (name, tvs, mx) c NONE |
|
558 |
(rtac exI 1 THEN |
|
559 |
QUIET_BREADTH_FIRST (has_fewer_prems 1) |
|
19403 | 560 |
(resolve_tac rep_intrs 1))) thy |> (fn (r, thy) => |
17870 | 561 |
let |
562 |
val permT = mk_permT (TFree (variant tvs "'a", HOLogic.typeS)); |
|
563 |
val pi = Free ("pi", permT); |
|
564 |
val tvs' = map (fn s => TFree (s, the (AList.lookup op = sorts' s))) tvs; |
|
565 |
val T = Type (Sign.intern_type thy name, tvs'); |
|
566 |
val Const (_, Type (_, [U])) = c |
|
18366 | 567 |
in apfst (pair r o hd) |
19635 | 568 |
(PureThy.add_defs_unchecked_i true [(("prm_" ^ name ^ "_def", Logic.mk_equals |
19494 | 569 |
(Const ("Nominal.perm", permT --> T --> T) $ pi $ Free ("x", T), |
17870 | 570 |
Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $ |
19494 | 571 |
(Const ("Nominal.perm", permT --> U --> U) $ pi $ |
17870 | 572 |
(Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $ |
573 |
Free ("x", T))))), [])] thy) |
|
574 |
end)) |
|
18366 | 575 |
(types_syntax ~~ tyvars ~~ |
576 |
(List.take (ind_consts, length new_type_names)) ~~ new_type_names) thy5; |
|
17870 | 577 |
|
578 |
val perm_defs = map snd typedefs; |
|
579 |
val Abs_inverse_thms = map (#Abs_inverse o fst) typedefs; |
|
18016 | 580 |
val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs; |
17870 | 581 |
val Rep_thms = map (#Rep o fst) typedefs; |
582 |
||
18016 | 583 |
val big_name = space_implode "_" new_type_names; |
584 |
||
585 |
||
17870 | 586 |
(** prove that new types are in class pt_<name> **) |
587 |
||
588 |
val _ = warning "prove that new types are in class pt_<name> ..."; |
|
589 |
||
590 |
fun pt_instance ((class, atom), perm_closed_thms) = |
|
591 |
fold (fn (((({Abs_inverse, Rep_inverse, Rep, ...}, |
|
592 |
perm_def), name), tvs), perm_closed) => fn thy => |
|
19275 | 593 |
AxClass.prove_arity |
17870 | 594 |
(Sign.intern_type thy name, |
595 |
replicate (length tvs) (classes @ cp_classes), [class]) |
|
19133
7e84a1a3741c
Adapted to Florian's recent changes to the AxClass package.
berghofe
parents:
18759
diff
changeset
|
596 |
(EVERY [ClassPackage.intro_classes_tac [], |
17870 | 597 |
rewrite_goals_tac [perm_def], |
598 |
asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1, |
|
599 |
asm_full_simp_tac (simpset_of thy addsimps |
|
600 |
[Rep RS perm_closed RS Abs_inverse]) 1, |
|
601 |
asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy |
|
602 |
(Name ("pt_" ^ Sign.base_name atom ^ "3"))]) 1]) thy) |
|
603 |
(typedefs ~~ new_type_names ~~ tyvars ~~ perm_closed_thms); |
|
604 |
||
605 |
||
606 |
(** prove that new types are in class cp_<name1>_<name2> **) |
|
607 |
||
608 |
val _ = warning "prove that new types are in class cp_<name1>_<name2> ..."; |
|
609 |
||
610 |
fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy = |
|
611 |
let |
|
612 |
val name = "cp_" ^ Sign.base_name atom1 ^ "_" ^ Sign.base_name atom2; |
|
613 |
val class = Sign.intern_class thy name; |
|
614 |
val cp1' = PureThy.get_thm thy (Name (name ^ "_inst")) RS cp1 |
|
615 |
in fold (fn ((((({Abs_inverse, Rep_inverse, Rep, ...}, |
|
616 |
perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy => |
|
19275 | 617 |
AxClass.prove_arity |
17870 | 618 |
(Sign.intern_type thy name, |
619 |
replicate (length tvs) (classes @ cp_classes), [class]) |
|
19133
7e84a1a3741c
Adapted to Florian's recent changes to the AxClass package.
berghofe
parents:
18759
diff
changeset
|
620 |
(EVERY [ClassPackage.intro_classes_tac [], |
17870 | 621 |
rewrite_goals_tac [perm_def], |
622 |
asm_full_simp_tac (simpset_of thy addsimps |
|
623 |
((Rep RS perm_closed1 RS Abs_inverse) :: |
|
624 |
(if atom1 = atom2 then [] |
|
625 |
else [Rep RS perm_closed2 RS Abs_inverse]))) 1, |
|
18016 | 626 |
cong_tac 1, |
17870 | 627 |
rtac refl 1, |
628 |
rtac cp1' 1]) thy) |
|
629 |
(typedefs ~~ new_type_names ~~ tyvars ~~ perm_closed_thms1 ~~ |
|
630 |
perm_closed_thms2) thy |
|
631 |
end; |
|
632 |
||
633 |
val thy7 = fold (fn x => fn thy => thy |> |
|
634 |
pt_instance x |> |
|
635 |
fold (cp_instance (apfst snd x)) (atoms ~~ perm_closed_thmss)) |
|
636 |
(classes ~~ atoms ~~ perm_closed_thmss) thy6; |
|
637 |
||
638 |
(**** constructors ****) |
|
639 |
||
640 |
fun mk_abs_fun (x, t) = |
|
641 |
let |
|
642 |
val T = fastype_of x; |
|
643 |
val U = fastype_of t |
|
644 |
in |
|
19494 | 645 |
Const ("Nominal.abs_fun", T --> U --> T --> |
646 |
Type ("Nominal.noption", [U])) $ x $ t |
|
17870 | 647 |
end; |
648 |
||
18016 | 649 |
val (ty_idxs, _) = foldl |
19494 | 650 |
(fn ((i, ("Nominal.noption", _, _)), p) => p |
18016 | 651 |
| ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr; |
652 |
||
653 |
fun reindex (DtType (s, dts)) = DtType (s, map reindex dts) |
|
654 |
| reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i)) |
|
655 |
| reindex dt = dt; |
|
656 |
||
657 |
fun strip_suffix i s = implode (List.take (explode s, size s - i)); |
|
658 |
||
659 |
(** strips the "_Rep" in type names *) |
|
18045 | 660 |
fun strip_nth_name i s = |
661 |
let val xs = NameSpace.unpack s; |
|
662 |
in NameSpace.pack (Library.nth_map (length xs - i) (strip_suffix 4) xs) end; |
|
18016 | 663 |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
664 |
val (descr'', ndescr) = ListPair.unzip (List.mapPartial |
19494 | 665 |
(fn (i, ("Nominal.noption", _, _)) => NONE |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
666 |
| (i, (s, dts, constrs)) => |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
667 |
let |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
668 |
val SOME index = AList.lookup op = ty_idxs i; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
669 |
val (constrs1, constrs2) = ListPair.unzip |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
670 |
(map (fn (cname, cargs) => apfst (pair (strip_nth_name 2 cname)) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
671 |
(foldl_map (fn (dts, dt) => |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
672 |
let val (dts', dt') = strip_option dt |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
673 |
in (dts @ dts' @ [reindex dt'], (length dts, length dts')) end) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
674 |
([], cargs))) constrs) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
675 |
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
|
676 |
(index, constrs2)) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
677 |
end) descr); |
18045 | 678 |
|
19489 | 679 |
val (descr1, descr2) = chop (length new_type_names) descr''; |
18016 | 680 |
val descr' = [descr1, descr2]; |
681 |
||
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
682 |
val typ_of_dtyp' = replace_types' o typ_of_dtyp descr sorts'; |
17870 | 683 |
|
684 |
val rep_names = map (fn s => |
|
685 |
Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names; |
|
686 |
val abs_names = map (fn s => |
|
687 |
Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names; |
|
688 |
||
18016 | 689 |
val recTs' = List.mapPartial |
19494 | 690 |
(fn ((_, ("Nominal.noption", _, _)), T) => NONE |
18016 | 691 |
| (_, T) => SOME T) (descr ~~ get_rec_types descr sorts'); |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
692 |
val recTs = get_rec_types descr'' sorts'; |
18016 | 693 |
val newTs' = Library.take (length new_type_names, recTs'); |
694 |
val newTs = Library.take (length new_type_names, recTs); |
|
17870 | 695 |
|
696 |
val full_new_type_names = map (Sign.full_name (sign_of thy)) new_type_names; |
|
697 |
||
698 |
fun make_constr_def tname T T' ((thy, defs, eqns), ((cname, cargs), (cname', mx))) = |
|
699 |
let |
|
700 |
fun constr_arg (dt, (j, l_args, r_args)) = |
|
701 |
let |
|
702 |
val x' = mk_Free "x" (typ_of_dtyp' dt) j; |
|
703 |
val (dts, dt') = strip_option dt; |
|
704 |
val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp' dt) i) |
|
705 |
(dts ~~ (j upto j + length dts - 1)) |
|
706 |
val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts) |
|
707 |
val (dts', dt'') = strip_dtyp dt' |
|
18261
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
708 |
in |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
709 |
(j + length dts + 1, |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
710 |
xs @ x :: l_args, |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
711 |
foldr mk_abs_fun |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
712 |
(case dt'' of |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
713 |
DtRec k => if k < length new_type_names then |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
714 |
list_abs (map (pair "z" o typ_of_dtyp') dts', |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
715 |
Const (List.nth (rep_names, k), typ_of_dtyp' dt'' --> |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
716 |
typ_of_dtyp descr sorts' dt'') $ app_bnds x (length dts')) |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
717 |
else error "nested recursion not (yet) supported" |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
718 |
| _ => x) xs :: r_args) |
17870 | 719 |
end |
720 |
||
721 |
val (_, l_args, r_args) = foldr constr_arg (1, [], []) cargs; |
|
722 |
val abs_name = Sign.intern_const (Theory.sign_of thy) ("Abs_" ^ tname); |
|
723 |
val rep_name = Sign.intern_const (Theory.sign_of thy) ("Rep_" ^ tname); |
|
724 |
val constrT = map fastype_of l_args ---> T; |
|
725 |
val lhs = list_comb (Const (Sign.full_name thy (Sign.base_name cname), |
|
726 |
constrT), l_args); |
|
727 |
val rhs = list_comb (Const (cname, map fastype_of r_args ---> T'), r_args); |
|
728 |
val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs); |
|
729 |
val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
730 |
(Const (rep_name, T --> T') $ lhs, rhs)); |
|
731 |
val def_name = (Sign.base_name cname) ^ "_def"; |
|
18366 | 732 |
val ([def_thm], thy') = thy |> |
17870 | 733 |
Theory.add_consts_i [(cname', constrT, mx)] |> |
734 |
(PureThy.add_defs_i false o map Thm.no_attributes) [(def_name, def)] |
|
735 |
in (thy', defs @ [def_thm], eqns @ [eqn]) end; |
|
736 |
||
737 |
fun dt_constr_defs ((thy, defs, eqns, dist_lemmas), |
|
738 |
(((((_, (_, _, constrs)), tname), T), T'), constr_syntax)) = |
|
739 |
let |
|
740 |
val rep_const = cterm_of thy |
|
741 |
(Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T')); |
|
742 |
val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma); |
|
743 |
val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T') |
|
744 |
((Theory.add_path tname thy, defs, []), constrs ~~ constr_syntax) |
|
745 |
in |
|
746 |
(parent_path flat_names thy', defs', eqns @ [eqns'], dist_lemmas @ [dist]) |
|
747 |
end; |
|
748 |
||
749 |
val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs |
|
750 |
((thy7, [], [], []), List.take (descr, length new_type_names) ~~ |
|
751 |
new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax); |
|
752 |
||
753 |
val abs_inject_thms = map (fn tname => |
|
754 |
PureThy.get_thm thy8 (Name ("Abs_" ^ tname ^ "_inject"))) new_type_names; |
|
755 |
||
756 |
val rep_inject_thms = map (fn tname => |
|
757 |
PureThy.get_thm thy8 (Name ("Rep_" ^ tname ^ "_inject"))) new_type_names; |
|
758 |
||
759 |
val rep_thms = map (fn tname => |
|
760 |
PureThy.get_thm thy8 (Name ("Rep_" ^ tname))) new_type_names; |
|
761 |
||
762 |
val rep_inverse_thms = map (fn tname => |
|
763 |
PureThy.get_thm thy8 (Name ("Rep_" ^ tname ^ "_inverse"))) new_type_names; |
|
764 |
||
765 |
(* prove theorem Rep_i (Constr_j ...) = Constr'_j ... *) |
|
766 |
||
767 |
fun prove_constr_rep_thm eqn = |
|
768 |
let |
|
769 |
val inj_thms = map (fn r => r RS iffD1) abs_inject_thms; |
|
770 |
val rewrites = constr_defs @ map mk_meta_eq rep_inverse_thms |
|
18010 | 771 |
in standard (Goal.prove thy8 [] [] eqn (fn _ => EVERY |
17870 | 772 |
[resolve_tac inj_thms 1, |
773 |
rewrite_goals_tac rewrites, |
|
774 |
rtac refl 3, |
|
775 |
resolve_tac rep_intrs 2, |
|
18010 | 776 |
REPEAT (resolve_tac rep_thms 1)])) |
17870 | 777 |
end; |
778 |
||
779 |
val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns; |
|
780 |
||
781 |
(* prove theorem pi \<bullet> Rep_i x = Rep_i (pi \<bullet> x) *) |
|
782 |
||
783 |
fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th => |
|
784 |
let |
|
785 |
val _ $ (_ $ (Rep $ x) $ _) = Logic.unvarify (prop_of th); |
|
786 |
val Type ("fun", [T, U]) = fastype_of Rep; |
|
787 |
val permT = mk_permT (Type (atom, [])); |
|
788 |
val pi = Free ("pi", permT); |
|
789 |
in |
|
18010 | 790 |
standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq |
19494 | 791 |
(Const ("Nominal.perm", permT --> U --> U) $ pi $ (Rep $ x), |
792 |
Rep $ (Const ("Nominal.perm", permT --> T --> T) $ pi $ x)))) |
|
18010 | 793 |
(fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @ |
794 |
perm_closed_thms @ Rep_thms)) 1)) |
|
17870 | 795 |
end) Rep_thms; |
796 |
||
797 |
val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm |
|
798 |
(atoms ~~ perm_closed_thmss)); |
|
799 |
||
800 |
(* prove distinctness theorems *) |
|
801 |
||
18016 | 802 |
val distinct_props = setmp DatatypeProp.dtK 1000 |
803 |
(DatatypeProp.make_distincts new_type_names descr' sorts') thy8; |
|
17870 | 804 |
|
805 |
val dist_rewrites = map (fn (rep_thms, dist_lemma) => |
|
806 |
dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])) |
|
807 |
(constr_rep_thmss ~~ dist_lemmas); |
|
808 |
||
809 |
fun prove_distinct_thms (_, []) = [] |
|
810 |
| prove_distinct_thms (p as (rep_thms, dist_lemma), t::ts) = |
|
811 |
let |
|
18010 | 812 |
val dist_thm = standard (Goal.prove thy8 [] [] t (fn _ => |
813 |
simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1)) |
|
17870 | 814 |
in dist_thm::(standard (dist_thm RS not_sym)):: |
815 |
(prove_distinct_thms (p, ts)) |
|
816 |
end; |
|
817 |
||
818 |
val distinct_thms = map prove_distinct_thms |
|
819 |
(constr_rep_thmss ~~ dist_lemmas ~~ distinct_props); |
|
820 |
||
821 |
(** prove equations for permutation functions **) |
|
822 |
||
823 |
val abs_perm = PureThy.get_thms thy8 (Name "abs_perm"); (* FIXME *) |
|
824 |
||
825 |
val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => |
|
826 |
let val T = replace_types' (nth_dtyp i) |
|
827 |
in List.concat (map (fn (atom, perm_closed_thms) => |
|
828 |
map (fn ((cname, dts), constr_rep_thm) => |
|
829 |
let |
|
830 |
val cname = Sign.intern_const thy8 |
|
831 |
(NameSpace.append tname (Sign.base_name cname)); |
|
832 |
val permT = mk_permT (Type (atom, [])); |
|
833 |
val pi = Free ("pi", permT); |
|
834 |
||
835 |
fun perm t = |
|
836 |
let val T = fastype_of t |
|
19494 | 837 |
in Const ("Nominal.perm", permT --> T --> T) $ pi $ t end; |
17870 | 838 |
|
839 |
fun constr_arg (dt, (j, l_args, r_args)) = |
|
840 |
let |
|
841 |
val x' = mk_Free "x" (typ_of_dtyp' dt) j; |
|
842 |
val (dts, dt') = strip_option dt; |
|
843 |
val Ts = map typ_of_dtyp' dts; |
|
844 |
val xs = map (fn (T, i) => mk_Free "x" T i) |
|
845 |
(Ts ~~ (j upto j + length dts - 1)) |
|
846 |
val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts); |
|
847 |
val (dts', dt'') = strip_dtyp dt'; |
|
18261
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
848 |
in |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
849 |
(j + length dts + 1, |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
850 |
xs @ x :: l_args, |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
851 |
map perm (xs @ [x]) @ r_args) |
17870 | 852 |
end |
853 |
||
854 |
val (_, l_args, r_args) = foldr constr_arg (1, [], []) dts; |
|
855 |
val c = Const (cname, map fastype_of l_args ---> T) |
|
856 |
in |
|
18010 | 857 |
standard (Goal.prove thy8 [] [] |
17870 | 858 |
(HOLogic.mk_Trueprop (HOLogic.mk_eq |
18010 | 859 |
(perm (list_comb (c, l_args)), list_comb (c, r_args)))) |
860 |
(fn _ => EVERY |
|
17870 | 861 |
[simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1, |
862 |
simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @ |
|
863 |
constr_defs @ perm_closed_thms)) 1, |
|
864 |
TRY (simp_tac (HOL_basic_ss addsimps |
|
865 |
(symmetric perm_fun_def :: abs_perm)) 1), |
|
866 |
TRY (simp_tac (HOL_basic_ss addsimps |
|
867 |
(perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @ |
|
18010 | 868 |
perm_closed_thms)) 1)])) |
17870 | 869 |
end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss)) |
870 |
end) (List.take (descr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); |
|
871 |
||
872 |
(** prove injectivity of constructors **) |
|
873 |
||
874 |
val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms; |
|
875 |
val alpha = PureThy.get_thms thy8 (Name "alpha"); |
|
876 |
val abs_fresh = PureThy.get_thms thy8 (Name "abs_fresh"); |
|
877 |
val fresh_def = PureThy.get_thm thy8 (Name "fresh_def"); |
|
878 |
val supp_def = PureThy.get_thm thy8 (Name "supp_def"); |
|
879 |
||
880 |
val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => |
|
881 |
let val T = replace_types' (nth_dtyp i) |
|
882 |
in List.mapPartial (fn ((cname, dts), constr_rep_thm) => |
|
883 |
if null dts then NONE else SOME |
|
884 |
let |
|
885 |
val cname = Sign.intern_const thy8 |
|
886 |
(NameSpace.append tname (Sign.base_name cname)); |
|
887 |
||
888 |
fun make_inj (dt, (j, args1, args2, eqs)) = |
|
889 |
let |
|
890 |
val x' = mk_Free "x" (typ_of_dtyp' dt) j; |
|
891 |
val y' = mk_Free "y" (typ_of_dtyp' dt) j; |
|
892 |
val (dts, dt') = strip_option dt; |
|
893 |
val Ts_idx = map typ_of_dtyp' dts ~~ (j upto j + length dts - 1); |
|
894 |
val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx; |
|
895 |
val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx; |
|
896 |
val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts); |
|
897 |
val y = mk_Free "y" (typ_of_dtyp' dt') (j + length dts); |
|
898 |
val (dts', dt'') = strip_dtyp dt'; |
|
18261
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
899 |
in |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
900 |
(j + length dts + 1, |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
901 |
xs @ (x :: args1), ys @ (y :: args2), |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
902 |
HOLogic.mk_eq |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
903 |
(foldr mk_abs_fun x xs, foldr mk_abs_fun y ys) :: eqs) |
17870 | 904 |
end; |
905 |
||
906 |
val (_, args1, args2, eqs) = foldr make_inj (1, [], [], []) dts; |
|
907 |
val Ts = map fastype_of args1; |
|
908 |
val c = Const (cname, Ts ---> T) |
|
909 |
in |
|
18010 | 910 |
standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq |
17870 | 911 |
(HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)), |
18010 | 912 |
foldr1 HOLogic.mk_conj eqs))) |
913 |
(fn _ => EVERY |
|
17870 | 914 |
[asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: |
915 |
rep_inject_thms')) 1, |
|
916 |
TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def :: |
|
917 |
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
|
918 |
perm_rep_perm_thms)) 1), |
8be65cf94d2e
Improved proof of injectivity theorems to make it work on
berghofe
parents:
17873
diff
changeset
|
919 |
TRY (asm_full_simp_tac (HOL_basic_ss addsimps (perm_fun_def :: |
18010 | 920 |
expand_fun_eq :: rep_inject_thms @ perm_rep_perm_thms)) 1)])) |
17870 | 921 |
end) (constrs ~~ constr_rep_thms) |
922 |
end) (List.take (descr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); |
|
923 |
||
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
924 |
(** equations for support and freshness **) |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
925 |
|
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
926 |
val Un_assoc = PureThy.get_thm thy8 (Name "Un_assoc"); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
927 |
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
|
928 |
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
|
929 |
val finite_Un = PureThy.get_thm thy8 (Name "finite_Un"); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
930 |
|
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
931 |
val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
932 |
(map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') => |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
933 |
let val T = replace_types' (nth_dtyp i) |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
934 |
in List.concat (map (fn (cname, dts) => map (fn atom => |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
935 |
let |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
936 |
val cname = Sign.intern_const thy8 |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
937 |
(NameSpace.append tname (Sign.base_name cname)); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
938 |
val atomT = Type (atom, []); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
939 |
|
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
940 |
fun process_constr (dt, (j, args1, args2)) = |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
941 |
let |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
942 |
val x' = mk_Free "x" (typ_of_dtyp' dt) j; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
943 |
val (dts, dt') = strip_option dt; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
944 |
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
|
945 |
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
|
946 |
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
|
947 |
val (dts', dt'') = strip_dtyp dt'; |
18261
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
948 |
in |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
949 |
(j + length dts + 1, |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
950 |
xs @ (x :: args1), foldr mk_abs_fun x xs :: args2) |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
951 |
end; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
952 |
|
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
953 |
val (_, args1, args2) = foldr process_constr (1, [], []) dts; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
954 |
val Ts = map fastype_of args1; |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
955 |
val c = list_comb (Const (cname, Ts ---> T), args1); |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
956 |
fun supp t = |
19494 | 957 |
Const ("Nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t; |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
958 |
fun fresh t = |
19494 | 959 |
Const ("Nominal.fresh", atomT --> fastype_of t --> HOLogic.boolT) $ |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
960 |
Free ("a", atomT) $ t; |
18010 | 961 |
val supp_thm = standard (Goal.prove thy8 [] [] |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
962 |
(HOLogic.mk_Trueprop (HOLogic.mk_eq |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
963 |
(supp c, |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
964 |
if null dts then Const ("{}", HOLogic.mk_setT atomT) |
18010 | 965 |
else foldr1 (HOLogic.mk_binop "op Un") (map supp args2)))) |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
966 |
(fn _ => |
18010 | 967 |
simp_tac (HOL_basic_ss addsimps (supp_def :: |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
968 |
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
|
969 |
symmetric empty_def :: Finites.emptyI :: simp_thms @ |
18010 | 970 |
abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1)) |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
971 |
in |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
972 |
(supp_thm, |
18010 | 973 |
standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
974 |
(fresh c, |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
975 |
if null dts then HOLogic.true_const |
18010 | 976 |
else foldr1 HOLogic.mk_conj (map fresh args2)))) |
17872
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
977 |
(fn _ => |
18010 | 978 |
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
|
979 |
end) atoms) constrs) |
f08fc98a164a
Implemented proofs for support and freshness theorems.
berghofe
parents:
17870
diff
changeset
|
980 |
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
|
981 |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
982 |
(**** weak induction theorem ****) |
18016 | 983 |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
984 |
val arities = get_arities descr''; |
18016 | 985 |
|
986 |
fun mk_funs_inv thm = |
|
987 |
let |
|
988 |
val {sign, prop, ...} = rep_thm thm; |
|
989 |
val _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ S)) $ |
|
990 |
(_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop; |
|
991 |
val used = add_term_tfree_names (a, []); |
|
992 |
||
993 |
fun mk_thm i = |
|
994 |
let |
|
995 |
val Ts = map (TFree o rpair HOLogic.typeS) |
|
996 |
(variantlist (replicate i "'t", used)); |
|
997 |
val f = Free ("f", Ts ---> U) |
|
998 |
in standard (Goal.prove sign [] [] (Logic.mk_implies |
|
999 |
(HOLogic.mk_Trueprop (HOLogic.list_all |
|
1000 |
(map (pair "x") Ts, HOLogic.mk_mem (app_bnds f i, S))), |
|
1001 |
HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts, |
|
1002 |
r $ (a $ app_bnds f i)), f)))) |
|
1003 |
(fn _ => EVERY [REPEAT (rtac ext 1), REPEAT (etac allE 1), rtac thm 1, atac 1])) |
|
1004 |
end |
|
1005 |
in map (fn r => r RS subst) (thm :: map mk_thm arities) end; |
|
1006 |
||
1007 |
fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) = |
|
1008 |
let |
|
1009 |
val Rep_t = Const (List.nth (rep_names, i), T --> U) $ |
|
1010 |
mk_Free "x" T i; |
|
1011 |
||
1012 |
val Abs_t = Const (List.nth (abs_names, i), U --> T) |
|
1013 |
||
1014 |
in (prems @ [HOLogic.imp $ HOLogic.mk_mem (Rep_t, |
|
1015 |
Const (List.nth (rep_set_names, i), HOLogic.mk_setT U)) $ |
|
1016 |
(mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))], |
|
1017 |
concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i]) |
|
1018 |
end; |
|
1019 |
||
1020 |
val (indrule_lemma_prems, indrule_lemma_concls) = |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1021 |
Library.foldl mk_indrule_lemma (([], []), (descr'' ~~ recTs ~~ recTs')); |
18016 | 1022 |
|
1023 |
val indrule_lemma = standard (Goal.prove thy8 [] [] |
|
1024 |
(Logic.mk_implies |
|
1025 |
(HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems), |
|
1026 |
HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY |
|
1027 |
[REPEAT (etac conjE 1), |
|
1028 |
REPEAT (EVERY |
|
1029 |
[TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1, |
|
1030 |
etac mp 1, resolve_tac Rep_thms 1])])); |
|
1031 |
||
1032 |
val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma))); |
|
1033 |
val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else |
|
1034 |
map (Free o apfst fst o dest_Var) Ps; |
|
1035 |
val indrule_lemma' = cterm_instantiate |
|
1036 |
(map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma; |
|
1037 |
||
1038 |
val Abs_inverse_thms' = List.concat (map mk_funs_inv Abs_inverse_thms); |
|
1039 |
||
1040 |
val dt_induct_prop = DatatypeProp.make_ind descr' sorts'; |
|
1041 |
val dt_induct = standard (Goal.prove thy8 [] |
|
1042 |
(Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) |
|
1043 |
(fn prems => EVERY |
|
1044 |
[rtac indrule_lemma' 1, |
|
1045 |
(DatatypeAux.indtac rep_induct THEN_ALL_NEW ObjectLogic.atomize_tac) 1, |
|
1046 |
EVERY (map (fn (prem, r) => (EVERY |
|
1047 |
[REPEAT (eresolve_tac Abs_inverse_thms' 1), |
|
1048 |
simp_tac (HOL_basic_ss addsimps [symmetric r]) 1, |
|
1049 |
DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)])) |
|
1050 |
(prems ~~ constr_defs))])); |
|
1051 |
||
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1052 |
val case_names_induct = mk_case_names_induct descr''; |
18016 | 1053 |
|
18066
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1054 |
(**** prove that new datatypes have finite support ****) |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1055 |
|
18246
676d2e625d98
added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
18245
diff
changeset
|
1056 |
val _ = warning "proving finite support for the new datatype"; |
676d2e625d98
added fsub.thy (poplmark challenge) to the examples
urbanc
parents:
18245
diff
changeset
|
1057 |
|
18066
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1058 |
val indnames = DatatypeProp.make_tnames recTs; |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1059 |
|
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1060 |
val abs_supp = PureThy.get_thms thy8 (Name "abs_supp"); |
18067 | 1061 |
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
|
1062 |
|
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1063 |
val finite_supp_thms = map (fn atom => |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1064 |
let val atomT = Type (atom, []) |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1065 |
in map standard (List.take |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1066 |
(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
|
1067 |
(foldr1 HOLogic.mk_conj (map (fn (s, T) => HOLogic.mk_mem |
19494 | 1068 |
(Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T), |
18066
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1069 |
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
|
1070 |
(indnames ~~ recTs)))) |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1071 |
(fn _ => indtac dt_induct indnames 1 THEN |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1072 |
ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps |
18067 | 1073 |
(abs_supp @ supp_atm @ |
18066
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1074 |
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
|
1075 |
List.concat supp_thms))))), |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1076 |
length new_type_names)) |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1077 |
end) atoms; |
d1e47ee13070
Added code for proving that new datatype has finite support.
berghofe
parents:
18054
diff
changeset
|
1078 |
|
18759 | 1079 |
val simp_atts = replicate (length new_type_names) [Simplifier.simp_add]; |
18658 | 1080 |
|
1081 |
val (_, thy9) = thy8 |> |
|
1082 |
Theory.add_path big_name |> |
|
1083 |
PureThy.add_thms [(("induct_weak", dt_induct), [case_names_induct])] ||>> |
|
1084 |
PureThy.add_thmss [(("inducts_weak", projections dt_induct), [case_names_induct])] ||> |
|
1085 |
Theory.parent_path ||>> |
|
1086 |
DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>> |
|
1087 |
DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>> |
|
1088 |
DatatypeAux.store_thmss_atts "perm" new_type_names simp_atts perm_simps' ||>> |
|
1089 |
DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>> |
|
1090 |
DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>> |
|
1091 |
DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||> |
|
1092 |
fold (fn (atom, ths) => fn thy => |
|
1093 |
let val class = Sign.intern_class thy ("fs_" ^ Sign.base_name atom) |
|
19275 | 1094 |
in fold (fn T => AxClass.prove_arity |
18658 | 1095 |
(fst (dest_Type T), |
1096 |
replicate (length sorts) [class], [class]) |
|
19133
7e84a1a3741c
Adapted to Florian's recent changes to the AxClass package.
berghofe
parents:
18759
diff
changeset
|
1097 |
(ClassPackage.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy |
18658 | 1098 |
end) (atoms ~~ finite_supp_thms); |
1099 |
||
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1100 |
(**** strong induction theorem ****) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1101 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1102 |
val pnames = if length descr'' = 1 then ["P"] |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1103 |
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
|
1104 |
val ind_sort = if null dt_atomTs then HOLogic.typeS |
18658 | 1105 |
else norm_sort thy9 (map (fn T => Sign.intern_class thy9 ("fs_" ^ |
1106 |
Sign.base_name (fst (dest_Type T)))) dt_atomTs); |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1107 |
val fsT = TFree ("'n", ind_sort); |
18658 | 1108 |
val fsT' = TFree ("'n", HOLogic.typeS); |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1109 |
|
18658 | 1110 |
val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T))) |
1111 |
(DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs); |
|
1112 |
||
1113 |
fun make_pred fsT i T = |
|
18302
577e5d19b33c
Changed order of predicate arguments and quantifiers in strong induction rule.
berghofe
parents:
18280
diff
changeset
|
1114 |
Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT); |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1115 |
|
18658 | 1116 |
fun make_ind_prem fsT f k T ((cname, cargs), idxs) = |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1117 |
let |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1118 |
val recs = List.filter is_rec_type cargs; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1119 |
val Ts = map (typ_of_dtyp descr'' sorts') cargs; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1120 |
val recTs' = map (typ_of_dtyp descr'' sorts') recs; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1121 |
val tnames = variantlist (DatatypeProp.make_tnames Ts, pnames); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1122 |
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
|
1123 |
val frees = tnames ~~ Ts; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1124 |
val z = (variant tnames "z", fsT); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1125 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1126 |
fun mk_prem ((dt, s), T) = |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1127 |
let |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1128 |
val (Us, U) = strip_type T; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1129 |
val l = length Us |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1130 |
in list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop |
18658 | 1131 |
(make_pred fsT (body_index dt) U $ Bound l $ app_bnds (Free (s, T)) l)) |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1132 |
end; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1133 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1134 |
val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs'); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1135 |
val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop |
18658 | 1136 |
(f T (Free p) (Free z))) |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1137 |
(map (curry List.nth frees) (List.concat (map (fn (m, n) => |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1138 |
m upto m + n - 1) idxs))) |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1139 |
|
18302
577e5d19b33c
Changed order of predicate arguments and quantifiers in strong induction rule.
berghofe
parents:
18280
diff
changeset
|
1140 |
in list_all_free (frees @ [z], Logic.list_implies (prems' @ prems, |
18658 | 1141 |
HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $ |
18302
577e5d19b33c
Changed order of predicate arguments and quantifiers in strong induction rule.
berghofe
parents:
18280
diff
changeset
|
1142 |
list_comb (Const (cname, Ts ---> T), map Free frees)))) |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1143 |
end; |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1144 |
|
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1145 |
val ind_prems = List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) => |
18658 | 1146 |
map (make_ind_prem fsT (fn T => fn t => fn u => |
19494 | 1147 |
Const ("Nominal.fresh", T --> fsT --> HOLogic.boolT) $ t $ u) i T) |
18658 | 1148 |
(constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs)); |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1149 |
val tnames = DatatypeProp.make_tnames recTs; |
18658 | 1150 |
val zs = variantlist (replicate (length descr'') "z", tnames); |
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1151 |
val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") |
18658 | 1152 |
(map (fn ((((i, _), T), tname), z) => |
1153 |
make_pred fsT i T $ Free (z, fsT) $ Free (tname, T)) |
|
1154 |
(descr'' ~~ recTs ~~ tnames ~~ zs))); |
|
18107
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1155 |
val induct = Logic.list_implies (ind_prems, ind_concl); |
ee6b4d3af498
Added strong induction theorem (currently only axiomatized!).
berghofe
parents:
18104
diff
changeset
|
1156 |
|
18658 | 1157 |
val ind_prems' = |
1158 |
map (fn (_, f as Free (_, T)) => list_all_free ([("x", fsT')], |
|
1159 |
HOLogic.mk_Trueprop (HOLogic.mk_mem (f $ Free ("x", fsT'), |
|
1160 |
Const ("Finite_Set.Finites", HOLogic.mk_setT (body_type T)))))) fresh_fs @ |
|
1161 |
List.concat (map (fn (((i, (_, _, constrs)), (_, idxss)), T) => |
|
1162 |
map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $ |
|
1163 |
HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T) |
|
1164 |
(constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs)); |
|
1165 |
val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") |
|
1166 |
(map (fn ((((i, _), T), tname), z) => |
|
1167 |
make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T)) |
|
1168 |
(descr'' ~~ recTs ~~ tnames ~~ zs))); |
|
1169 |
val induct' = Logic.list_implies (ind_prems', ind_concl'); |
|
1170 |
||
1171 |
fun mk_perm Ts (t, u) = |
|
1172 |
let |
|
1173 |
val T = fastype_of1 (Ts, t); |
|
1174 |
val U = fastype_of1 (Ts, u) |
|
19494 | 1175 |
in Const ("Nominal.perm", T --> U --> U) $ t $ u end; |
18658 | 1176 |
|
1177 |
val aux_ind_vars = |
|
1178 |
(DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~ |
|
1179 |
map mk_permT dt_atomTs) @ [("z", fsT')]; |
|
1180 |
val aux_ind_Ts = rev (map snd aux_ind_vars); |
|
1181 |
val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") |
|
1182 |
(map (fn (((i, _), T), tname) => |
|
1183 |
HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $ |
|
1184 |
foldr (mk_perm aux_ind_Ts) (Free (tname, T)) |
|
1185 |
(map Bound (length dt_atomTs downto 1)))) |
|
1186 |
(descr'' ~~ recTs ~~ tnames))); |
|
1187 |
||
1188 |
fun mk_ind_perm i k p l vs j = |
|
1189 |
let |
|
1190 |
val n = length vs; |
|
1191 |
val Ts = map snd vs; |
|
1192 |
val T = List.nth (Ts, i - j); |
|
1193 |
val pT = NominalAtoms.mk_permT T |
|
1194 |
in |
|
1195 |
Const ("List.list.Cons", HOLogic.mk_prodT (T, T) --> pT --> pT) $ |
|
1196 |
(HOLogic.pair_const T T $ Bound (l - j) $ foldr (mk_perm Ts) |
|
1197 |
(Bound (i - j)) |
|
1198 |
(map (mk_ind_perm i k p l vs) (j - 1 downto 0) @ |
|
1199 |
map Bound (n - k - 1 downto n - k - p))) $ |
|
1200 |
Const ("List.list.Nil", pT) |
|
1201 |
end; |
|
1202 |
||
1203 |
fun mk_fresh i i' j k p l vs _ _ = |
|
1204 |
let |
|
1205 |
val n = length vs; |
|
1206 |
val Ts = map snd vs; |
|
1207 |
val T = List.nth (Ts, n - i - 1 - j); |
|
1208 |
val f = the (AList.lookup op = fresh_fs T); |
|
1209 |
val U = List.nth (Ts, n - i' - 1); |
|
1210 |
val S = HOLogic.mk_setT T; |
|
1211 |
val prms = map (mk_ind_perm (n - i) k p (n - l) (("a", T) :: vs)) |
|
1212 |
(j - 1 downto 0) @ |
|
1213 |
map Bound (n - k downto n - k - p + 1) |
|
1214 |
in |
|
1215 |
HOLogic.mk_Trueprop (Const ("Ex", (T --> HOLogic.boolT) --> HOLogic.boolT) $ |
|
1216 |
Abs ("a", T, HOLogic.Not $ |
|
1217 |
(Const ("op :", T --> S --> HOLogic.boolT) $ Bound 0 $ |
|
1218 |
(Const ("insert", T --> S --> S) $ |
|
1219 |
(foldr (mk_perm (T :: Ts)) (Bound (n - i - j)) prms) $ |
|
1220 |
(Const ("op Un", S --> S --> S) $ (f $ Bound (n - k - p)) $ |
|
19494 | 1221 |
(Const ("Nominal.supp", U --> S) $ |
18658 | 1222 |
foldr (mk_perm (T :: Ts)) (Bound (n - i')) prms)))))) |
1223 |
end; |
|
18104
dbe58b104cb9
added thms perm, distinct and fresh to the simplifier.
urbanc
parents:
18068
diff
changeset
|
1224 |
|
18658 | 1225 |
fun mk_fresh_constr is p vs _ concl = |
1226 |
let |
|
1227 |
val n = length vs; |
|
1228 |
val Ts = map snd vs; |
|
1229 |
val _ $ (_ $ _ $ t) = concl; |
|
1230 |
val c = head_of t; |
|
1231 |
val T = body_type (fastype_of c); |
|
1232 |
val k = foldr op + 0 (map (fn (_, i) => i + 1) is); |
|
1233 |
val ps = map Bound (n - k - 1 downto n - k - p); |
|
1234 |
val (_, _, ts, us) = foldl (fn ((_, i), (m, n, ts, us)) => |
|
1235 |
(m - i - 1, n - i, |
|
1236 |
ts @ map Bound (n downto n - i + 1) @ |
|
1237 |
[foldr (mk_perm Ts) (Bound (m - i)) |
|
1238 |
(map (mk_ind_perm m k p n vs) (i - 1 downto 0) @ ps)], |
|
1239 |
us @ map (fn j => foldr (mk_perm Ts) (Bound j) ps) (m downto m - i))) |
|
1240 |
(n - 1, n - k - p - 2, [], []) is |
|
1241 |
in |
|
1242 |
HOLogic.mk_Trueprop (HOLogic.eq_const T $ list_comb (c, ts) $ list_comb (c, us)) |
|
1243 |
end; |
|
1244 |
||
1245 |
val abs_fun_finite_supp = PureThy.get_thm thy9 (Name "abs_fun_finite_supp"); |
|
1246 |
||
1247 |
val at_finite_select = PureThy.get_thm thy9 (Name "at_finite_select"); |
|
1248 |
||
1249 |
val induct_aux_lemmas = List.concat (map (fn Type (s, _) => |
|
1250 |
[PureThy.get_thm thy9 (Name ("pt_" ^ Sign.base_name s ^ "_inst")), |
|
1251 |
PureThy.get_thm thy9 (Name ("fs_" ^ Sign.base_name s ^ "1")), |
|
1252 |
PureThy.get_thm thy9 (Name ("at_" ^ Sign.base_name s ^ "_inst"))]) dt_atomTs); |
|
1253 |
||
1254 |
val induct_aux_lemmas' = map (fn Type (s, _) => |
|
1255 |
PureThy.get_thm thy9 (Name ("pt_" ^ Sign.base_name s ^ "2")) RS sym) dt_atomTs; |
|
1256 |
||
1257 |
val induct_aux = standard (Goal.prove thy9 [] ind_prems' ind_concl' |
|
1258 |
(fn prems => EVERY |
|
1259 |
([mk_subgoal 1 (K (K (K aux_ind_concl))), |
|
1260 |
indtac dt_induct tnames 1] @ |
|
1261 |
List.concat (map (fn ((_, (_, _, constrs)), (_, constrs')) => |
|
1262 |
List.concat (map (fn ((cname, cargs), is) => |
|
1263 |
[simp_tac (HOL_basic_ss addsimps List.concat perm_simps') 1, |
|
1264 |
REPEAT (rtac allI 1)] @ |
|
1265 |
List.concat (map |
|
1266 |
(fn ((_, 0), _) => [] |
|
1267 |
| ((i, j), k) => List.concat (map (fn j' => |
|
1268 |
let |
|
1269 |
val DtType (tname, _) = List.nth (cargs, i + j'); |
|
1270 |
val atom = Sign.base_name tname |
|
1271 |
in |
|
1272 |
[mk_subgoal 1 (mk_fresh i (i + j) j' |
|
1273 |
(length cargs) (length dt_atomTs) |
|
1274 |
(length cargs + length dt_atomTs + 1 + i - k)), |
|
1275 |
rtac at_finite_select 1, |
|
1276 |
rtac (PureThy.get_thm thy9 (Name ("at_" ^ atom ^ "_inst"))) 1, |
|
1277 |
asm_full_simp_tac (simpset_of thy9 addsimps |
|
1278 |
[PureThy.get_thm thy9 (Name ("fs_" ^ atom ^ "1"))]) 1, |
|
1279 |
resolve_tac prems 1, |
|
1280 |
etac exE 1, |
|
1281 |
asm_full_simp_tac (simpset_of thy9 addsimps |
|
1282 |
[symmetric fresh_def]) 1] |
|
1283 |
end) (0 upto j - 1))) (is ~~ (0 upto length is - 1))) @ |
|
1284 |
(if exists (not o equal 0 o snd) is then |
|
1285 |
[mk_subgoal 1 (mk_fresh_constr is (length dt_atomTs)), |
|
1286 |
asm_full_simp_tac (simpset_of thy9 addsimps |
|
1287 |
(List.concat inject_thms @ |
|
1288 |
alpha @ abs_perm @ abs_fresh @ [abs_fun_finite_supp] @ |
|
1289 |
induct_aux_lemmas)) 1, |
|
1290 |
dtac sym 1, |
|
1291 |
asm_full_simp_tac (simpset_of thy9 |
|
1292 |
addsimps induct_aux_lemmas' |
|
1293 |
addsimprocs [perm_simproc]) 1, |
|
1294 |
REPEAT (etac conjE 1)] |
|
1295 |
else |
|
1296 |
[]) @ |
|
1297 |
[(resolve_tac prems THEN_ALL_NEW |
|
1298 |
(atac ORELSE' ((REPEAT o etac allE) THEN' atac))) 1]) |
|
1299 |
(constrs ~~ constrs'))) (descr'' ~~ ndescr)) @ |
|
1300 |
[REPEAT (eresolve_tac [conjE, allE_Nil] 1), |
|
1301 |
REPEAT (etac allE 1), |
|
1302 |
REPEAT (TRY (rtac conjI 1) THEN asm_full_simp_tac (simpset_of thy9) 1)]))); |
|
1303 |
||
1304 |
val induct_aux' = Thm.instantiate ([], |
|
1305 |
map (fn (s, T) => |
|
1306 |
let val pT = TVar (("'n", 0), HOLogic.typeS) --> T --> HOLogic.boolT |
|
1307 |
in (cterm_of thy9 (Var ((s, 0), pT)), cterm_of thy9 (Free (s, pT))) end) |
|
1308 |
(pnames ~~ recTs) @ |
|
1309 |
map (fn (_, f) => |
|
1310 |
let val f' = Logic.varify f |
|
1311 |
in (cterm_of thy9 f', |
|
19494 | 1312 |
cterm_of thy9 (Const ("Nominal.supp", fastype_of f'))) |
18658 | 1313 |
end) fresh_fs) induct_aux; |
1314 |
||
1315 |
val induct = standard (Goal.prove thy9 [] ind_prems ind_concl |
|
1316 |
(fn prems => EVERY |
|
1317 |
[rtac induct_aux' 1, |
|
1318 |
REPEAT (resolve_tac induct_aux_lemmas 1), |
|
1319 |
REPEAT ((resolve_tac prems THEN_ALL_NEW |
|
1320 |
(etac meta_spec ORELSE' full_simp_tac (HOL_basic_ss addsimps [fresh_def]))) 1)])) |
|
1321 |
||
1322 |
val (_, thy10) = thy9 |> |
|
18016 | 1323 |
Theory.add_path big_name |> |
18658 | 1324 |
PureThy.add_thms [(("induct'", induct_aux), [])] ||>> |
1325 |
PureThy.add_thms [(("induct", induct), [case_names_induct])] ||>> |
|
1326 |
PureThy.add_thmss [(("inducts", projections induct), [case_names_induct])] ||> |
|
1327 |
Theory.parent_path; |
|
1328 |
||
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1329 |
(**** recursion combinator ****) |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1330 |
|
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1331 |
val _ = warning "defining recursion combinator ..."; |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1332 |
|
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1333 |
val used = foldr add_typ_tfree_names [] recTs; |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1334 |
|
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1335 |
val (rec_result_Ts, rec_fn_Ts) = DatatypeProp.make_primrec_Ts descr' sorts' used; |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1336 |
|
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1337 |
val permTs = map mk_permT dt_atomTs; |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1338 |
val perms = map Free |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1339 |
(DatatypeProp.indexify_names (replicate (length permTs) "pi") ~~ permTs); |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1340 |
|
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1341 |
val rec_set_Ts = map (fn (T1, T2) => rec_fn_Ts ---> HOLogic.mk_setT |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1342 |
(HOLogic.mk_prodT (T1, permTs ---> T2))) (recTs ~~ rec_result_Ts); |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1343 |
|
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1344 |
val big_rec_name = big_name ^ "_rec_set"; |
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1345 |
val rec_set_names = map (Sign.full_name (Theory.sign_of thy10)) |
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1346 |
(if length descr'' = 1 then [big_rec_name] else |
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1347 |
(map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int) |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1348 |
(1 upto (length descr'')))); |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1349 |
|
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1350 |
val rec_fns = map (uncurry (mk_Free "f")) |
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1351 |
(rec_fn_Ts ~~ (1 upto (length rec_fn_Ts))); |
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1352 |
val rec_sets = map (fn c => list_comb (Const c, rec_fns)) |
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1353 |
(rec_set_names ~~ rec_set_Ts); |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1354 |
|
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1355 |
(* introduction rules for graph of recursion function *) |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1356 |
|
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1357 |
fun partition_cargs idxs xs = map (fn (i, j) => |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1358 |
(List.take (List.drop (xs, i), j), List.nth (xs, i + j))) idxs; |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1359 |
|
19494 | 1360 |
fun mk_fresh_fun (a, t) = Const ("Nominal.fresh_fun", |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1361 |
(fastype_of a --> fastype_of t) --> fastype_of t) $ lambda a t; |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1362 |
|
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1363 |
fun make_rec_intr T rec_set ((rec_intr_ts, l), ((cname, cargs), idxs)) = |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1364 |
let |
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1365 |
fun mk_prem ((dts, (dt, U)), (j, k, prems, t1s, t2s, t3s, atoms)) = |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1366 |
let |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1367 |
val free1 = mk_Free "x" U (j + length dts); |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1368 |
val Us = map snd dts; |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1369 |
val xs = Us ~~ (j upto j + length dts - 1); |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1370 |
val frees = map (uncurry (mk_Free "x")) xs; |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1371 |
val frees' = map (uncurry (mk_Free "x'")) xs; |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1372 |
val frees'' = Us ~~ (frees ~~ frees'); |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1373 |
val pis = map (fn (T, p) => |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1374 |
case filter (equal T o fst) frees'' of |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1375 |
[] => p |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1376 |
| xs => HOLogic.mk_binop "List.op @" (p, |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1377 |
HOLogic.mk_list (HOLogic.mk_prod o snd) |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1378 |
(HOLogic.mk_prodT (T, T)) xs)) |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1379 |
(dt_atomTs ~~ perms) |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1380 |
in (case dt of |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1381 |
DtRec m => |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1382 |
let val free2 = mk_Free "y" |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1383 |
(permTs ---> List.nth (rec_result_Ts, m)) k |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1384 |
in (j + length dts + 1, k + 1, |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1385 |
HOLogic.mk_Trueprop |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1386 |
(HOLogic.mk_mem (HOLogic.mk_prod |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1387 |
(free1, free2), |
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1388 |
List.nth (rec_sets, m))) :: prems, |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1389 |
frees @ free1 :: t1s, |
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1390 |
frees' @ foldr (mk_perm []) free1 pis :: t2s, |
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1391 |
list_comb (free2, pis) :: t3s, |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1392 |
frees' @ atoms) |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1393 |
end |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1394 |
| _ => (j + length dts + 1, k, prems, |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1395 |
frees @ free1 :: t1s, |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1396 |
frees' @ foldr (mk_perm []) free1 pis :: t2s, |
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1397 |
t3s, |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1398 |
frees' @ atoms)) |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1399 |
end; |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1400 |
|
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1401 |
val Ts = map (typ_of_dtyp descr'' sorts') cargs; |
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1402 |
val (_, _, prems, t1s, t2s, t3s, atoms) = foldr mk_prem (1, 1, [], [], [], [], []) |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1403 |
(partition_cargs idxs (cargs ~~ Ts)) |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1404 |
|
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1405 |
in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop (HOLogic.mk_mem |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1406 |
(HOLogic.mk_prod (list_comb (Const (cname, Ts ---> T), t1s), |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1407 |
foldr (uncurry lambda) |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1408 |
(foldr mk_fresh_fun |
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1409 |
(list_comb (List.nth (rec_fns, l), t2s @ t3s)) atoms) |
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1410 |
perms), rec_set)))], l + 1) |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1411 |
end; |
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1412 |
|
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1413 |
val (rec_intr_ts, _) = Library.foldl (fn (x, (((d, d'), T), rec_set)) => |
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1414 |
Library.foldl (make_rec_intr T rec_set) (x, #3 (snd d) ~~ snd d')) |
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1415 |
(([], 0), descr'' ~~ ndescr ~~ recTs ~~ rec_sets); |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1416 |
|
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1417 |
val (thy11, {intrs = rec_intrs, elims = rec_elims, ...}) = |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1418 |
setmp InductivePackage.quiet_mode (!quiet_mode) |
19322
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1419 |
(InductivePackage.add_inductive_i false true big_rec_name false false false |
bf84bdf05f14
Replaced iteration combinator by recursion combinator.
berghofe
parents:
19275
diff
changeset
|
1420 |
rec_sets (map (fn x => (("", x), [])) rec_intr_ts) []) thy10; |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1421 |
|
17870 | 1422 |
in |
19251
6bc0dda66f32
First version of function for defining graph of iteration combinator.
berghofe
parents:
19134
diff
changeset
|
1423 |
thy11 |
17870 | 1424 |
end; |
1425 |
||
1426 |
val add_nominal_datatype = gen_add_nominal_datatype read_typ true; |
|
1427 |
||
1428 |
||
1429 |
(* FIXME: The following stuff should be exported by DatatypePackage *) |
|
1430 |
||
1431 |
local structure P = OuterParse and K = OuterKeyword in |
|
1432 |
||
1433 |
val datatype_decl = |
|
1434 |
Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix -- |
|
1435 |
(P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix)); |
|
1436 |
||
1437 |
fun mk_datatype args = |
|
1438 |
let |
|
1439 |
val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args; |
|
1440 |
val specs = map (fn ((((_, vs), t), mx), cons) => |
|
1441 |
(vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args; |
|
18068 | 1442 |
in add_nominal_datatype false names specs end; |
17870 | 1443 |
|
1444 |
val nominal_datatypeP = |
|
1445 |
OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl |
|
1446 |
(P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype)); |
|
1447 |
||
1448 |
val _ = OuterSyntax.add_parsers [nominal_datatypeP]; |
|
1449 |
||
1450 |
end; |
|
1451 |
||
18261
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
1452 |
end |
1318955d57ac
Corrected treatment of non-recursive abstraction types.
berghofe
parents:
18246
diff
changeset
|
1453 |