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