author | urbanc |
Mon, 07 Nov 2005 11:17:45 +0100 | |
changeset 18100 | 193c3382bbfe |
parent 18068 | e8c3d371594e |
child 18262 | d0fcd4d684f5 |
permissions | -rw-r--r-- |
18068 | 1 |
(* $Id$ *) |
2 |
||
3 |
signature NOMINAL_ATOMS = |
|
4 |
sig |
|
5 |
val create_nom_typedecls : string list -> theory -> theory |
|
6 |
val atoms_of : theory -> string list |
|
7 |
val mk_permT : typ -> typ |
|
8 |
val setup : (theory -> theory) list |
|
9 |
end |
|
10 |
||
11 |
structure NominalAtoms : NOMINAL_ATOMS = |
|
12 |
struct |
|
13 |
||
14 |
(* data kind 'HOL/nominal' *) |
|
15 |
||
16 |
structure NominalArgs = |
|
17 |
struct |
|
18 |
val name = "HOL/nominal"; |
|
19 |
type T = unit Symtab.table; |
|
20 |
||
21 |
val empty = Symtab.empty; |
|
22 |
val copy = I; |
|
23 |
val extend = I; |
|
24 |
fun merge _ x = Symtab.merge (K true) x; |
|
25 |
||
26 |
fun print sg tab = (); |
|
27 |
end; |
|
28 |
||
29 |
structure NominalData = TheoryDataFun(NominalArgs); |
|
30 |
||
31 |
fun atoms_of thy = map fst (Symtab.dest (NominalData.get thy)); |
|
32 |
||
33 |
(* FIXME: add to hologic.ML ? *) |
|
34 |
fun mk_listT T = Type ("List.list", [T]); |
|
35 |
fun mk_permT T = mk_listT (HOLogic.mk_prodT (T, T)); |
|
36 |
||
37 |
fun mk_Cons x xs = |
|
38 |
let val T = fastype_of x |
|
39 |
in Const ("List.list.Cons", T --> mk_listT T --> mk_listT T) $ x $ xs end; |
|
40 |
||
41 |
||
42 |
(* this function sets up all matters related to atom- *) |
|
43 |
(* kinds; the user specifies a list of atom-kind names *) |
|
44 |
(* atom_decl <ak1> ... <akn> *) |
|
45 |
fun create_nom_typedecls ak_names thy = |
|
46 |
let |
|
47 |
(* declares a type-decl for every atom-kind: *) |
|
48 |
(* that is typedecl <ak> *) |
|
49 |
val thy1 = TypedefPackage.add_typedecls (map (fn x => (x,[],NoSyn)) ak_names) thy; |
|
50 |
||
51 |
(* produces a list consisting of pairs: *) |
|
52 |
(* fst component is the atom-kind name *) |
|
53 |
(* snd component is its type *) |
|
54 |
val full_ak_names = map (Sign.intern_type (sign_of thy1)) ak_names; |
|
55 |
val ak_names_types = ak_names ~~ map (Type o rpair []) full_ak_names; |
|
56 |
||
57 |
(* adds for every atom-kind an axiom *) |
|
58 |
(* <ak>_infinite: infinite (UNIV::<ak_type> set) *) |
|
59 |
val (thy2,inf_axs) = PureThy.add_axioms_i (map (fn (ak_name, T) => |
|
60 |
let |
|
61 |
val name = ak_name ^ "_infinite" |
|
62 |
val axiom = HOLogic.mk_Trueprop (HOLogic.mk_not |
|
63 |
(HOLogic.mk_mem (HOLogic.mk_UNIV T, |
|
64 |
Const ("Finite_Set.Finites", HOLogic.mk_setT (HOLogic.mk_setT T))))) |
|
65 |
in |
|
66 |
((name, axiom), []) |
|
67 |
end) ak_names_types) thy1; |
|
68 |
||
69 |
(* declares a swapping function for every atom-kind, it is *) |
|
70 |
(* const swap_<ak> :: <akT> * <akT> => <akT> => <akT> *) |
|
71 |
(* swap_<ak> (a,b) c = (if a=c then b (else if b=c then a else c)) *) |
|
72 |
(* overloades then the general swap-function *) |
|
73 |
val (thy3, swap_eqs) = foldl_map (fn (thy, (ak_name, T)) => |
|
74 |
let |
|
75 |
val swapT = HOLogic.mk_prodT (T, T) --> T --> T; |
|
76 |
val swap_name = Sign.full_name (sign_of thy) ("swap_" ^ ak_name); |
|
77 |
val a = Free ("a", T); |
|
78 |
val b = Free ("b", T); |
|
79 |
val c = Free ("c", T); |
|
80 |
val ab = Free ("ab", HOLogic.mk_prodT (T, T)) |
|
81 |
val cif = Const ("HOL.If", HOLogic.boolT --> T --> T --> T); |
|
82 |
val cswap_akname = Const (swap_name, swapT); |
|
83 |
val cswap = Const ("nominal.swap", swapT) |
|
84 |
||
85 |
val name = "swap_"^ak_name^"_def"; |
|
86 |
val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
87 |
(cswap_akname $ HOLogic.mk_prod (a,b) $ c, |
|
88 |
cif $ HOLogic.mk_eq (a,c) $ b $ (cif $ HOLogic.mk_eq (b,c) $ a $ c))) |
|
89 |
val def2 = Logic.mk_equals (cswap $ ab $ c, cswap_akname $ ab $ c) |
|
90 |
in |
|
91 |
thy |> Theory.add_consts_i [("swap_" ^ ak_name, swapT, NoSyn)] |
|
92 |
|> (#1 o PureThy.add_defs_i true [((name, def2),[])]) |
|
93 |
|> PrimrecPackage.add_primrec_i "" [(("", def1),[])] |
|
94 |
end) (thy2, ak_names_types); |
|
95 |
||
96 |
(* declares a permutation function for every atom-kind acting *) |
|
97 |
(* on such atoms *) |
|
98 |
(* const <ak>_prm_<ak> :: (<akT> * <akT>)list => akT => akT *) |
|
99 |
(* <ak>_prm_<ak> [] a = a *) |
|
100 |
(* <ak>_prm_<ak> (x#xs) a = swap_<ak> x (perm xs a) *) |
|
101 |
val (thy4, prm_eqs) = foldl_map (fn (thy, (ak_name, T)) => |
|
102 |
let |
|
103 |
val swapT = HOLogic.mk_prodT (T, T) --> T --> T; |
|
104 |
val swap_name = Sign.full_name (sign_of thy) ("swap_" ^ ak_name) |
|
105 |
val prmT = mk_permT T --> T --> T; |
|
106 |
val prm_name = ak_name ^ "_prm_" ^ ak_name; |
|
107 |
val qu_prm_name = Sign.full_name (sign_of thy) prm_name; |
|
108 |
val x = Free ("x", HOLogic.mk_prodT (T, T)); |
|
109 |
val xs = Free ("xs", mk_permT T); |
|
110 |
val a = Free ("a", T) ; |
|
111 |
||
112 |
val cnil = Const ("List.list.Nil", mk_permT T); |
|
113 |
||
114 |
val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (qu_prm_name, prmT) $ cnil $ a, a)); |
|
115 |
||
116 |
val def2 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
117 |
(Const (qu_prm_name, prmT) $ mk_Cons x xs $ a, |
|
118 |
Const (swap_name, swapT) $ x $ (Const (qu_prm_name, prmT) $ xs $ a))); |
|
119 |
in |
|
120 |
thy |> Theory.add_consts_i [(prm_name, mk_permT T --> T --> T, NoSyn)] |
|
121 |
|> PrimrecPackage.add_primrec_i "" [(("", def1), []),(("", def2), [])] |
|
122 |
end) (thy3, ak_names_types); |
|
123 |
||
124 |
(* defines permutation functions for all combinations of atom-kinds; *) |
|
125 |
(* there are a trivial cases and non-trivial cases *) |
|
126 |
(* non-trivial case: *) |
|
127 |
(* <ak>_prm_<ak>_def: perm pi a == <ak>_prm_<ak> pi a *) |
|
128 |
(* trivial case with <ak> != <ak'> *) |
|
129 |
(* <ak>_prm<ak'>_def[simp]: perm pi a == a *) |
|
130 |
(* *) |
|
131 |
(* the trivial cases are added to the simplifier, while the non- *) |
|
132 |
(* have their own rules proved below *) |
|
133 |
val (thy5, perm_defs) = foldl_map (fn (thy, (ak_name, T)) => |
|
134 |
foldl_map (fn (thy', (ak_name', T')) => |
|
135 |
let |
|
136 |
val perm_def_name = ak_name ^ "_prm_" ^ ak_name'; |
|
137 |
val pi = Free ("pi", mk_permT T); |
|
138 |
val a = Free ("a", T'); |
|
139 |
val cperm = Const ("nominal.perm", mk_permT T --> T' --> T'); |
|
140 |
val cperm_def = Const (Sign.full_name (sign_of thy') perm_def_name, mk_permT T --> T' --> T'); |
|
141 |
||
142 |
val name = ak_name ^ "_prm_" ^ ak_name' ^ "_def"; |
|
143 |
val def = Logic.mk_equals |
|
144 |
(cperm $ pi $ a, if ak_name = ak_name' then cperm_def $ pi $ a else a) |
|
145 |
in |
|
146 |
thy' |> PureThy.add_defs_i true [((name, def),[])] |
|
147 |
end) (thy, ak_names_types)) (thy4, ak_names_types); |
|
148 |
||
149 |
(* proves that every atom-kind is an instance of at *) |
|
150 |
(* lemma at_<ak>_inst: *) |
|
151 |
(* at TYPE(<ak>) *) |
|
152 |
val (thy6, prm_cons_thms) = |
|
153 |
thy5 |> PureThy.add_thms (map (fn (ak_name, T) => |
|
154 |
let |
|
155 |
val ak_name_qu = Sign.full_name (sign_of thy5) (ak_name); |
|
156 |
val i_type = Type(ak_name_qu,[]); |
|
157 |
val cat = Const ("nominal.at",(Term.itselfT i_type) --> HOLogic.boolT); |
|
158 |
val at_type = Logic.mk_type i_type; |
|
159 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy5 |
|
160 |
[Name "at_def", |
|
161 |
Name (ak_name ^ "_prm_" ^ ak_name ^ "_def"), |
|
162 |
Name (ak_name ^ "_prm_" ^ ak_name ^ ".simps"), |
|
163 |
Name ("swap_" ^ ak_name ^ "_def"), |
|
164 |
Name ("swap_" ^ ak_name ^ ".simps"), |
|
165 |
Name (ak_name ^ "_infinite")] |
|
166 |
||
167 |
val name = "at_"^ak_name^ "_inst"; |
|
168 |
val statement = HOLogic.mk_Trueprop (cat $ at_type); |
|
169 |
||
170 |
val proof = fn _ => auto_tac (claset(),simp_s); |
|
171 |
||
172 |
in |
|
173 |
((name, standard (Goal.prove thy5 [] [] statement proof)), []) |
|
174 |
end) ak_names_types); |
|
175 |
||
176 |
(* declares a perm-axclass for every atom-kind *) |
|
177 |
(* axclass pt_<ak> *) |
|
178 |
(* pt_<ak>1[simp]: perm [] x = x *) |
|
179 |
(* pt_<ak>2: perm (pi1@pi2) x = perm pi1 (perm pi2 x) *) |
|
180 |
(* pt_<ak>3: pi1 ~ pi2 ==> perm pi1 x = perm pi2 x *) |
|
181 |
val (thy7, pt_ax_classes) = foldl_map (fn (thy, (ak_name, T)) => |
|
182 |
let |
|
183 |
val cl_name = "pt_"^ak_name; |
|
184 |
val ty = TFree("'a",["HOL.type"]); |
|
185 |
val x = Free ("x", ty); |
|
186 |
val pi1 = Free ("pi1", mk_permT T); |
|
187 |
val pi2 = Free ("pi2", mk_permT T); |
|
188 |
val cperm = Const ("nominal.perm", mk_permT T --> ty --> ty); |
|
189 |
val cnil = Const ("List.list.Nil", mk_permT T); |
|
190 |
val cappend = Const ("List.op @",mk_permT T --> mk_permT T --> mk_permT T); |
|
191 |
val cprm_eq = Const ("nominal.prm_eq",mk_permT T --> mk_permT T --> HOLogic.boolT); |
|
192 |
(* nil axiom *) |
|
193 |
val axiom1 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
194 |
(cperm $ cnil $ x, x)); |
|
195 |
(* append axiom *) |
|
196 |
val axiom2 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
197 |
(cperm $ (cappend $ pi1 $ pi2) $ x, cperm $ pi1 $ (cperm $ pi2 $ x))); |
|
198 |
(* perm-eq axiom *) |
|
199 |
val axiom3 = Logic.mk_implies |
|
200 |
(HOLogic.mk_Trueprop (cprm_eq $ pi1 $ pi2), |
|
201 |
HOLogic.mk_Trueprop (HOLogic.mk_eq (cperm $ pi1 $ x, cperm $ pi2 $ x))); |
|
202 |
in |
|
203 |
thy |> AxClass.add_axclass_i (cl_name, ["HOL.type"]) |
|
204 |
[((cl_name^"1", axiom1),[Simplifier.simp_add_global]), |
|
205 |
((cl_name^"2", axiom2),[]), |
|
206 |
((cl_name^"3", axiom3),[])] |
|
207 |
end) (thy6,ak_names_types); |
|
208 |
||
209 |
(* proves that every pt_<ak>-type together with <ak>-type *) |
|
210 |
(* instance of pt *) |
|
211 |
(* lemma pt_<ak>_inst: *) |
|
212 |
(* pt TYPE('x::pt_<ak>) TYPE(<ak>) *) |
|
213 |
val (thy8, prm_inst_thms) = |
|
214 |
thy7 |> PureThy.add_thms (map (fn (ak_name, T) => |
|
215 |
let |
|
216 |
val ak_name_qu = Sign.full_name (sign_of thy7) (ak_name); |
|
217 |
val pt_name_qu = Sign.full_name (sign_of thy7) ("pt_"^ak_name); |
|
218 |
val i_type1 = TFree("'x",[pt_name_qu]); |
|
219 |
val i_type2 = Type(ak_name_qu,[]); |
|
220 |
val cpt = Const ("nominal.pt",(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); |
|
221 |
val pt_type = Logic.mk_type i_type1; |
|
222 |
val at_type = Logic.mk_type i_type2; |
|
223 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy7 |
|
224 |
[Name "pt_def", |
|
225 |
Name ("pt_" ^ ak_name ^ "1"), |
|
226 |
Name ("pt_" ^ ak_name ^ "2"), |
|
227 |
Name ("pt_" ^ ak_name ^ "3")]; |
|
228 |
||
229 |
val name = "pt_"^ak_name^ "_inst"; |
|
230 |
val statement = HOLogic.mk_Trueprop (cpt $ pt_type $ at_type); |
|
231 |
||
232 |
val proof = fn _ => auto_tac (claset(),simp_s); |
|
233 |
in |
|
234 |
((name, standard (Goal.prove thy7 [] [] statement proof)), []) |
|
235 |
end) ak_names_types); |
|
236 |
||
237 |
(* declares an fs-axclass for every atom-kind *) |
|
238 |
(* axclass fs_<ak> *) |
|
239 |
(* fs_<ak>1: finite ((supp x)::<ak> set) *) |
|
240 |
val (thy11, fs_ax_classes) = foldl_map (fn (thy, (ak_name, T)) => |
|
241 |
let |
|
242 |
val cl_name = "fs_"^ak_name; |
|
243 |
val pt_name = Sign.full_name (sign_of thy) ("pt_"^ak_name); |
|
244 |
val ty = TFree("'a",["HOL.type"]); |
|
245 |
val x = Free ("x", ty); |
|
246 |
val csupp = Const ("nominal.supp", ty --> HOLogic.mk_setT T); |
|
247 |
val cfinites = Const ("Finite_Set.Finites", HOLogic.mk_setT (HOLogic.mk_setT T)) |
|
248 |
||
249 |
val axiom1 = HOLogic.mk_Trueprop (HOLogic.mk_mem (csupp $ x, cfinites)); |
|
250 |
||
251 |
in |
|
252 |
thy |> AxClass.add_axclass_i (cl_name, [pt_name]) [((cl_name^"1", axiom1),[])] |
|
253 |
end) (thy8,ak_names_types); |
|
254 |
||
255 |
(* proves that every fs_<ak>-type together with <ak>-type *) |
|
256 |
(* instance of fs-type *) |
|
257 |
(* lemma abst_<ak>_inst: *) |
|
258 |
(* fs TYPE('x::pt_<ak>) TYPE (<ak>) *) |
|
259 |
val (thy12, fs_inst_thms) = |
|
260 |
thy11 |> PureThy.add_thms (map (fn (ak_name, T) => |
|
261 |
let |
|
262 |
val ak_name_qu = Sign.full_name (sign_of thy11) (ak_name); |
|
263 |
val fs_name_qu = Sign.full_name (sign_of thy11) ("fs_"^ak_name); |
|
264 |
val i_type1 = TFree("'x",[fs_name_qu]); |
|
265 |
val i_type2 = Type(ak_name_qu,[]); |
|
266 |
val cfs = Const ("nominal.fs", |
|
267 |
(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); |
|
268 |
val fs_type = Logic.mk_type i_type1; |
|
269 |
val at_type = Logic.mk_type i_type2; |
|
270 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy11 |
|
271 |
[Name "fs_def", |
|
272 |
Name ("fs_" ^ ak_name ^ "1")]; |
|
273 |
||
274 |
val name = "fs_"^ak_name^ "_inst"; |
|
275 |
val statement = HOLogic.mk_Trueprop (cfs $ fs_type $ at_type); |
|
276 |
||
277 |
val proof = fn _ => auto_tac (claset(),simp_s); |
|
278 |
in |
|
279 |
((name, standard (Goal.prove thy11 [] [] statement proof)), []) |
|
280 |
end) ak_names_types); |
|
281 |
||
282 |
(* declares for every atom-kind combination an axclass *) |
|
283 |
(* cp_<ak1>_<ak2> giving a composition property *) |
|
284 |
(* cp_<ak1>_<ak2>1: pi1 o pi2 o x = (pi1 o pi2) o (pi1 o x) *) |
|
285 |
val (thy12b,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
286 |
foldl_map (fn (thy', (ak_name', T')) => |
|
287 |
let |
|
288 |
val cl_name = "cp_"^ak_name^"_"^ak_name'; |
|
289 |
val ty = TFree("'a",["HOL.type"]); |
|
290 |
val x = Free ("x", ty); |
|
291 |
val pi1 = Free ("pi1", mk_permT T); |
|
292 |
val pi2 = Free ("pi2", mk_permT T'); |
|
293 |
val cperm1 = Const ("nominal.perm", mk_permT T --> ty --> ty); |
|
294 |
val cperm2 = Const ("nominal.perm", mk_permT T' --> ty --> ty); |
|
295 |
val cperm3 = Const ("nominal.perm", mk_permT T --> mk_permT T' --> mk_permT T'); |
|
296 |
||
297 |
val ax1 = HOLogic.mk_Trueprop |
|
298 |
(HOLogic.mk_eq (cperm1 $ pi1 $ (cperm2 $ pi2 $ x), |
|
299 |
cperm2 $ (cperm3 $ pi1 $ pi2) $ (cperm1 $ pi1 $ x))); |
|
300 |
in |
|
301 |
(fst (AxClass.add_axclass_i (cl_name, ["HOL.type"]) [((cl_name^"1", ax1),[])] thy'),()) |
|
302 |
end) |
|
303 |
(thy, ak_names_types)) (thy12, ak_names_types) |
|
304 |
||
305 |
(* proves for every <ak>-combination a cp_<ak1>_<ak2>_inst theorem; *) |
|
306 |
(* lemma cp_<ak1>_<ak2>_inst: *) |
|
307 |
(* cp TYPE('a::cp_<ak1>_<ak2>) TYPE(<ak1>) TYPE(<ak2>) *) |
|
308 |
val (thy12c, cp_thms) = foldl_map (fn (thy, (ak_name, T)) => |
|
309 |
foldl_map (fn (thy', (ak_name', T')) => |
|
310 |
let |
|
311 |
val ak_name_qu = Sign.full_name (sign_of thy') (ak_name); |
|
312 |
val ak_name_qu' = Sign.full_name (sign_of thy') (ak_name'); |
|
313 |
val cp_name_qu = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name'); |
|
314 |
val i_type0 = TFree("'a",[cp_name_qu]); |
|
315 |
val i_type1 = Type(ak_name_qu,[]); |
|
316 |
val i_type2 = Type(ak_name_qu',[]); |
|
317 |
val ccp = Const ("nominal.cp", |
|
318 |
(Term.itselfT i_type0)-->(Term.itselfT i_type1)--> |
|
319 |
(Term.itselfT i_type2)-->HOLogic.boolT); |
|
320 |
val at_type = Logic.mk_type i_type1; |
|
321 |
val at_type' = Logic.mk_type i_type2; |
|
322 |
val cp_type = Logic.mk_type i_type0; |
|
323 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy' [(Name "cp_def")]; |
|
324 |
val cp1 = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"1")); |
|
325 |
||
326 |
val name = "cp_"^ak_name^ "_"^ak_name'^"_inst"; |
|
327 |
val statement = HOLogic.mk_Trueprop (ccp $ cp_type $ at_type $ at_type'); |
|
328 |
||
329 |
val proof = fn _ => EVERY [auto_tac (claset(),simp_s), rtac cp1 1]; |
|
330 |
in |
|
331 |
thy' |> PureThy.add_thms |
|
332 |
[((name, standard (Goal.prove thy' [] [] statement proof)), [])] |
|
333 |
end) |
|
334 |
(thy, ak_names_types)) (thy12b, ak_names_types); |
|
335 |
||
336 |
(* proves for every non-trivial <ak>-combination a disjointness *) |
|
337 |
(* theorem; i.e. <ak1> != <ak2> *) |
|
338 |
(* lemma ds_<ak1>_<ak2>: *) |
|
339 |
(* dj TYPE(<ak1>) TYPE(<ak2>) *) |
|
340 |
val (thy12d, dj_thms) = foldl_map (fn (thy, (ak_name, T)) => |
|
341 |
foldl_map (fn (thy', (ak_name', T')) => |
|
342 |
(if not (ak_name = ak_name') |
|
343 |
then |
|
344 |
let |
|
345 |
val ak_name_qu = Sign.full_name (sign_of thy') (ak_name); |
|
346 |
val ak_name_qu' = Sign.full_name (sign_of thy') (ak_name'); |
|
347 |
val i_type1 = Type(ak_name_qu,[]); |
|
348 |
val i_type2 = Type(ak_name_qu',[]); |
|
349 |
val cdj = Const ("nominal.disjoint", |
|
350 |
(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); |
|
351 |
val at_type = Logic.mk_type i_type1; |
|
352 |
val at_type' = Logic.mk_type i_type2; |
|
353 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy' |
|
354 |
[Name "disjoint_def", |
|
355 |
Name (ak_name^"_prm_"^ak_name'^"_def"), |
|
356 |
Name (ak_name'^"_prm_"^ak_name^"_def")]; |
|
357 |
||
358 |
val name = "dj_"^ak_name^"_"^ak_name'; |
|
359 |
val statement = HOLogic.mk_Trueprop (cdj $ at_type $ at_type'); |
|
360 |
||
361 |
val proof = fn _ => auto_tac (claset(),simp_s); |
|
362 |
in |
|
363 |
thy' |> PureThy.add_thms |
|
364 |
[((name, standard (Goal.prove thy' [] [] statement proof)), []) ] |
|
365 |
end |
|
366 |
else |
|
367 |
(thy',[]))) (* do nothing branch, if ak_name = ak_name' *) |
|
368 |
(thy, ak_names_types)) (thy12c, ak_names_types); |
|
369 |
||
370 |
(*<<<<<<< pt_<ak> class instances >>>>>>>*) |
|
371 |
(*=========================================*) |
|
372 |
||
373 |
(* some frequently used theorems *) |
|
374 |
val pt1 = PureThy.get_thm thy12c (Name "pt1"); |
|
375 |
val pt2 = PureThy.get_thm thy12c (Name "pt2"); |
|
376 |
val pt3 = PureThy.get_thm thy12c (Name "pt3"); |
|
377 |
val at_pt_inst = PureThy.get_thm thy12c (Name "at_pt_inst"); |
|
378 |
val pt_bool_inst = PureThy.get_thm thy12c (Name "pt_bool_inst"); |
|
379 |
val pt_set_inst = PureThy.get_thm thy12c (Name "pt_set_inst"); |
|
380 |
val pt_unit_inst = PureThy.get_thm thy12c (Name "pt_unit_inst"); |
|
381 |
val pt_prod_inst = PureThy.get_thm thy12c (Name "pt_prod_inst"); |
|
382 |
val pt_list_inst = PureThy.get_thm thy12c (Name "pt_list_inst"); |
|
383 |
val pt_optn_inst = PureThy.get_thm thy12c (Name "pt_option_inst"); |
|
384 |
val pt_noptn_inst = PureThy.get_thm thy12c (Name "pt_noption_inst"); |
|
385 |
val pt_fun_inst = PureThy.get_thm thy12c (Name "pt_fun_inst"); |
|
386 |
||
387 |
(* for all atom-kind combination shows that *) |
|
388 |
(* every <ak> is an instance of pt_<ai> *) |
|
389 |
val (thy13,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
390 |
foldl_map (fn (thy', (ak_name', T')) => |
|
391 |
(if ak_name = ak_name' |
|
392 |
then |
|
393 |
let |
|
394 |
val qu_name = Sign.full_name (sign_of thy') ak_name; |
|
395 |
val qu_class = Sign.full_name (sign_of thy') ("pt_"^ak_name); |
|
396 |
val at_inst = PureThy.get_thm thy' (Name ("at_"^ak_name ^"_inst")); |
|
397 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
398 |
rtac ((at_inst RS at_pt_inst) RS pt1) 1, |
|
399 |
rtac ((at_inst RS at_pt_inst) RS pt2) 1, |
|
400 |
rtac ((at_inst RS at_pt_inst) RS pt3) 1, |
|
401 |
atac 1]; |
|
402 |
in |
|
403 |
(AxClass.add_inst_arity_i (qu_name,[],[qu_class]) proof thy',()) |
|
404 |
end |
|
405 |
else |
|
406 |
let |
|
407 |
val qu_name' = Sign.full_name (sign_of thy') ak_name'; |
|
408 |
val qu_class = Sign.full_name (sign_of thy') ("pt_"^ak_name); |
|
409 |
val simp_s = HOL_basic_ss addsimps |
|
410 |
PureThy.get_thmss thy' [Name (ak_name^"_prm_"^ak_name'^"_def")]; |
|
411 |
val proof = EVERY [AxClass.intro_classes_tac [], auto_tac (claset(),simp_s)]; |
|
412 |
in |
|
413 |
(AxClass.add_inst_arity_i (qu_name',[],[qu_class]) proof thy',()) |
|
414 |
end)) |
|
415 |
(thy, ak_names_types)) (thy12c, ak_names_types); |
|
416 |
||
417 |
(* shows that bool is an instance of pt_<ak> *) |
|
418 |
(* uses the theorem pt_bool_inst *) |
|
419 |
val (thy14,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
420 |
let |
|
421 |
val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name); |
|
422 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
423 |
rtac (pt_bool_inst RS pt1) 1, |
|
424 |
rtac (pt_bool_inst RS pt2) 1, |
|
425 |
rtac (pt_bool_inst RS pt3) 1, |
|
426 |
atac 1]; |
|
427 |
in |
|
428 |
(AxClass.add_inst_arity_i ("bool",[],[qu_class]) proof thy,()) |
|
429 |
end) (thy13,ak_names_types); |
|
430 |
||
431 |
(* shows that set(pt_<ak>) is an instance of pt_<ak> *) |
|
432 |
(* unfolds the permutation definition and applies pt_<ak>i *) |
|
433 |
val (thy15,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
434 |
let |
|
435 |
val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name); |
|
436 |
val pt_inst = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst")); |
|
437 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
438 |
rtac ((pt_inst RS pt_set_inst) RS pt1) 1, |
|
439 |
rtac ((pt_inst RS pt_set_inst) RS pt2) 1, |
|
440 |
rtac ((pt_inst RS pt_set_inst) RS pt3) 1, |
|
441 |
atac 1]; |
|
442 |
in |
|
443 |
(AxClass.add_inst_arity_i ("set",[[qu_class]],[qu_class]) proof thy,()) |
|
444 |
end) (thy14,ak_names_types); |
|
445 |
||
446 |
(* shows that unit is an instance of pt_<ak> *) |
|
447 |
val (thy16,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
448 |
let |
|
449 |
val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name); |
|
450 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
451 |
rtac (pt_unit_inst RS pt1) 1, |
|
452 |
rtac (pt_unit_inst RS pt2) 1, |
|
453 |
rtac (pt_unit_inst RS pt3) 1, |
|
454 |
atac 1]; |
|
455 |
in |
|
456 |
(AxClass.add_inst_arity_i ("Product_Type.unit",[],[qu_class]) proof thy,()) |
|
457 |
end) (thy15,ak_names_types); |
|
458 |
||
459 |
(* shows that *(pt_<ak>,pt_<ak>) is an instance of pt_<ak> *) |
|
460 |
(* uses the theorem pt_prod_inst and pt_<ak>_inst *) |
|
461 |
val (thy17,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
462 |
let |
|
463 |
val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name); |
|
464 |
val pt_inst = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst")); |
|
465 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
466 |
rtac ((pt_inst RS (pt_inst RS pt_prod_inst)) RS pt1) 1, |
|
467 |
rtac ((pt_inst RS (pt_inst RS pt_prod_inst)) RS pt2) 1, |
|
468 |
rtac ((pt_inst RS (pt_inst RS pt_prod_inst)) RS pt3) 1, |
|
469 |
atac 1]; |
|
470 |
in |
|
471 |
(AxClass.add_inst_arity_i ("*",[[qu_class],[qu_class]],[qu_class]) proof thy,()) |
|
472 |
end) (thy16,ak_names_types); |
|
473 |
||
474 |
(* shows that list(pt_<ak>) is an instance of pt_<ak> *) |
|
475 |
(* uses the theorem pt_list_inst and pt_<ak>_inst *) |
|
476 |
val (thy18,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
477 |
let |
|
478 |
val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name); |
|
479 |
val pt_inst = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst")); |
|
480 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
481 |
rtac ((pt_inst RS pt_list_inst) RS pt1) 1, |
|
482 |
rtac ((pt_inst RS pt_list_inst) RS pt2) 1, |
|
483 |
rtac ((pt_inst RS pt_list_inst) RS pt3) 1, |
|
484 |
atac 1]; |
|
485 |
in |
|
486 |
(AxClass.add_inst_arity_i ("List.list",[[qu_class]],[qu_class]) proof thy,()) |
|
487 |
end) (thy17,ak_names_types); |
|
488 |
||
489 |
(* shows that option(pt_<ak>) is an instance of pt_<ak> *) |
|
490 |
(* uses the theorem pt_option_inst and pt_<ak>_inst *) |
|
491 |
val (thy18a,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
492 |
let |
|
493 |
val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name); |
|
494 |
val pt_inst = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst")); |
|
495 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
496 |
rtac ((pt_inst RS pt_optn_inst) RS pt1) 1, |
|
497 |
rtac ((pt_inst RS pt_optn_inst) RS pt2) 1, |
|
498 |
rtac ((pt_inst RS pt_optn_inst) RS pt3) 1, |
|
499 |
atac 1]; |
|
500 |
in |
|
501 |
(AxClass.add_inst_arity_i ("Datatype.option",[[qu_class]],[qu_class]) proof thy,()) |
|
502 |
end) (thy18,ak_names_types); |
|
503 |
||
504 |
(* shows that nOption(pt_<ak>) is an instance of pt_<ak> *) |
|
505 |
(* uses the theorem pt_option_inst and pt_<ak>_inst *) |
|
506 |
val (thy18b,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
507 |
let |
|
508 |
val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name); |
|
509 |
val pt_inst = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst")); |
|
510 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
511 |
rtac ((pt_inst RS pt_noptn_inst) RS pt1) 1, |
|
512 |
rtac ((pt_inst RS pt_noptn_inst) RS pt2) 1, |
|
513 |
rtac ((pt_inst RS pt_noptn_inst) RS pt3) 1, |
|
514 |
atac 1]; |
|
515 |
in |
|
516 |
(AxClass.add_inst_arity_i ("nominal.nOption",[[qu_class]],[qu_class]) proof thy,()) |
|
517 |
end) (thy18a,ak_names_types); |
|
518 |
||
519 |
||
520 |
(* shows that fun(pt_<ak>,pt_<ak>) is an instance of pt_<ak> *) |
|
521 |
(* uses the theorem pt_list_inst and pt_<ak>_inst *) |
|
522 |
val (thy19,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
523 |
let |
|
524 |
val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name); |
|
525 |
val at_thm = PureThy.get_thm thy (Name ("at_"^ak_name^"_inst")); |
|
526 |
val pt_inst = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst")); |
|
527 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
528 |
rtac ((at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst))) RS pt1) 1, |
|
529 |
rtac ((at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst))) RS pt2) 1, |
|
530 |
rtac ((at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst))) RS pt3) 1, |
|
531 |
atac 1]; |
|
532 |
in |
|
533 |
(AxClass.add_inst_arity_i ("fun",[[qu_class],[qu_class]],[qu_class]) proof thy,()) |
|
534 |
end) (thy18b,ak_names_types); |
|
535 |
||
536 |
(*<<<<<<< fs_<ak> class instances >>>>>>>*) |
|
537 |
(*=========================================*) |
|
538 |
val fs1 = PureThy.get_thm thy19 (Name "fs1"); |
|
539 |
val fs_at_inst = PureThy.get_thm thy19 (Name "fs_at_inst"); |
|
540 |
val fs_unit_inst = PureThy.get_thm thy19 (Name "fs_unit_inst"); |
|
541 |
val fs_bool_inst = PureThy.get_thm thy19 (Name "fs_bool_inst"); |
|
542 |
val fs_prod_inst = PureThy.get_thm thy19 (Name "fs_prod_inst"); |
|
543 |
val fs_list_inst = PureThy.get_thm thy19 (Name "fs_list_inst"); |
|
544 |
||
545 |
(* shows that <ak> is an instance of fs_<ak> *) |
|
546 |
(* uses the theorem at_<ak>_inst *) |
|
547 |
val (thy20,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
548 |
let |
|
549 |
val qu_name = Sign.full_name (sign_of thy) ak_name; |
|
550 |
val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name); |
|
551 |
val at_thm = PureThy.get_thm thy (Name ("at_"^ak_name^"_inst")); |
|
552 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
553 |
rtac ((at_thm RS fs_at_inst) RS fs1) 1]; |
|
554 |
in |
|
555 |
(AxClass.add_inst_arity_i (qu_name,[],[qu_class]) proof thy,()) |
|
556 |
end) (thy19,ak_names_types); |
|
557 |
||
558 |
(* shows that unit is an instance of fs_<ak> *) |
|
559 |
(* uses the theorem fs_unit_inst *) |
|
560 |
val (thy21,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
561 |
let |
|
562 |
val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name); |
|
563 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
564 |
rtac (fs_unit_inst RS fs1) 1]; |
|
565 |
in |
|
566 |
(AxClass.add_inst_arity_i ("Product_Type.unit",[],[qu_class]) proof thy,()) |
|
567 |
end) (thy20,ak_names_types); |
|
568 |
||
569 |
(* shows that bool is an instance of fs_<ak> *) |
|
570 |
(* uses the theorem fs_bool_inst *) |
|
571 |
val (thy22,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
572 |
let |
|
573 |
val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name); |
|
574 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
575 |
rtac (fs_bool_inst RS fs1) 1]; |
|
576 |
in |
|
577 |
(AxClass.add_inst_arity_i ("bool",[],[qu_class]) proof thy,()) |
|
578 |
end) (thy21,ak_names_types); |
|
579 |
||
580 |
(* shows that *(fs_<ak>,fs_<ak>) is an instance of fs_<ak> *) |
|
581 |
(* uses the theorem fs_prod_inst *) |
|
582 |
val (thy23,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
583 |
let |
|
584 |
val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name); |
|
585 |
val fs_inst = PureThy.get_thm thy (Name ("fs_"^ak_name^"_inst")); |
|
586 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
587 |
rtac ((fs_inst RS (fs_inst RS fs_prod_inst)) RS fs1) 1]; |
|
588 |
in |
|
589 |
(AxClass.add_inst_arity_i ("*",[[qu_class],[qu_class]],[qu_class]) proof thy,()) |
|
590 |
end) (thy22,ak_names_types); |
|
591 |
||
592 |
(* shows that list(fs_<ak>) is an instance of fs_<ak> *) |
|
593 |
(* uses the theorem fs_list_inst *) |
|
594 |
val (thy24,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
595 |
let |
|
596 |
val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name); |
|
597 |
val fs_inst = PureThy.get_thm thy (Name ("fs_"^ak_name^"_inst")); |
|
598 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
599 |
rtac ((fs_inst RS fs_list_inst) RS fs1) 1]; |
|
600 |
in |
|
601 |
(AxClass.add_inst_arity_i ("List.list",[[qu_class]],[qu_class]) proof thy,()) |
|
602 |
end) (thy23,ak_names_types); |
|
603 |
||
604 |
(*<<<<<<< cp_<ak>_<ai> class instances >>>>>>>*) |
|
605 |
(*==============================================*) |
|
606 |
val cp1 = PureThy.get_thm thy24 (Name "cp1"); |
|
607 |
val cp_unit_inst = PureThy.get_thm thy24 (Name "cp_unit_inst"); |
|
608 |
val cp_bool_inst = PureThy.get_thm thy24 (Name "cp_bool_inst"); |
|
609 |
val cp_prod_inst = PureThy.get_thm thy24 (Name "cp_prod_inst"); |
|
610 |
val cp_list_inst = PureThy.get_thm thy24 (Name "cp_list_inst"); |
|
611 |
val cp_fun_inst = PureThy.get_thm thy24 (Name "cp_fun_inst"); |
|
612 |
val cp_option_inst = PureThy.get_thm thy24 (Name "cp_option_inst"); |
|
613 |
val cp_noption_inst = PureThy.get_thm thy24 (Name "cp_noption_inst"); |
|
614 |
val pt_perm_compose = PureThy.get_thm thy24 (Name "pt_perm_compose"); |
|
615 |
val dj_pp_forget = PureThy.get_thm thy24 (Name "dj_perm_perm_forget"); |
|
616 |
||
617 |
(* shows that <aj> is an instance of cp_<ak>_<ai> *) |
|
618 |
(* that needs a three-nested loop *) |
|
619 |
val (thy25,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
620 |
foldl_map (fn (thy', (ak_name', T')) => |
|
621 |
foldl_map (fn (thy'', (ak_name'', T'')) => |
|
622 |
let |
|
623 |
val qu_name = Sign.full_name (sign_of thy'') ak_name; |
|
624 |
val qu_class = Sign.full_name (sign_of thy'') ("cp_"^ak_name'^"_"^ak_name''); |
|
625 |
val proof = |
|
626 |
(if (ak_name'=ak_name'') then |
|
627 |
(let |
|
628 |
val pt_inst = PureThy.get_thm thy'' (Name ("pt_"^ak_name''^"_inst")); |
|
629 |
val at_inst = PureThy.get_thm thy'' (Name ("at_"^ak_name''^"_inst")); |
|
630 |
in |
|
631 |
EVERY [AxClass.intro_classes_tac [], |
|
632 |
rtac (at_inst RS (pt_inst RS pt_perm_compose)) 1] |
|
633 |
end) |
|
634 |
else |
|
635 |
(let |
|
636 |
val dj_inst = PureThy.get_thm thy'' (Name ("dj_"^ak_name''^"_"^ak_name')); |
|
637 |
val simp_s = HOL_basic_ss addsimps |
|
638 |
((dj_inst RS dj_pp_forget):: |
|
639 |
(PureThy.get_thmss thy'' |
|
640 |
[Name (ak_name' ^"_prm_"^ak_name^"_def"), |
|
641 |
Name (ak_name''^"_prm_"^ak_name^"_def")])); |
|
642 |
in |
|
643 |
EVERY [AxClass.intro_classes_tac [], simp_tac simp_s 1] |
|
644 |
end)) |
|
645 |
in |
|
646 |
(AxClass.add_inst_arity_i (qu_name,[],[qu_class]) proof thy'',()) |
|
647 |
end) |
|
648 |
(thy', ak_names_types)) (thy, ak_names_types)) (thy24, ak_names_types); |
|
649 |
||
650 |
(* shows that unit is an instance of cp_<ak>_<ai> *) |
|
651 |
(* for every <ak>-combination *) |
|
652 |
val (thy26,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
653 |
foldl_map (fn (thy', (ak_name', T')) => |
|
654 |
let |
|
655 |
val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name'); |
|
656 |
val proof = EVERY [AxClass.intro_classes_tac [],rtac (cp_unit_inst RS cp1) 1]; |
|
657 |
in |
|
658 |
(AxClass.add_inst_arity_i ("Product_Type.unit",[],[qu_class]) proof thy',()) |
|
659 |
end) |
|
660 |
(thy, ak_names_types)) (thy25, ak_names_types); |
|
661 |
||
662 |
(* shows that bool is an instance of cp_<ak>_<ai> *) |
|
663 |
(* for every <ak>-combination *) |
|
664 |
val (thy27,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
665 |
foldl_map (fn (thy', (ak_name', T')) => |
|
666 |
let |
|
667 |
val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name'); |
|
668 |
val proof = EVERY [AxClass.intro_classes_tac [], rtac (cp_bool_inst RS cp1) 1]; |
|
669 |
in |
|
670 |
(AxClass.add_inst_arity_i ("bool",[],[qu_class]) proof thy',()) |
|
671 |
end) |
|
672 |
(thy, ak_names_types)) (thy26, ak_names_types); |
|
673 |
||
674 |
(* shows that prod is an instance of cp_<ak>_<ai> *) |
|
675 |
(* for every <ak>-combination *) |
|
676 |
val (thy28,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
677 |
foldl_map (fn (thy', (ak_name', T')) => |
|
678 |
let |
|
679 |
val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name'); |
|
680 |
val cp_inst = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst")); |
|
681 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
682 |
rtac ((cp_inst RS (cp_inst RS cp_prod_inst)) RS cp1) 1]; |
|
683 |
in |
|
684 |
(AxClass.add_inst_arity_i ("*",[[qu_class],[qu_class]],[qu_class]) proof thy',()) |
|
685 |
end) |
|
686 |
(thy, ak_names_types)) (thy27, ak_names_types); |
|
687 |
||
688 |
(* shows that list is an instance of cp_<ak>_<ai> *) |
|
689 |
(* for every <ak>-combination *) |
|
690 |
val (thy29,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
691 |
foldl_map (fn (thy', (ak_name', T')) => |
|
692 |
let |
|
693 |
val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name'); |
|
694 |
val cp_inst = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst")); |
|
695 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
696 |
rtac ((cp_inst RS cp_list_inst) RS cp1) 1]; |
|
697 |
in |
|
698 |
(AxClass.add_inst_arity_i ("List.list",[[qu_class]],[qu_class]) proof thy',()) |
|
699 |
end) |
|
700 |
(thy, ak_names_types)) (thy28, ak_names_types); |
|
701 |
||
702 |
(* shows that function is an instance of cp_<ak>_<ai> *) |
|
703 |
(* for every <ak>-combination *) |
|
704 |
val (thy30,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
705 |
foldl_map (fn (thy', (ak_name', T')) => |
|
706 |
let |
|
707 |
val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name'); |
|
708 |
val cp_inst = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst")); |
|
709 |
val pt_inst = PureThy.get_thm thy' (Name ("pt_"^ak_name^"_inst")); |
|
710 |
val at_inst = PureThy.get_thm thy' (Name ("at_"^ak_name^"_inst")); |
|
711 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
712 |
rtac ((at_inst RS (pt_inst RS (cp_inst RS (cp_inst RS cp_fun_inst)))) RS cp1) 1]; |
|
713 |
in |
|
714 |
(AxClass.add_inst_arity_i ("fun",[[qu_class],[qu_class]],[qu_class]) proof thy',()) |
|
715 |
end) |
|
716 |
(thy, ak_names_types)) (thy29, ak_names_types); |
|
717 |
||
718 |
(* shows that option is an instance of cp_<ak>_<ai> *) |
|
719 |
(* for every <ak>-combination *) |
|
720 |
val (thy31,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
721 |
foldl_map (fn (thy', (ak_name', T')) => |
|
722 |
let |
|
723 |
val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name'); |
|
724 |
val cp_inst = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst")); |
|
725 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
726 |
rtac ((cp_inst RS cp_option_inst) RS cp1) 1]; |
|
727 |
in |
|
728 |
(AxClass.add_inst_arity_i ("Datatype.option",[[qu_class]],[qu_class]) proof thy',()) |
|
729 |
end) |
|
730 |
(thy, ak_names_types)) (thy30, ak_names_types); |
|
731 |
||
732 |
(* shows that nOption is an instance of cp_<ak>_<ai> *) |
|
733 |
(* for every <ak>-combination *) |
|
734 |
val (thy32,_) = foldl_map (fn (thy, (ak_name, T)) => |
|
735 |
foldl_map (fn (thy', (ak_name', T')) => |
|
736 |
let |
|
737 |
val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name'); |
|
738 |
val cp_inst = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst")); |
|
739 |
val proof = EVERY [AxClass.intro_classes_tac [], |
|
740 |
rtac ((cp_inst RS cp_noption_inst) RS cp1) 1]; |
|
741 |
in |
|
742 |
(AxClass.add_inst_arity_i ("nominal.nOption",[[qu_class]],[qu_class]) proof thy',()) |
|
743 |
end) |
|
744 |
(thy, ak_names_types)) (thy31, ak_names_types); |
|
745 |
||
746 |
(* abbreviations for some collection of rules *) |
|
747 |
(*============================================*) |
|
748 |
val abs_fun_pi = PureThy.get_thm thy32 (Name ("nominal.abs_fun_pi")); |
|
749 |
val abs_fun_pi_ineq = PureThy.get_thm thy32 (Name ("nominal.abs_fun_pi_ineq")); |
|
750 |
val abs_fun_eq = PureThy.get_thm thy32 (Name ("nominal.abs_fun_eq")); |
|
751 |
val dj_perm_forget = PureThy.get_thm thy32 (Name ("nominal.dj_perm_forget")); |
|
752 |
val dj_pp_forget = PureThy.get_thm thy32 (Name ("nominal.dj_perm_perm_forget")); |
|
753 |
val fresh_iff = PureThy.get_thm thy32 (Name ("nominal.fresh_abs_fun_iff")); |
|
754 |
val fresh_iff_ineq = PureThy.get_thm thy32 (Name ("nominal.fresh_abs_fun_iff_ineq")); |
|
755 |
val abs_fun_supp = PureThy.get_thm thy32 (Name ("nominal.abs_fun_supp")); |
|
756 |
val abs_fun_supp_ineq = PureThy.get_thm thy32 (Name ("nominal.abs_fun_supp_ineq")); |
|
757 |
val pt_swap_bij = PureThy.get_thm thy32 (Name ("nominal.pt_swap_bij")); |
|
758 |
val pt_fresh_fresh = PureThy.get_thm thy32 (Name ("nominal.pt_fresh_fresh")); |
|
759 |
val pt_bij = PureThy.get_thm thy32 (Name ("nominal.pt_bij")); |
|
760 |
val pt_perm_compose = PureThy.get_thm thy32 (Name ("nominal.pt_perm_compose")); |
|
761 |
val perm_eq_app = PureThy.get_thm thy32 (Name ("nominal.perm_eq_app")); |
|
762 |
val at_fresh = PureThy.get_thm thy32 (Name ("nominal.at_fresh")); |
|
763 |
val at_calc = PureThy.get_thms thy32 (Name ("nominal.at_calc")); |
|
764 |
val at_supp = PureThy.get_thm thy32 (Name ("nominal.at_supp")); |
|
765 |
val dj_supp = PureThy.get_thm thy32 (Name ("nominal.dj_supp")); |
|
766 |
||
767 |
(* abs_perm collects all lemmas for simplifying a permutation *) |
|
768 |
(* in front of an abs_fun *) |
|
769 |
val (thy33,_) = |
|
770 |
let |
|
771 |
val name = "abs_perm" |
|
772 |
val thm_list = Library.flat (map (fn (ak_name, T) => |
|
773 |
let |
|
774 |
val at_inst = PureThy.get_thm thy32 (Name ("at_"^ak_name^"_inst")); |
|
775 |
val pt_inst = PureThy.get_thm thy32 (Name ("pt_"^ak_name^"_inst")); |
|
776 |
val thm = [pt_inst, at_inst] MRS abs_fun_pi |
|
777 |
val thm_list = map (fn (ak_name', T') => |
|
778 |
let |
|
779 |
val cp_inst = PureThy.get_thm thy32 (Name ("cp_"^ak_name^"_"^ak_name'^"_inst")); |
|
780 |
in |
|
781 |
[pt_inst, pt_inst, at_inst, cp_inst] MRS abs_fun_pi_ineq |
|
782 |
end) ak_names_types; |
|
783 |
in thm::thm_list end) (ak_names_types)) |
|
784 |
in |
|
785 |
(PureThy.add_thmss [((name, thm_list),[])] thy32) |
|
786 |
end; |
|
787 |
||
788 |
val (thy34,_) = |
|
789 |
let |
|
790 |
(* takes a theorem and a list of theorems *) |
|
791 |
(* produces a list of theorems of the form *) |
|
792 |
(* [t1 RS thm,..,tn RS thm] where t1..tn in thms *) |
|
793 |
fun instantiate thm thms = map (fn ti => ti RS thm) thms; |
|
794 |
||
795 |
(* takes two theorem lists (hopefully of the same length) *) |
|
796 |
(* produces a list of theorems of the form *) |
|
797 |
(* [t1 RS m1,..,tn RS mn] where t1..tn in thms1 and m1..mn in thms2 *) |
|
798 |
fun instantiate_zip thms1 thms2 = |
|
799 |
map (fn (t1,t2) => t1 RS t2) (thms1 ~~ thms2); |
|
800 |
||
801 |
(* list of all at_inst-theorems *) |
|
802 |
val ats = map (fn ak => PureThy.get_thm thy33 (Name ("at_"^ak^"_inst"))) ak_names |
|
803 |
(* list of all pt_inst-theorems *) |
|
804 |
val pts = map (fn ak => PureThy.get_thm thy33 (Name ("pt_"^ak^"_inst"))) ak_names |
|
805 |
(* list of all cp_inst-theorems *) |
|
806 |
val cps = |
|
807 |
let fun cps_fun (ak1,ak2) = PureThy.get_thm thy33 (Name ("cp_"^ak1^"_"^ak2^"_inst")) |
|
18100
193c3382bbfe
used the function Library.product for the cprod from Stefan
urbanc
parents:
18068
diff
changeset
|
808 |
in map cps_fun (Library.product ak_names ak_names) end; |
18068 | 809 |
(* list of all dj_inst-theorems *) |
810 |
val djs = |
|
811 |
let fun djs_fun (ak1,ak2) = |
|
812 |
if ak1=ak2 |
|
813 |
then NONE |
|
814 |
else SOME(PureThy.get_thm thy33 (Name ("dj_"^ak1^"_"^ak2))) |
|
18100
193c3382bbfe
used the function Library.product for the cprod from Stefan
urbanc
parents:
18068
diff
changeset
|
815 |
in List.mapPartial I (map djs_fun (Library.product ak_names ak_names)) end; |
18068 | 816 |
|
817 |
fun inst_pt thms = Library.flat (map (fn ti => instantiate ti pts) thms); |
|
818 |
fun inst_at thms = Library.flat (map (fn ti => instantiate ti ats) thms); |
|
819 |
fun inst_pt_at thms = instantiate_zip ats (inst_pt thms); |
|
820 |
fun inst_dj thms = Library.flat (map (fn ti => instantiate ti djs) thms); |
|
821 |
||
822 |
in |
|
823 |
thy33 |
|
824 |
|> PureThy.add_thmss [(("alpha", inst_pt_at [abs_fun_eq]),[])] |
|
825 |
|>>> PureThy.add_thmss [(("perm_swap", inst_pt_at [pt_swap_bij]),[])] |
|
826 |
|>>> PureThy.add_thmss [(("perm_fresh_fresh", inst_pt_at [pt_fresh_fresh]),[])] |
|
827 |
|>>> PureThy.add_thmss [(("perm_bij", inst_pt_at [pt_bij]),[])] |
|
828 |
|>>> PureThy.add_thmss [(("perm_compose", inst_pt_at [pt_perm_compose]),[])] |
|
829 |
|>>> PureThy.add_thmss [(("perm_app_eq", inst_pt_at [perm_eq_app]),[])] |
|
830 |
|>>> PureThy.add_thmss [(("supp_atm", (inst_at [at_supp]) @ (inst_dj [dj_supp])),[])] |
|
831 |
|>>> PureThy.add_thmss [(("fresh_atm", inst_at [at_fresh]),[])] |
|
832 |
|>>> PureThy.add_thmss [(("calc_atm", inst_at at_calc),[])] |
|
833 |
||
834 |
end; |
|
835 |
||
836 |
(* perm_dj collects all lemmas that forget an permutation *) |
|
837 |
(* when it acts on an atom of different type *) |
|
838 |
val (thy35,_) = |
|
839 |
let |
|
840 |
val name = "perm_dj" |
|
841 |
val thm_list = Library.flat (map (fn (ak_name, T) => |
|
842 |
Library.flat (map (fn (ak_name', T') => |
|
843 |
if not (ak_name = ak_name') |
|
844 |
then |
|
845 |
let |
|
846 |
val dj_inst = PureThy.get_thm thy34 (Name ("dj_"^ak_name^"_"^ak_name')); |
|
847 |
in |
|
848 |
[dj_inst RS dj_perm_forget, dj_inst RS dj_pp_forget] |
|
849 |
end |
|
850 |
else []) ak_names_types)) ak_names_types) |
|
851 |
in |
|
852 |
(PureThy.add_thmss [((name, thm_list),[])] thy34) |
|
853 |
end; |
|
854 |
||
855 |
(* abs_fresh collects all lemmas for simplifying a freshness *) |
|
856 |
(* proposition involving an abs_fun *) |
|
857 |
||
858 |
val (thy36,_) = |
|
859 |
let |
|
860 |
val name = "abs_fresh" |
|
861 |
val thm_list = Library.flat (map (fn (ak_name, T) => |
|
862 |
let |
|
863 |
val at_inst = PureThy.get_thm thy35 (Name ("at_"^ak_name^"_inst")); |
|
864 |
val pt_inst = PureThy.get_thm thy35 (Name ("pt_"^ak_name^"_inst")); |
|
865 |
val fs_inst = PureThy.get_thm thy35 (Name ("fs_"^ak_name^"_inst")); |
|
866 |
val thm = [pt_inst, at_inst, (fs_inst RS fs1)] MRS fresh_iff |
|
867 |
val thm_list = Library.flat (map (fn (ak_name', T') => |
|
868 |
(if (not (ak_name = ak_name')) |
|
869 |
then |
|
870 |
let |
|
871 |
val cp_inst = PureThy.get_thm thy35 (Name ("cp_"^ak_name^"_"^ak_name'^"_inst")); |
|
872 |
val dj_inst = PureThy.get_thm thy35 (Name ("dj_"^ak_name'^"_"^ak_name)); |
|
873 |
in |
|
874 |
[[pt_inst, pt_inst, at_inst, cp_inst, dj_inst] MRS fresh_iff_ineq] |
|
875 |
end |
|
876 |
else [])) ak_names_types); |
|
877 |
in thm::thm_list end) (ak_names_types)) |
|
878 |
in |
|
879 |
(PureThy.add_thmss [((name, thm_list),[])] thy35) |
|
880 |
end; |
|
881 |
||
882 |
(* abs_supp collects all lemmas for simplifying *) |
|
883 |
(* support proposition involving an abs_fun *) |
|
884 |
||
885 |
val (thy37,_) = |
|
886 |
let |
|
887 |
val name = "abs_supp" |
|
888 |
val thm_list = Library.flat (map (fn (ak_name, T) => |
|
889 |
let |
|
890 |
val at_inst = PureThy.get_thm thy36 (Name ("at_"^ak_name^"_inst")); |
|
891 |
val pt_inst = PureThy.get_thm thy36 (Name ("pt_"^ak_name^"_inst")); |
|
892 |
val fs_inst = PureThy.get_thm thy36 (Name ("fs_"^ak_name^"_inst")); |
|
893 |
val thm1 = [pt_inst, at_inst, (fs_inst RS fs1)] MRS abs_fun_supp |
|
894 |
val thm2 = [pt_inst, at_inst] MRS abs_fun_supp |
|
895 |
val thm_list = Library.flat (map (fn (ak_name', T') => |
|
896 |
(if (not (ak_name = ak_name')) |
|
897 |
then |
|
898 |
let |
|
899 |
val cp_inst = PureThy.get_thm thy36 (Name ("cp_"^ak_name^"_"^ak_name'^"_inst")); |
|
900 |
val dj_inst = PureThy.get_thm thy36 (Name ("dj_"^ak_name'^"_"^ak_name)); |
|
901 |
in |
|
902 |
[[pt_inst, pt_inst, at_inst, cp_inst, dj_inst] MRS abs_fun_supp_ineq] |
|
903 |
end |
|
904 |
else [])) ak_names_types); |
|
905 |
in thm1::thm2::thm_list end) (ak_names_types)) |
|
906 |
in |
|
907 |
(PureThy.add_thmss [((name, thm_list),[])] thy36) |
|
908 |
end; |
|
909 |
||
910 |
in NominalData.put (fold Symtab.update (map (rpair ()) full_ak_names) |
|
911 |
(NominalData.get thy11)) thy37 |
|
912 |
end; |
|
913 |
||
914 |
||
915 |
(* syntax und parsing *) |
|
916 |
structure P = OuterParse and K = OuterKeyword; |
|
917 |
||
918 |
val atom_declP = |
|
919 |
OuterSyntax.command "atom_decl" "Declare new kinds of atoms" K.thy_decl |
|
920 |
(Scan.repeat1 P.name >> (Toplevel.theory o create_nom_typedecls)); |
|
921 |
||
922 |
val _ = OuterSyntax.add_parsers [atom_declP]; |
|
923 |
||
924 |
val setup = [NominalData.init]; |
|
925 |
||
926 |
end; |