author | urbanc |
Wed, 11 Jan 2006 12:21:01 +0100 | |
changeset 18652 | 3930a060d71b |
parent 18651 | 0cb5a8f501aa |
child 18707 | 9d6154f76476 |
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) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
59 |
val (inf_axs,thy2) = PureThy.add_axioms_i (map (fn (ak_name, T) => |
18068 | 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)] |
|
18366 | 92 |
|> (#2 o PureThy.add_defs_i true [((name, def2),[])]) |
18068 | 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 *) |
|
18366 | 133 |
val (perm_defs, thy5) = fold_map (fn (ak_name, T) => fn thy => |
134 |
fold_map (fn (ak_name', T') => fn thy' => |
|
18068 | 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 |
|
18366 | 146 |
PureThy.add_defs_i true [((name, def),[])] thy' |
147 |
end) ak_names_types thy) ak_names_types thy4; |
|
18068 | 148 |
|
149 |
(* proves that every atom-kind is an instance of at *) |
|
150 |
(* lemma at_<ak>_inst: *) |
|
151 |
(* at TYPE(<ak>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
152 |
val (prm_cons_thms,thy6) = |
18068 | 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 *) |
|
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
181 |
val (pt_ax_classes,thy7) = fold_map (fn (ak_name, T) => fn thy => |
18068 | 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 |
|
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
203 |
AxClass.add_axclass_i (cl_name, ["HOL.type"]) |
18068 | 204 |
[((cl_name^"1", axiom1),[Simplifier.simp_add_global]), |
205 |
((cl_name^"2", axiom2),[]), |
|
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
206 |
((cl_name^"3", axiom3),[])] thy |
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
207 |
end) ak_names_types thy6; |
18068 | 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>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
213 |
val (prm_inst_thms,thy8) = |
18068 | 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) *) |
|
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
240 |
val (fs_ax_classes,thy11) = fold_map (fn (ak_name, T) => fn thy => |
18068 | 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 |
|
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
252 |
AxClass.add_axclass_i (cl_name, [pt_name]) [((cl_name^"1", axiom1),[])] thy |
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
253 |
end) ak_names_types thy8; |
18068 | 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>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
259 |
val (fs_inst_thms,thy12) = |
18068 | 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) *) |
|
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
285 |
val (_,thy12b) = fold_map (fn (ak_name, T) => fn thy => |
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
286 |
fold_map (fn (ak_name', T') => fn thy' => |
18068 | 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 |
|
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
301 |
AxClass.add_axclass_i (cl_name, ["HOL.type"]) [((cl_name^"1", ax1),[])] thy' |
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
302 |
end) ak_names_types thy) ak_names_types thy12; |
18068 | 303 |
|
304 |
(* proves for every <ak>-combination a cp_<ak1>_<ak2>_inst theorem; *) |
|
305 |
(* lemma cp_<ak1>_<ak2>_inst: *) |
|
306 |
(* cp TYPE('a::cp_<ak1>_<ak2>) TYPE(<ak1>) TYPE(<ak2>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
307 |
val (cp_thms,thy12c) = fold_map (fn (ak_name, T) => fn thy => |
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
308 |
fold_map (fn (ak_name', T') => fn thy' => |
18068 | 309 |
let |
310 |
val ak_name_qu = Sign.full_name (sign_of thy') (ak_name); |
|
311 |
val ak_name_qu' = Sign.full_name (sign_of thy') (ak_name'); |
|
312 |
val cp_name_qu = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name'); |
|
313 |
val i_type0 = TFree("'a",[cp_name_qu]); |
|
314 |
val i_type1 = Type(ak_name_qu,[]); |
|
315 |
val i_type2 = Type(ak_name_qu',[]); |
|
316 |
val ccp = Const ("nominal.cp", |
|
317 |
(Term.itselfT i_type0)-->(Term.itselfT i_type1)--> |
|
318 |
(Term.itselfT i_type2)-->HOLogic.boolT); |
|
319 |
val at_type = Logic.mk_type i_type1; |
|
320 |
val at_type' = Logic.mk_type i_type2; |
|
321 |
val cp_type = Logic.mk_type i_type0; |
|
322 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy' [(Name "cp_def")]; |
|
323 |
val cp1 = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"1")); |
|
324 |
||
325 |
val name = "cp_"^ak_name^ "_"^ak_name'^"_inst"; |
|
326 |
val statement = HOLogic.mk_Trueprop (ccp $ cp_type $ at_type $ at_type'); |
|
327 |
||
328 |
val proof = fn _ => EVERY [auto_tac (claset(),simp_s), rtac cp1 1]; |
|
329 |
in |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
330 |
PureThy.add_thms [((name, standard (Goal.prove thy' [] [] statement proof)), [])] thy' |
18068 | 331 |
end) |
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
332 |
ak_names_types thy) ak_names_types thy12b; |
18068 | 333 |
|
334 |
(* proves for every non-trivial <ak>-combination a disjointness *) |
|
335 |
(* theorem; i.e. <ak1> != <ak2> *) |
|
336 |
(* lemma ds_<ak1>_<ak2>: *) |
|
337 |
(* dj TYPE(<ak1>) TYPE(<ak2>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
338 |
val (dj_thms, thy12d) = fold_map (fn (ak_name,T) => fn thy => |
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
339 |
fold_map (fn (ak_name',T') => fn thy' => |
18068 | 340 |
(if not (ak_name = ak_name') |
341 |
then |
|
342 |
let |
|
343 |
val ak_name_qu = Sign.full_name (sign_of thy') (ak_name); |
|
344 |
val ak_name_qu' = Sign.full_name (sign_of thy') (ak_name'); |
|
345 |
val i_type1 = Type(ak_name_qu,[]); |
|
346 |
val i_type2 = Type(ak_name_qu',[]); |
|
347 |
val cdj = Const ("nominal.disjoint", |
|
348 |
(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); |
|
349 |
val at_type = Logic.mk_type i_type1; |
|
350 |
val at_type' = Logic.mk_type i_type2; |
|
351 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy' |
|
352 |
[Name "disjoint_def", |
|
353 |
Name (ak_name^"_prm_"^ak_name'^"_def"), |
|
354 |
Name (ak_name'^"_prm_"^ak_name^"_def")]; |
|
355 |
||
356 |
val name = "dj_"^ak_name^"_"^ak_name'; |
|
357 |
val statement = HOLogic.mk_Trueprop (cdj $ at_type $ at_type'); |
|
358 |
||
359 |
val proof = fn _ => auto_tac (claset(),simp_s); |
|
360 |
in |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
361 |
PureThy.add_thms [((name, standard (Goal.prove thy' [] [] statement proof)), [])] thy' |
18068 | 362 |
end |
363 |
else |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
364 |
([],thy'))) (* do nothing branch, if ak_name = ak_name' *) |
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
365 |
ak_names_types thy) ak_names_types thy12c; |
18068 | 366 |
|
367 |
(*<<<<<<< pt_<ak> class instances >>>>>>>*) |
|
368 |
(*=========================================*) |
|
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
369 |
(* some abbreviations for theorems *) |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
370 |
val pt1 = thm "pt1"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
371 |
val pt2 = thm "pt2"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
372 |
val pt3 = thm "pt3"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
373 |
val at_pt_inst = thm "at_pt_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
374 |
val pt_set_inst = thm "pt_set_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
375 |
val pt_unit_inst = thm "pt_unit_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
376 |
val pt_prod_inst = thm "pt_prod_inst"; |
18600 | 377 |
val pt_nprod_inst = thm "pt_nprod_inst"; |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
378 |
val pt_list_inst = thm "pt_list_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
379 |
val pt_optn_inst = thm "pt_option_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
380 |
val pt_noptn_inst = thm "pt_noption_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
381 |
val pt_fun_inst = thm "pt_fun_inst"; |
18068 | 382 |
|
18435 | 383 |
(* for all atom-kind combinations <ak>/<ak'> show that *) |
384 |
(* every <ak> is an instance of pt_<ak'>; the proof for *) |
|
385 |
(* ak!=ak' is by definition; the case ak=ak' uses at_pt_inst. *) |
|
18431 | 386 |
val thy13 = fold (fn ak_name => fn thy => |
387 |
fold (fn ak_name' => fn thy' => |
|
388 |
let |
|
389 |
val qu_name = Sign.full_name (sign_of thy') ak_name'; |
|
390 |
val cls_name = Sign.full_name (sign_of thy') ("pt_"^ak_name); |
|
391 |
val at_inst = PureThy.get_thm thy' (Name ("at_"^ak_name'^"_inst")); |
|
392 |
||
393 |
val proof1 = EVERY [AxClass.intro_classes_tac [], |
|
18068 | 394 |
rtac ((at_inst RS at_pt_inst) RS pt1) 1, |
395 |
rtac ((at_inst RS at_pt_inst) RS pt2) 1, |
|
396 |
rtac ((at_inst RS at_pt_inst) RS pt3) 1, |
|
397 |
atac 1]; |
|
18431 | 398 |
val simp_s = HOL_basic_ss addsimps |
399 |
PureThy.get_thmss thy' [Name (ak_name^"_prm_"^ak_name'^"_def")]; |
|
400 |
val proof2 = EVERY [AxClass.intro_classes_tac [], REPEAT (asm_simp_tac simp_s 1)]; |
|
401 |
||
402 |
in |
|
403 |
thy' |
|
404 |
|> AxClass.add_inst_arity_i (qu_name,[],[cls_name]) |
|
405 |
(if ak_name = ak_name' then proof1 else proof2) |
|
406 |
end) ak_names thy) ak_names thy12c; |
|
18068 | 407 |
|
18430 | 408 |
(* show that *) |
409 |
(* fun(pt_<ak>,pt_<ak>) *) |
|
18579
002d371401f5
changed the name of the type "nOption" to "noption".
urbanc
parents:
18438
diff
changeset
|
410 |
(* noption(pt_<ak>) *) |
18430 | 411 |
(* option(pt_<ak>) *) |
412 |
(* list(pt_<ak>) *) |
|
413 |
(* *(pt_<ak>,pt_<ak>) *) |
|
18600 | 414 |
(* nprod(pt_<ak>,pt_<ak>) *) |
18430 | 415 |
(* unit *) |
416 |
(* set(pt_<ak>) *) |
|
417 |
(* are instances of pt_<ak> *) |
|
18431 | 418 |
val thy18 = fold (fn ak_name => fn thy => |
18068 | 419 |
let |
18430 | 420 |
val cls_name = Sign.full_name (sign_of thy) ("pt_"^ak_name); |
18068 | 421 |
val at_thm = PureThy.get_thm thy (Name ("at_"^ak_name^"_inst")); |
422 |
val pt_inst = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst")); |
|
18430 | 423 |
|
424 |
fun pt_proof thm = |
|
425 |
EVERY [AxClass.intro_classes_tac [], |
|
426 |
rtac (thm RS pt1) 1, rtac (thm RS pt2) 1, rtac (thm RS pt3) 1, atac 1]; |
|
427 |
||
428 |
val pt_thm_fun = at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst)); |
|
429 |
val pt_thm_noptn = pt_inst RS pt_noptn_inst; |
|
430 |
val pt_thm_optn = pt_inst RS pt_optn_inst; |
|
431 |
val pt_thm_list = pt_inst RS pt_list_inst; |
|
432 |
val pt_thm_prod = pt_inst RS (pt_inst RS pt_prod_inst); |
|
18600 | 433 |
val pt_thm_nprod = pt_inst RS (pt_inst RS pt_nprod_inst); |
18430 | 434 |
val pt_thm_unit = pt_unit_inst; |
435 |
val pt_thm_set = pt_inst RS pt_set_inst |
|
18068 | 436 |
in |
18430 | 437 |
thy |
438 |
|> AxClass.add_inst_arity_i ("fun",[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_fun) |
|
18579
002d371401f5
changed the name of the type "nOption" to "noption".
urbanc
parents:
18438
diff
changeset
|
439 |
|> AxClass.add_inst_arity_i ("nominal.noption",[[cls_name]],[cls_name]) (pt_proof pt_thm_noptn) |
18430 | 440 |
|> AxClass.add_inst_arity_i ("Datatype.option",[[cls_name]],[cls_name]) (pt_proof pt_thm_optn) |
441 |
|> AxClass.add_inst_arity_i ("List.list",[[cls_name]],[cls_name]) (pt_proof pt_thm_list) |
|
442 |
|> AxClass.add_inst_arity_i ("*",[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_prod) |
|
18600 | 443 |
|> AxClass.add_inst_arity_i ("nominal.nprod",[[cls_name],[cls_name]],[cls_name]) |
444 |
(pt_proof pt_thm_nprod) |
|
18430 | 445 |
|> AxClass.add_inst_arity_i ("Product_Type.unit",[],[cls_name]) (pt_proof pt_thm_unit) |
446 |
|> AxClass.add_inst_arity_i ("set",[[cls_name]],[cls_name]) (pt_proof pt_thm_set) |
|
447 |
end) ak_names thy13; |
|
18068 | 448 |
|
449 |
(*<<<<<<< fs_<ak> class instances >>>>>>>*) |
|
450 |
(*=========================================*) |
|
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
451 |
(* abbreviations for some lemmas *) |
18431 | 452 |
val fs1 = thm "fs1"; |
453 |
val fs_at_inst = thm "fs_at_inst"; |
|
454 |
val fs_unit_inst = thm "fs_unit_inst"; |
|
455 |
val fs_prod_inst = thm "fs_prod_inst"; |
|
18600 | 456 |
val fs_nprod_inst = thm "fs_nprod_inst"; |
18431 | 457 |
val fs_list_inst = thm "fs_list_inst"; |
458 |
val fs_option_inst = thm "fs_option_inst"; |
|
18437
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
459 |
val dj_supp = thm "dj_supp" |
18068 | 460 |
|
461 |
(* shows that <ak> is an instance of fs_<ak> *) |
|
462 |
(* uses the theorem at_<ak>_inst *) |
|
18431 | 463 |
val thy20 = fold (fn ak_name => fn thy => |
18437
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
464 |
fold (fn ak_name' => fn thy' => |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
465 |
let |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
466 |
val qu_name = Sign.full_name (sign_of thy') ak_name'; |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
467 |
val qu_class = Sign.full_name (sign_of thy') ("fs_"^ak_name); |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
468 |
val proof = |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
469 |
(if ak_name = ak_name' |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
470 |
then |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
471 |
let val at_thm = PureThy.get_thm thy' (Name ("at_"^ak_name^"_inst")); |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
472 |
in EVERY [AxClass.intro_classes_tac [], |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
473 |
rtac ((at_thm RS fs_at_inst) RS fs1) 1] end |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
474 |
else |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
475 |
let val dj_inst = PureThy.get_thm thy' (Name ("dj_"^ak_name'^"_"^ak_name)); |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
476 |
val simp_s = HOL_basic_ss addsimps [dj_inst RS dj_supp, Finites.emptyI]; |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
477 |
in EVERY [AxClass.intro_classes_tac [], asm_simp_tac simp_s 1] end) |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
478 |
in |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
479 |
AxClass.add_inst_arity_i (qu_name,[],[qu_class]) proof thy' |
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
480 |
end) ak_names thy) ak_names thy18; |
18068 | 481 |
|
18431 | 482 |
(* shows that *) |
483 |
(* unit *) |
|
484 |
(* *(fs_<ak>,fs_<ak>) *) |
|
18600 | 485 |
(* nprod(fs_<ak>,fs_<ak>) *) |
18431 | 486 |
(* list(fs_<ak>) *) |
487 |
(* option(fs_<ak>) *) |
|
488 |
(* are instances of fs_<ak> *) |
|
18068 | 489 |
|
18431 | 490 |
val thy24 = fold (fn ak_name => fn thy => |
491 |
let |
|
492 |
val cls_name = Sign.full_name (sign_of thy) ("fs_"^ak_name); |
|
18068 | 493 |
val fs_inst = PureThy.get_thm thy (Name ("fs_"^ak_name^"_inst")); |
18431 | 494 |
fun fs_proof thm = EVERY [AxClass.intro_classes_tac [], rtac (thm RS fs1) 1]; |
18068 | 495 |
|
18600 | 496 |
val fs_thm_unit = fs_unit_inst; |
497 |
val fs_thm_prod = fs_inst RS (fs_inst RS fs_prod_inst); |
|
498 |
val fs_thm_nprod = fs_inst RS (fs_inst RS fs_nprod_inst); |
|
499 |
val fs_thm_list = fs_inst RS fs_list_inst; |
|
500 |
val fs_thm_optn = fs_inst RS fs_option_inst; |
|
18431 | 501 |
in |
502 |
thy |
|
503 |
|> AxClass.add_inst_arity_i ("Product_Type.unit",[],[cls_name]) (fs_proof fs_thm_unit) |
|
18600 | 504 |
|> AxClass.add_inst_arity_i ("*",[[cls_name],[cls_name]],[cls_name]) (fs_proof fs_thm_prod) |
505 |
|> AxClass.add_inst_arity_i ("nominal.nprod",[[cls_name],[cls_name]],[cls_name]) |
|
506 |
(fs_proof fs_thm_nprod) |
|
18431 | 507 |
|> AxClass.add_inst_arity_i ("List.list",[[cls_name]],[cls_name]) (fs_proof fs_thm_list) |
508 |
|> AxClass.add_inst_arity_i ("Datatype.option",[[cls_name]],[cls_name]) (fs_proof fs_thm_optn) |
|
509 |
end) ak_names thy20; |
|
510 |
||
18068 | 511 |
(*<<<<<<< cp_<ak>_<ai> class instances >>>>>>>*) |
512 |
(*==============================================*) |
|
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
513 |
(* abbreviations for some lemmas *) |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
514 |
val cp1 = thm "cp1"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
515 |
val cp_unit_inst = thm "cp_unit_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
516 |
val cp_bool_inst = thm "cp_bool_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
517 |
val cp_prod_inst = thm "cp_prod_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
518 |
val cp_list_inst = thm "cp_list_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
519 |
val cp_fun_inst = thm "cp_fun_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
520 |
val cp_option_inst = thm "cp_option_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
521 |
val cp_noption_inst = thm "cp_noption_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
522 |
val pt_perm_compose = thm "pt_perm_compose"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
523 |
val dj_pp_forget = thm "dj_perm_perm_forget"; |
18068 | 524 |
|
525 |
(* shows that <aj> is an instance of cp_<ak>_<ai> *) |
|
18432 | 526 |
(* for every <ak>/<ai>-combination *) |
527 |
val thy25 = fold (fn ak_name => fn thy => |
|
528 |
fold (fn ak_name' => fn thy' => |
|
529 |
fold (fn ak_name'' => fn thy'' => |
|
18068 | 530 |
let |
18432 | 531 |
val name = Sign.full_name (sign_of thy'') ak_name; |
532 |
val cls_name = Sign.full_name (sign_of thy'') ("cp_"^ak_name'^"_"^ak_name''); |
|
18068 | 533 |
val proof = |
534 |
(if (ak_name'=ak_name'') then |
|
535 |
(let |
|
536 |
val pt_inst = PureThy.get_thm thy'' (Name ("pt_"^ak_name''^"_inst")); |
|
537 |
val at_inst = PureThy.get_thm thy'' (Name ("at_"^ak_name''^"_inst")); |
|
538 |
in |
|
539 |
EVERY [AxClass.intro_classes_tac [], |
|
540 |
rtac (at_inst RS (pt_inst RS pt_perm_compose)) 1] |
|
541 |
end) |
|
542 |
else |
|
543 |
(let |
|
544 |
val dj_inst = PureThy.get_thm thy'' (Name ("dj_"^ak_name''^"_"^ak_name')); |
|
545 |
val simp_s = HOL_basic_ss addsimps |
|
546 |
((dj_inst RS dj_pp_forget):: |
|
547 |
(PureThy.get_thmss thy'' |
|
548 |
[Name (ak_name' ^"_prm_"^ak_name^"_def"), |
|
549 |
Name (ak_name''^"_prm_"^ak_name^"_def")])); |
|
550 |
in |
|
551 |
EVERY [AxClass.intro_classes_tac [], simp_tac simp_s 1] |
|
552 |
end)) |
|
553 |
in |
|
18432 | 554 |
AxClass.add_inst_arity_i (name,[],[cls_name]) proof thy'' |
555 |
end) ak_names thy') ak_names thy) ak_names thy24; |
|
18068 | 556 |
|
18432 | 557 |
(* shows that *) |
558 |
(* units *) |
|
559 |
(* products *) |
|
560 |
(* lists *) |
|
561 |
(* functions *) |
|
562 |
(* options *) |
|
563 |
(* noptions *) |
|
564 |
(* are instances of cp_<ak>_<ai> for every <ak>/<ai>-combination *) |
|
565 |
val thy26 = fold (fn ak_name => fn thy => |
|
566 |
fold (fn ak_name' => fn thy' => |
|
567 |
let |
|
568 |
val cls_name = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name'); |
|
18068 | 569 |
val cp_inst = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst")); |
570 |
val pt_inst = PureThy.get_thm thy' (Name ("pt_"^ak_name^"_inst")); |
|
571 |
val at_inst = PureThy.get_thm thy' (Name ("at_"^ak_name^"_inst")); |
|
18432 | 572 |
|
573 |
fun cp_proof thm = EVERY [AxClass.intro_classes_tac [],rtac (thm RS cp1) 1]; |
|
574 |
||
575 |
val cp_thm_unit = cp_unit_inst; |
|
576 |
val cp_thm_prod = cp_inst RS (cp_inst RS cp_prod_inst); |
|
577 |
val cp_thm_list = cp_inst RS cp_list_inst; |
|
578 |
val cp_thm_fun = at_inst RS (pt_inst RS (cp_inst RS (cp_inst RS cp_fun_inst))); |
|
579 |
val cp_thm_optn = cp_inst RS cp_option_inst; |
|
580 |
val cp_thm_noptn = cp_inst RS cp_noption_inst; |
|
581 |
in |
|
582 |
thy' |
|
583 |
|> AxClass.add_inst_arity_i ("Product_Type.unit",[],[cls_name]) (cp_proof cp_thm_unit) |
|
584 |
|> AxClass.add_inst_arity_i ("*",[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_prod) |
|
585 |
|> AxClass.add_inst_arity_i ("List.list",[[cls_name]],[cls_name]) (cp_proof cp_thm_list) |
|
586 |
|> AxClass.add_inst_arity_i ("fun",[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_fun) |
|
587 |
|> AxClass.add_inst_arity_i ("Datatype.option",[[cls_name]],[cls_name]) (cp_proof cp_thm_optn) |
|
18579
002d371401f5
changed the name of the type "nOption" to "noption".
urbanc
parents:
18438
diff
changeset
|
588 |
|> AxClass.add_inst_arity_i ("nominal.noption",[[cls_name]],[cls_name]) (cp_proof cp_thm_noptn) |
18432 | 589 |
end) ak_names thy) ak_names thy25; |
590 |
||
591 |
(* show that discrete nominal types are permutation types, finitely *) |
|
592 |
(* supported and have the commutation property *) |
|
593 |
(* discrete types have a permutation operation defined as pi o x = x; *) |
|
594 |
(* which renders the proofs to be simple "simp_all"-proofs. *) |
|
595 |
val thy32 = |
|
596 |
let |
|
597 |
fun discrete_pt_inst discrete_ty defn = |
|
598 |
fold (fn ak_name => fn thy => |
|
599 |
let |
|
600 |
val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name); |
|
601 |
val simp_s = HOL_basic_ss addsimps [defn]; |
|
602 |
val proof = EVERY [AxClass.intro_classes_tac [], REPEAT (asm_simp_tac simp_s 1)]; |
|
603 |
in |
|
604 |
AxClass.add_inst_arity_i (discrete_ty,[],[qu_class]) proof thy |
|
605 |
end) ak_names; |
|
18068 | 606 |
|
18432 | 607 |
fun discrete_fs_inst discrete_ty defn = |
608 |
fold (fn ak_name => fn thy => |
|
609 |
let |
|
610 |
val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name); |
|
611 |
val supp_def = thm "nominal.supp_def"; |
|
612 |
val simp_s = HOL_ss addsimps [supp_def,Collect_const,Finites.emptyI,defn]; |
|
613 |
val proof = EVERY [AxClass.intro_classes_tac [], asm_simp_tac simp_s 1]; |
|
614 |
in |
|
615 |
AxClass.add_inst_arity_i (discrete_ty,[],[qu_class]) proof thy |
|
616 |
end) ak_names; |
|
18068 | 617 |
|
18432 | 618 |
fun discrete_cp_inst discrete_ty defn = |
619 |
fold (fn ak_name' => (fold (fn ak_name => fn thy => |
|
620 |
let |
|
621 |
val qu_class = Sign.full_name (sign_of thy) ("cp_"^ak_name^"_"^ak_name'); |
|
622 |
val supp_def = thm "nominal.supp_def"; |
|
623 |
val simp_s = HOL_ss addsimps [defn]; |
|
624 |
val proof = EVERY [AxClass.intro_classes_tac [], asm_simp_tac simp_s 1]; |
|
625 |
in |
|
626 |
AxClass.add_inst_arity_i (discrete_ty,[],[qu_class]) proof thy |
|
627 |
end) ak_names)) ak_names; |
|
628 |
||
629 |
in |
|
630 |
thy26 |
|
631 |
|> discrete_pt_inst "nat" (thm "perm_nat_def") |
|
632 |
|> discrete_fs_inst "nat" (thm "perm_nat_def") |
|
633 |
|> discrete_cp_inst "nat" (thm "perm_nat_def") |
|
634 |
|> discrete_pt_inst "bool" (thm "perm_bool") |
|
635 |
|> discrete_fs_inst "bool" (thm "perm_bool") |
|
636 |
|> discrete_cp_inst "bool" (thm "perm_bool") |
|
637 |
|> discrete_pt_inst "IntDef.int" (thm "perm_int_def") |
|
638 |
|> discrete_fs_inst "IntDef.int" (thm "perm_int_def") |
|
639 |
|> discrete_cp_inst "IntDef.int" (thm "perm_int_def") |
|
640 |
|> discrete_pt_inst "List.char" (thm "perm_char_def") |
|
641 |
|> discrete_fs_inst "List.char" (thm "perm_char_def") |
|
642 |
|> discrete_cp_inst "List.char" (thm "perm_char_def") |
|
643 |
end; |
|
644 |
||
18068 | 645 |
|
18262 | 646 |
(* abbreviations for some lemmas *) |
647 |
(*===============================*) |
|
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
648 |
val abs_fun_pi = thm "nominal.abs_fun_pi"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
649 |
val abs_fun_pi_ineq = thm "nominal.abs_fun_pi_ineq"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
650 |
val abs_fun_eq = thm "nominal.abs_fun_eq"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
651 |
val dj_perm_forget = thm "nominal.dj_perm_forget"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
652 |
val dj_pp_forget = thm "nominal.dj_perm_perm_forget"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
653 |
val fresh_iff = thm "nominal.fresh_abs_fun_iff"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
654 |
val fresh_iff_ineq = thm "nominal.fresh_abs_fun_iff_ineq"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
655 |
val abs_fun_supp = thm "nominal.abs_fun_supp"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
656 |
val abs_fun_supp_ineq = thm "nominal.abs_fun_supp_ineq"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
657 |
val pt_swap_bij = thm "nominal.pt_swap_bij"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
658 |
val pt_fresh_fresh = thm "nominal.pt_fresh_fresh"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
659 |
val pt_bij = thm "nominal.pt_bij"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
660 |
val pt_perm_compose = thm "nominal.pt_perm_compose"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
661 |
val perm_eq_app = thm "nominal.perm_eq_app"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
662 |
val at_fresh = thm "nominal.at_fresh"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
663 |
val at_calc = thms "nominal.at_calc"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
664 |
val at_supp = thm "nominal.at_supp"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
665 |
val dj_supp = thm "nominal.dj_supp"; |
18396
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
666 |
val fresh_left_ineq = thm "nominal.pt_fresh_left_ineq"; |
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
667 |
val fresh_left = thm "nominal.pt_fresh_left"; |
18426 | 668 |
val fresh_bij_ineq = thm "nominal.pt_fresh_bij_ineq"; |
669 |
val fresh_bij = thm "nominal.pt_fresh_bij"; |
|
18068 | 670 |
|
18262 | 671 |
(* Now we collect and instantiate some lemmas w.r.t. all atom *) |
672 |
(* types; this allows for example to use abs_perm (which is a *) |
|
673 |
(* collection of theorems) instead of thm abs_fun_pi with explicit *) |
|
674 |
(* instantiations. *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
675 |
val (_,thy33) = |
18262 | 676 |
let |
18651
0cb5a8f501aa
added the thms-collection "pt_id" (collection of all pt_<ak>1 lemmas)
urbanc
parents:
18626
diff
changeset
|
677 |
|
0cb5a8f501aa
added the thms-collection "pt_id" (collection of all pt_<ak>1 lemmas)
urbanc
parents:
18626
diff
changeset
|
678 |
|
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
679 |
(* takes a theorem thm and a list of theorems [t1,..,tn] *) |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
680 |
(* produces a list of theorems of the form [t1 RS thm,..,tn RS thm] *) |
18262 | 681 |
fun instR thm thms = map (fn ti => ti RS thm) thms; |
18068 | 682 |
|
18262 | 683 |
(* takes two theorem lists (hopefully of the same length ;o) *) |
684 |
(* produces a list of theorems of the form *) |
|
685 |
(* [t1 RS m1,..,tn RS mn] where [t1,..,tn] is thms1 and [m1,..,mn] is thms2 *) |
|
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
686 |
fun inst_zip thms1 thms2 = map (fn (t1,t2) => t1 RS t2) (thms1 ~~ thms2); |
18068 | 687 |
|
18262 | 688 |
(* takes a theorem list of the form [l1,...,ln] *) |
689 |
(* and a list of theorem lists of the form *) |
|
690 |
(* [[h11,...,h1m],....,[hk1,....,hkm] *) |
|
691 |
(* produces the list of theorem lists *) |
|
692 |
(* [[l1 RS h11,...,l1 RS h1m],...,[ln RS hk1,...,ln RS hkm]] *) |
|
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
693 |
fun inst_mult thms thmss = map (fn (t,ts) => instR t ts) (thms ~~ thmss); |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
694 |
|
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
695 |
(* FIXME: these lists do not need to be created dynamically again *) |
18262 | 696 |
|
18068 | 697 |
(* list of all at_inst-theorems *) |
18262 | 698 |
val ats = map (fn ak => PureThy.get_thm thy32 (Name ("at_"^ak^"_inst"))) ak_names |
18068 | 699 |
(* list of all pt_inst-theorems *) |
18262 | 700 |
val pts = map (fn ak => PureThy.get_thm thy32 (Name ("pt_"^ak^"_inst"))) ak_names |
701 |
(* list of all cp_inst-theorems as a collection of lists*) |
|
18068 | 702 |
val cps = |
18262 | 703 |
let fun cps_fun ak1 ak2 = PureThy.get_thm thy32 (Name ("cp_"^ak1^"_"^ak2^"_inst")) |
704 |
in map (fn aki => (map (cps_fun aki) ak_names)) ak_names end; |
|
705 |
(* list of all cp_inst-theorems that have different atom types *) |
|
706 |
val cps' = |
|
707 |
let fun cps'_fun ak1 ak2 = |
|
708 |
if ak1=ak2 then NONE else SOME(PureThy.get_thm thy32 (Name ("cp_"^ak1^"_"^ak2^"_inst"))) |
|
709 |
in map (fn aki => (List.mapPartial I (map (cps'_fun aki) ak_names))) ak_names end; |
|
18068 | 710 |
(* list of all dj_inst-theorems *) |
711 |
val djs = |
|
712 |
let fun djs_fun (ak1,ak2) = |
|
18262 | 713 |
if ak1=ak2 then NONE else SOME(PureThy.get_thm thy32 (Name ("dj_"^ak2^"_"^ak1))) |
714 |
in List.mapPartial I (map djs_fun (Library.product ak_names ak_names)) end; |
|
715 |
(* list of all fs_inst-theorems *) |
|
716 |
val fss = map (fn ak => PureThy.get_thm thy32 (Name ("fs_"^ak^"_inst"))) ak_names |
|
18651
0cb5a8f501aa
added the thms-collection "pt_id" (collection of all pt_<ak>1 lemmas)
urbanc
parents:
18626
diff
changeset
|
717 |
|
18262 | 718 |
fun inst_pt thms = Library.flat (map (fn ti => instR ti pts) thms); |
719 |
fun inst_at thms = Library.flat (map (fn ti => instR ti ats) thms); |
|
720 |
fun inst_fs thms = Library.flat (map (fn ti => instR ti fss) thms); |
|
18436
9649e24bc10e
added thms to perm_compose (so far only composition
urbanc
parents:
18435
diff
changeset
|
721 |
fun inst_cp thms cps = Library.flat (inst_mult thms cps); |
18262 | 722 |
fun inst_pt_at thms = inst_zip ats (inst_pt thms); |
723 |
fun inst_dj thms = Library.flat (map (fn ti => instR ti djs) thms); |
|
18436
9649e24bc10e
added thms to perm_compose (so far only composition
urbanc
parents:
18435
diff
changeset
|
724 |
fun inst_pt_pt_at_cp thms = inst_cp (inst_zip ats (inst_zip pts (inst_pt thms))) cps; |
18262 | 725 |
fun inst_pt_at_fs thms = inst_zip (inst_fs [fs1]) (inst_zip ats (inst_pt thms)); |
18396
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
726 |
fun inst_pt_pt_at_cp thms = |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
727 |
let val i_pt_pt_at = inst_zip ats (inst_zip pts (inst_pt thms)); |
18436
9649e24bc10e
added thms to perm_compose (so far only composition
urbanc
parents:
18435
diff
changeset
|
728 |
val i_pt_pt_at_cp = inst_cp i_pt_pt_at cps'; |
18396
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
729 |
in i_pt_pt_at_cp end; |
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
730 |
fun inst_pt_pt_at_cp_dj thms = inst_zip djs (inst_pt_pt_at_cp thms); |
18068 | 731 |
in |
18262 | 732 |
thy32 |
18652
3930a060d71b
rolled back the last addition since these lemmas were already
urbanc
parents:
18651
diff
changeset
|
733 |
|> PureThy.add_thmss [(("alpha", inst_pt_at [abs_fun_eq]),[])] |
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
734 |
||>> PureThy.add_thmss [(("perm_swap", inst_pt_at [pt_swap_bij]),[])] |
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
735 |
||>> PureThy.add_thmss [(("perm_fresh_fresh", inst_pt_at [pt_fresh_fresh]),[])] |
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
736 |
||>> PureThy.add_thmss [(("perm_bij", inst_pt_at [pt_bij]),[])] |
18436
9649e24bc10e
added thms to perm_compose (so far only composition
urbanc
parents:
18435
diff
changeset
|
737 |
||>> PureThy.add_thmss |
9649e24bc10e
added thms to perm_compose (so far only composition
urbanc
parents:
18435
diff
changeset
|
738 |
let val thms1 = inst_pt_at [pt_perm_compose]; |
9649e24bc10e
added thms to perm_compose (so far only composition
urbanc
parents:
18435
diff
changeset
|
739 |
val thms2 = instR cp1 (Library.flat cps'); |
9649e24bc10e
added thms to perm_compose (so far only composition
urbanc
parents:
18435
diff
changeset
|
740 |
in [(("perm_compose",thms1 @ thms2),[])] end |
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
741 |
||>> PureThy.add_thmss [(("perm_app_eq", inst_pt_at [perm_eq_app]),[])] |
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
742 |
||>> PureThy.add_thmss [(("supp_atm", (inst_at [at_supp]) @ (inst_dj [dj_supp])),[])] |
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
743 |
||>> PureThy.add_thmss [(("fresh_atm", inst_at [at_fresh]),[])] |
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
744 |
||>> PureThy.add_thmss [(("calc_atm", inst_at at_calc),[])] |
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
745 |
||>> PureThy.add_thmss |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
746 |
let val thms1 = inst_pt_at [abs_fun_pi] |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
747 |
and thms2 = inst_pt_pt_at_cp [abs_fun_pi_ineq] |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
748 |
in [(("abs_perm", thms1 @ thms2),[])] end |
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
749 |
||>> PureThy.add_thmss |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
750 |
let val thms1 = inst_dj [dj_perm_forget] |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
751 |
and thms2 = inst_dj [dj_pp_forget] |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
752 |
in [(("perm_dj", thms1 @ thms2),[])] end |
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
753 |
||>> PureThy.add_thmss |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
754 |
let val thms1 = inst_pt_at_fs [fresh_iff] |
18626 | 755 |
and thms2 = inst_pt_at [fresh_iff] |
756 |
and thms3 = inst_pt_pt_at_cp_dj [fresh_iff_ineq] |
|
757 |
in [(("abs_fresh", thms1 @ thms2 @ thms3),[])] end |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
758 |
||>> PureThy.add_thmss |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
759 |
let val thms1 = inst_pt_at [abs_fun_supp] |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
760 |
and thms2 = inst_pt_at_fs [abs_fun_supp] |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
761 |
and thms3 = inst_pt_pt_at_cp_dj [abs_fun_supp_ineq] |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
762 |
in [(("abs_supp", thms1 @ thms2 @ thms3),[])] end |
18396
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
763 |
||>> PureThy.add_thmss |
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
764 |
let val thms1 = inst_pt_at [fresh_left] |
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
765 |
and thms2 = inst_pt_pt_at_cp [fresh_left_ineq] |
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
766 |
in [(("fresh_left", thms1 @ thms2),[])] end |
18426 | 767 |
||>> PureThy.add_thmss |
768 |
let val thms1 = inst_pt_at [fresh_bij] |
|
769 |
and thms2 = inst_pt_pt_at_cp [fresh_bij_ineq] |
|
770 |
in [(("fresh_eqvt", thms1 @ thms2),[])] end |
|
18068 | 771 |
end; |
772 |
||
773 |
in NominalData.put (fold Symtab.update (map (rpair ()) full_ak_names) |
|
18262 | 774 |
(NominalData.get thy11)) thy33 |
18068 | 775 |
end; |
776 |
||
777 |
||
778 |
(* syntax und parsing *) |
|
779 |
structure P = OuterParse and K = OuterKeyword; |
|
780 |
||
781 |
val atom_declP = |
|
782 |
OuterSyntax.command "atom_decl" "Declare new kinds of atoms" K.thy_decl |
|
783 |
(Scan.repeat1 P.name >> (Toplevel.theory o create_nom_typedecls)); |
|
784 |
||
785 |
val _ = OuterSyntax.add_parsers [atom_declP]; |
|
786 |
||
787 |
val setup = [NominalData.init]; |
|
788 |
||
789 |
end; |