author | urbanc |
Sat, 07 Apr 2007 12:40:32 +0200 | |
changeset 22611 | 0e008e3ddb1e |
parent 22610 | c8b5133045f3 |
child 22715 | 381e6c45f13b |
permissions | -rw-r--r-- |
19494 | 1 |
(* Title: HOL/Nominal/nominal_atoms.ML |
2 |
ID: $Id$ |
|
3 |
Author: Christian Urban and Stefan Berghofer, TU Muenchen |
|
4 |
||
5 |
Declaration of atom types to be used in nominal datatypes. |
|
6 |
*) |
|
18068 | 7 |
|
8 |
signature NOMINAL_ATOMS = |
|
9 |
sig |
|
10 |
val create_nom_typedecls : string list -> theory -> theory |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
11 |
type atom_info |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
12 |
val get_atom_infos : theory -> atom_info Symtab.table |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
13 |
val get_atom_info : theory -> string -> atom_info option |
18068 | 14 |
val atoms_of : theory -> string list |
15 |
val mk_permT : typ -> typ |
|
18746
a4ece70964ae
made the change for setup-functions not returning functions
urbanc
parents:
18707
diff
changeset
|
16 |
val setup : theory -> theory |
18068 | 17 |
end |
18 |
||
19 |
structure NominalAtoms : NOMINAL_ATOMS = |
|
20 |
struct |
|
21 |
||
22274 | 22 |
val finite_emptyI = thm "finite.emptyI"; |
21669 | 23 |
val Collect_const = thm "Collect_const"; |
24 |
||
25 |
||
18068 | 26 |
(* data kind 'HOL/nominal' *) |
27 |
||
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
28 |
type atom_info = |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
29 |
{pt_class : string, |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
30 |
fs_class : string, |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
31 |
cp_classes : (string * string) list}; |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
32 |
|
18068 | 33 |
structure NominalArgs = |
34 |
struct |
|
35 |
val name = "HOL/nominal"; |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
36 |
type T = atom_info Symtab.table; |
18068 | 37 |
|
38 |
val empty = Symtab.empty; |
|
39 |
val copy = I; |
|
40 |
val extend = I; |
|
41 |
fun merge _ x = Symtab.merge (K true) x; |
|
42 |
||
43 |
fun print sg tab = (); |
|
44 |
end; |
|
45 |
||
46 |
structure NominalData = TheoryDataFun(NominalArgs); |
|
47 |
||
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
48 |
fun make_atom_info ((pt_class, fs_class), cp_classes) = |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
49 |
{pt_class = pt_class, |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
50 |
fs_class = fs_class, |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
51 |
cp_classes = cp_classes}; |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
52 |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
53 |
val get_atom_infos = NominalData.get; |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
54 |
val get_atom_info = Symtab.lookup o NominalData.get; |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
55 |
|
18068 | 56 |
fun atoms_of thy = map fst (Symtab.dest (NominalData.get thy)); |
57 |
||
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
58 |
fun mk_permT T = HOLogic.listT (HOLogic.mk_prodT (T, T)); |
18068 | 59 |
|
60 |
fun mk_Cons x xs = |
|
61 |
let val T = fastype_of x |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
62 |
in Const ("List.list.Cons", T --> HOLogic.listT T --> HOLogic.listT T) $ x $ xs end; |
18068 | 63 |
|
64 |
||
65 |
(* this function sets up all matters related to atom- *) |
|
66 |
(* kinds; the user specifies a list of atom-kind names *) |
|
67 |
(* atom_decl <ak1> ... <akn> *) |
|
68 |
fun create_nom_typedecls ak_names thy = |
|
69 |
let |
|
70 |
(* declares a type-decl for every atom-kind: *) |
|
71 |
(* that is typedecl <ak> *) |
|
72 |
val thy1 = TypedefPackage.add_typedecls (map (fn x => (x,[],NoSyn)) ak_names) thy; |
|
73 |
||
74 |
(* produces a list consisting of pairs: *) |
|
75 |
(* fst component is the atom-kind name *) |
|
76 |
(* snd component is its type *) |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
77 |
val full_ak_names = map (Sign.intern_type thy1) ak_names; |
18068 | 78 |
val ak_names_types = ak_names ~~ map (Type o rpair []) full_ak_names; |
79 |
||
80 |
(* adds for every atom-kind an axiom *) |
|
81 |
(* <ak>_infinite: infinite (UNIV::<ak_type> set) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
82 |
val (inf_axs,thy2) = PureThy.add_axioms_i (map (fn (ak_name, T) => |
18068 | 83 |
let |
20179 | 84 |
val name = ak_name ^ "_infinite" |
18068 | 85 |
val axiom = HOLogic.mk_Trueprop (HOLogic.mk_not |
22274 | 86 |
(Const ("Finite_Set.finite", HOLogic.mk_setT T --> HOLogic.boolT) $ |
87 |
HOLogic.mk_UNIV T)) |
|
18068 | 88 |
in |
20179 | 89 |
((name, axiom), []) |
18068 | 90 |
end) ak_names_types) thy1; |
91 |
||
92 |
(* declares a swapping function for every atom-kind, it is *) |
|
93 |
(* const swap_<ak> :: <akT> * <akT> => <akT> => <akT> *) |
|
94 |
(* swap_<ak> (a,b) c = (if a=c then b (else if b=c then a else c)) *) |
|
95 |
(* overloades then the general swap-function *) |
|
20179 | 96 |
val (swap_eqs, thy3) = fold_map (fn (ak_name, T) => fn thy => |
18068 | 97 |
let |
98 |
val swapT = HOLogic.mk_prodT (T, T) --> T --> T; |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
99 |
val swap_name = Sign.full_name thy ("swap_" ^ ak_name); |
18068 | 100 |
val a = Free ("a", T); |
101 |
val b = Free ("b", T); |
|
102 |
val c = Free ("c", T); |
|
103 |
val ab = Free ("ab", HOLogic.mk_prodT (T, T)) |
|
104 |
val cif = Const ("HOL.If", HOLogic.boolT --> T --> T --> T); |
|
105 |
val cswap_akname = Const (swap_name, swapT); |
|
19494 | 106 |
val cswap = Const ("Nominal.swap", swapT) |
18068 | 107 |
|
108 |
val name = "swap_"^ak_name^"_def"; |
|
109 |
val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
20179 | 110 |
(cswap_akname $ HOLogic.mk_prod (a,b) $ c, |
18068 | 111 |
cif $ HOLogic.mk_eq (a,c) $ b $ (cif $ HOLogic.mk_eq (b,c) $ a $ c))) |
112 |
val def2 = Logic.mk_equals (cswap $ ab $ c, cswap_akname $ ab $ c) |
|
113 |
in |
|
114 |
thy |> Theory.add_consts_i [("swap_" ^ ak_name, swapT, NoSyn)] |
|
20179 | 115 |
|> PureThy.add_defs_unchecked_i true [((name, def2),[])] |
116 |
|> snd |
|
117 |
|> PrimrecPackage.add_primrec_unchecked_i "" [(("", def1),[])] |
|
118 |
end) ak_names_types thy2; |
|
18068 | 119 |
|
120 |
(* declares a permutation function for every atom-kind acting *) |
|
121 |
(* on such atoms *) |
|
122 |
(* const <ak>_prm_<ak> :: (<akT> * <akT>)list => akT => akT *) |
|
123 |
(* <ak>_prm_<ak> [] a = a *) |
|
124 |
(* <ak>_prm_<ak> (x#xs) a = swap_<ak> x (perm xs a) *) |
|
20179 | 125 |
val (prm_eqs, thy4) = fold_map (fn (ak_name, T) => fn thy => |
18068 | 126 |
let |
127 |
val swapT = HOLogic.mk_prodT (T, T) --> T --> T; |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
128 |
val swap_name = Sign.full_name thy ("swap_" ^ ak_name) |
18068 | 129 |
val prmT = mk_permT T --> T --> T; |
130 |
val prm_name = ak_name ^ "_prm_" ^ ak_name; |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
131 |
val qu_prm_name = Sign.full_name thy prm_name; |
18068 | 132 |
val x = Free ("x", HOLogic.mk_prodT (T, T)); |
133 |
val xs = Free ("xs", mk_permT T); |
|
134 |
val a = Free ("a", T) ; |
|
135 |
||
136 |
val cnil = Const ("List.list.Nil", mk_permT T); |
|
137 |
||
138 |
val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (qu_prm_name, prmT) $ cnil $ a, a)); |
|
139 |
||
140 |
val def2 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
141 |
(Const (qu_prm_name, prmT) $ mk_Cons x xs $ a, |
|
142 |
Const (swap_name, swapT) $ x $ (Const (qu_prm_name, prmT) $ xs $ a))); |
|
143 |
in |
|
144 |
thy |> Theory.add_consts_i [(prm_name, mk_permT T --> T --> T, NoSyn)] |
|
19635 | 145 |
|> PrimrecPackage.add_primrec_unchecked_i "" [(("", def1), []),(("", def2), [])] |
20179 | 146 |
end) ak_names_types thy3; |
18068 | 147 |
|
148 |
(* defines permutation functions for all combinations of atom-kinds; *) |
|
149 |
(* there are a trivial cases and non-trivial cases *) |
|
150 |
(* non-trivial case: *) |
|
151 |
(* <ak>_prm_<ak>_def: perm pi a == <ak>_prm_<ak> pi a *) |
|
152 |
(* trivial case with <ak> != <ak'> *) |
|
153 |
(* <ak>_prm<ak'>_def[simp]: perm pi a == a *) |
|
154 |
(* *) |
|
155 |
(* the trivial cases are added to the simplifier, while the non- *) |
|
156 |
(* have their own rules proved below *) |
|
18366 | 157 |
val (perm_defs, thy5) = fold_map (fn (ak_name, T) => fn thy => |
158 |
fold_map (fn (ak_name', T') => fn thy' => |
|
18068 | 159 |
let |
160 |
val perm_def_name = ak_name ^ "_prm_" ^ ak_name'; |
|
161 |
val pi = Free ("pi", mk_permT T); |
|
162 |
val a = Free ("a", T'); |
|
19494 | 163 |
val cperm = Const ("Nominal.perm", mk_permT T --> T' --> T'); |
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
164 |
val cperm_def = Const (Sign.full_name thy' perm_def_name, mk_permT T --> T' --> T'); |
18068 | 165 |
|
166 |
val name = ak_name ^ "_prm_" ^ ak_name' ^ "_def"; |
|
167 |
val def = Logic.mk_equals |
|
168 |
(cperm $ pi $ a, if ak_name = ak_name' then cperm_def $ pi $ a else a) |
|
169 |
in |
|
19635 | 170 |
PureThy.add_defs_unchecked_i true [((name, def),[])] thy' |
18366 | 171 |
end) ak_names_types thy) ak_names_types thy4; |
18068 | 172 |
|
173 |
(* proves that every atom-kind is an instance of at *) |
|
174 |
(* lemma at_<ak>_inst: *) |
|
175 |
(* at TYPE(<ak>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
176 |
val (prm_cons_thms,thy6) = |
18068 | 177 |
thy5 |> PureThy.add_thms (map (fn (ak_name, T) => |
178 |
let |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
179 |
val ak_name_qu = Sign.full_name thy5 (ak_name); |
18068 | 180 |
val i_type = Type(ak_name_qu,[]); |
19494 | 181 |
val cat = Const ("Nominal.at",(Term.itselfT i_type) --> HOLogic.boolT); |
18068 | 182 |
val at_type = Logic.mk_type i_type; |
183 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy5 |
|
184 |
[Name "at_def", |
|
185 |
Name (ak_name ^ "_prm_" ^ ak_name ^ "_def"), |
|
186 |
Name (ak_name ^ "_prm_" ^ ak_name ^ ".simps"), |
|
187 |
Name ("swap_" ^ ak_name ^ "_def"), |
|
188 |
Name ("swap_" ^ ak_name ^ ".simps"), |
|
189 |
Name (ak_name ^ "_infinite")] |
|
190 |
||
191 |
val name = "at_"^ak_name^ "_inst"; |
|
192 |
val statement = HOLogic.mk_Trueprop (cat $ at_type); |
|
193 |
||
194 |
val proof = fn _ => auto_tac (claset(),simp_s); |
|
195 |
||
196 |
in |
|
20046 | 197 |
((name, Goal.prove_global thy5 [] [] statement proof), []) |
18068 | 198 |
end) ak_names_types); |
199 |
||
200 |
(* declares a perm-axclass for every atom-kind *) |
|
201 |
(* axclass pt_<ak> *) |
|
202 |
(* pt_<ak>1[simp]: perm [] x = x *) |
|
203 |
(* pt_<ak>2: perm (pi1@pi2) x = perm pi1 (perm pi2 x) *) |
|
204 |
(* 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
|
205 |
val (pt_ax_classes,thy7) = fold_map (fn (ak_name, T) => fn thy => |
18068 | 206 |
let |
207 |
val cl_name = "pt_"^ak_name; |
|
208 |
val ty = TFree("'a",["HOL.type"]); |
|
209 |
val x = Free ("x", ty); |
|
210 |
val pi1 = Free ("pi1", mk_permT T); |
|
211 |
val pi2 = Free ("pi2", mk_permT T); |
|
19494 | 212 |
val cperm = Const ("Nominal.perm", mk_permT T --> ty --> ty); |
18068 | 213 |
val cnil = Const ("List.list.Nil", mk_permT T); |
214 |
val cappend = Const ("List.op @",mk_permT T --> mk_permT T --> mk_permT T); |
|
19494 | 215 |
val cprm_eq = Const ("Nominal.prm_eq",mk_permT T --> mk_permT T --> HOLogic.boolT); |
18068 | 216 |
(* nil axiom *) |
217 |
val axiom1 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
218 |
(cperm $ cnil $ x, x)); |
|
219 |
(* append axiom *) |
|
220 |
val axiom2 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
221 |
(cperm $ (cappend $ pi1 $ pi2) $ x, cperm $ pi1 $ (cperm $ pi2 $ x))); |
|
222 |
(* perm-eq axiom *) |
|
223 |
val axiom3 = Logic.mk_implies |
|
224 |
(HOLogic.mk_Trueprop (cprm_eq $ pi1 $ pi2), |
|
225 |
HOLogic.mk_Trueprop (HOLogic.mk_eq (cperm $ pi1 $ x, cperm $ pi2 $ x))); |
|
226 |
in |
|
19509 | 227 |
AxClass.define_class_i (cl_name, ["HOL.type"]) [] |
19488 | 228 |
[((cl_name ^ "1", [Simplifier.simp_add]), [axiom1]), |
229 |
((cl_name ^ "2", []), [axiom2]), |
|
230 |
((cl_name ^ "3", []), [axiom3])] thy |
|
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
231 |
end) ak_names_types thy6; |
18068 | 232 |
|
233 |
(* proves that every pt_<ak>-type together with <ak>-type *) |
|
234 |
(* instance of pt *) |
|
235 |
(* lemma pt_<ak>_inst: *) |
|
236 |
(* pt TYPE('x::pt_<ak>) TYPE(<ak>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
237 |
val (prm_inst_thms,thy8) = |
18068 | 238 |
thy7 |> PureThy.add_thms (map (fn (ak_name, T) => |
239 |
let |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
240 |
val ak_name_qu = Sign.full_name thy7 ak_name; |
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
241 |
val pt_name_qu = Sign.full_name thy7 ("pt_"^ak_name); |
18068 | 242 |
val i_type1 = TFree("'x",[pt_name_qu]); |
243 |
val i_type2 = Type(ak_name_qu,[]); |
|
19494 | 244 |
val cpt = Const ("Nominal.pt",(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); |
18068 | 245 |
val pt_type = Logic.mk_type i_type1; |
246 |
val at_type = Logic.mk_type i_type2; |
|
247 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy7 |
|
248 |
[Name "pt_def", |
|
249 |
Name ("pt_" ^ ak_name ^ "1"), |
|
250 |
Name ("pt_" ^ ak_name ^ "2"), |
|
251 |
Name ("pt_" ^ ak_name ^ "3")]; |
|
252 |
||
253 |
val name = "pt_"^ak_name^ "_inst"; |
|
254 |
val statement = HOLogic.mk_Trueprop (cpt $ pt_type $ at_type); |
|
255 |
||
256 |
val proof = fn _ => auto_tac (claset(),simp_s); |
|
257 |
in |
|
20046 | 258 |
((name, Goal.prove_global thy7 [] [] statement proof), []) |
18068 | 259 |
end) ak_names_types); |
260 |
||
261 |
(* declares an fs-axclass for every atom-kind *) |
|
262 |
(* axclass fs_<ak> *) |
|
263 |
(* 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
|
264 |
val (fs_ax_classes,thy11) = fold_map (fn (ak_name, T) => fn thy => |
18068 | 265 |
let |
266 |
val cl_name = "fs_"^ak_name; |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
267 |
val pt_name = Sign.full_name thy ("pt_"^ak_name); |
18068 | 268 |
val ty = TFree("'a",["HOL.type"]); |
269 |
val x = Free ("x", ty); |
|
19494 | 270 |
val csupp = Const ("Nominal.supp", ty --> HOLogic.mk_setT T); |
22274 | 271 |
val cfinite = Const ("Finite_Set.finite", HOLogic.mk_setT T --> HOLogic.boolT) |
18068 | 272 |
|
22274 | 273 |
val axiom1 = HOLogic.mk_Trueprop (cfinite $ (csupp $ x)); |
18068 | 274 |
|
275 |
in |
|
19509 | 276 |
AxClass.define_class_i (cl_name, [pt_name]) [] [((cl_name ^ "1", []), [axiom1])] thy |
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
277 |
end) ak_names_types thy8; |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
278 |
|
18068 | 279 |
(* proves that every fs_<ak>-type together with <ak>-type *) |
280 |
(* instance of fs-type *) |
|
281 |
(* lemma abst_<ak>_inst: *) |
|
282 |
(* fs TYPE('x::pt_<ak>) TYPE (<ak>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
283 |
val (fs_inst_thms,thy12) = |
18068 | 284 |
thy11 |> PureThy.add_thms (map (fn (ak_name, T) => |
285 |
let |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
286 |
val ak_name_qu = Sign.full_name thy11 ak_name; |
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
287 |
val fs_name_qu = Sign.full_name thy11 ("fs_"^ak_name); |
18068 | 288 |
val i_type1 = TFree("'x",[fs_name_qu]); |
289 |
val i_type2 = Type(ak_name_qu,[]); |
|
19494 | 290 |
val cfs = Const ("Nominal.fs", |
18068 | 291 |
(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); |
292 |
val fs_type = Logic.mk_type i_type1; |
|
293 |
val at_type = Logic.mk_type i_type2; |
|
294 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy11 |
|
295 |
[Name "fs_def", |
|
296 |
Name ("fs_" ^ ak_name ^ "1")]; |
|
297 |
||
298 |
val name = "fs_"^ak_name^ "_inst"; |
|
299 |
val statement = HOLogic.mk_Trueprop (cfs $ fs_type $ at_type); |
|
300 |
||
301 |
val proof = fn _ => auto_tac (claset(),simp_s); |
|
302 |
in |
|
20046 | 303 |
((name, Goal.prove_global thy11 [] [] statement proof), []) |
18068 | 304 |
end) ak_names_types); |
305 |
||
306 |
(* declares for every atom-kind combination an axclass *) |
|
307 |
(* cp_<ak1>_<ak2> giving a composition property *) |
|
308 |
(* cp_<ak1>_<ak2>1: pi1 o pi2 o x = (pi1 o pi2) o (pi1 o x) *) |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
309 |
val (cp_ax_classes,thy12b) = fold_map (fn (ak_name, T) => fn thy => |
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
310 |
fold_map (fn (ak_name', T') => fn thy' => |
18068 | 311 |
let |
312 |
val cl_name = "cp_"^ak_name^"_"^ak_name'; |
|
313 |
val ty = TFree("'a",["HOL.type"]); |
|
314 |
val x = Free ("x", ty); |
|
315 |
val pi1 = Free ("pi1", mk_permT T); |
|
316 |
val pi2 = Free ("pi2", mk_permT T'); |
|
19494 | 317 |
val cperm1 = Const ("Nominal.perm", mk_permT T --> ty --> ty); |
318 |
val cperm2 = Const ("Nominal.perm", mk_permT T' --> ty --> ty); |
|
319 |
val cperm3 = Const ("Nominal.perm", mk_permT T --> mk_permT T' --> mk_permT T'); |
|
18068 | 320 |
|
321 |
val ax1 = HOLogic.mk_Trueprop |
|
322 |
(HOLogic.mk_eq (cperm1 $ pi1 $ (cperm2 $ pi2 $ x), |
|
323 |
cperm2 $ (cperm3 $ pi1 $ pi2) $ (cperm1 $ pi1 $ x))); |
|
324 |
in |
|
19509 | 325 |
AxClass.define_class_i (cl_name, ["HOL.type"]) [] [((cl_name ^ "1", []), [ax1])] thy' |
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
326 |
end) ak_names_types thy) ak_names_types thy12; |
18068 | 327 |
|
328 |
(* proves for every <ak>-combination a cp_<ak1>_<ak2>_inst theorem; *) |
|
329 |
(* lemma cp_<ak1>_<ak2>_inst: *) |
|
330 |
(* 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
|
331 |
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
|
332 |
fold_map (fn (ak_name', T') => fn thy' => |
18068 | 333 |
let |
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
334 |
val ak_name_qu = Sign.full_name thy' (ak_name); |
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
335 |
val ak_name_qu' = Sign.full_name thy' (ak_name'); |
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
336 |
val cp_name_qu = Sign.full_name thy' ("cp_"^ak_name^"_"^ak_name'); |
18068 | 337 |
val i_type0 = TFree("'a",[cp_name_qu]); |
338 |
val i_type1 = Type(ak_name_qu,[]); |
|
339 |
val i_type2 = Type(ak_name_qu',[]); |
|
19494 | 340 |
val ccp = Const ("Nominal.cp", |
18068 | 341 |
(Term.itselfT i_type0)-->(Term.itselfT i_type1)--> |
342 |
(Term.itselfT i_type2)-->HOLogic.boolT); |
|
343 |
val at_type = Logic.mk_type i_type1; |
|
344 |
val at_type' = Logic.mk_type i_type2; |
|
345 |
val cp_type = Logic.mk_type i_type0; |
|
346 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy' [(Name "cp_def")]; |
|
347 |
val cp1 = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"1")); |
|
348 |
||
349 |
val name = "cp_"^ak_name^ "_"^ak_name'^"_inst"; |
|
350 |
val statement = HOLogic.mk_Trueprop (ccp $ cp_type $ at_type $ at_type'); |
|
351 |
||
352 |
val proof = fn _ => EVERY [auto_tac (claset(),simp_s), rtac cp1 1]; |
|
353 |
in |
|
20046 | 354 |
PureThy.add_thms [((name, Goal.prove_global thy' [] [] statement proof), [])] thy' |
18068 | 355 |
end) |
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
356 |
ak_names_types thy) ak_names_types thy12b; |
18068 | 357 |
|
358 |
(* proves for every non-trivial <ak>-combination a disjointness *) |
|
359 |
(* theorem; i.e. <ak1> != <ak2> *) |
|
360 |
(* lemma ds_<ak1>_<ak2>: *) |
|
361 |
(* dj TYPE(<ak1>) TYPE(<ak2>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
362 |
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
|
363 |
fold_map (fn (ak_name',T') => fn thy' => |
18068 | 364 |
(if not (ak_name = ak_name') |
365 |
then |
|
366 |
let |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
367 |
val ak_name_qu = Sign.full_name thy' ak_name; |
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
368 |
val ak_name_qu' = Sign.full_name thy' ak_name'; |
18068 | 369 |
val i_type1 = Type(ak_name_qu,[]); |
370 |
val i_type2 = Type(ak_name_qu',[]); |
|
19494 | 371 |
val cdj = Const ("Nominal.disjoint", |
18068 | 372 |
(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); |
373 |
val at_type = Logic.mk_type i_type1; |
|
374 |
val at_type' = Logic.mk_type i_type2; |
|
375 |
val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy' |
|
376 |
[Name "disjoint_def", |
|
377 |
Name (ak_name^"_prm_"^ak_name'^"_def"), |
|
378 |
Name (ak_name'^"_prm_"^ak_name^"_def")]; |
|
379 |
||
380 |
val name = "dj_"^ak_name^"_"^ak_name'; |
|
381 |
val statement = HOLogic.mk_Trueprop (cdj $ at_type $ at_type'); |
|
382 |
||
383 |
val proof = fn _ => auto_tac (claset(),simp_s); |
|
384 |
in |
|
20046 | 385 |
PureThy.add_thms [((name, Goal.prove_global thy' [] [] statement proof), [])] thy' |
18068 | 386 |
end |
387 |
else |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
388 |
([],thy'))) (* do nothing branch, if ak_name = ak_name' *) |
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
389 |
ak_names_types thy) ak_names_types thy12c; |
18068 | 390 |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
391 |
(******** pt_<ak> class instances ********) |
18068 | 392 |
(*=========================================*) |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
393 |
(* some abbreviations for theorems *) |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
394 |
val pt1 = thm "pt1"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
395 |
val pt2 = thm "pt2"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
396 |
val pt3 = thm "pt3"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
397 |
val at_pt_inst = thm "at_pt_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
398 |
val pt_set_inst = thm "pt_set_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
399 |
val pt_unit_inst = thm "pt_unit_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
400 |
val pt_prod_inst = thm "pt_prod_inst"; |
18600 | 401 |
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
|
402 |
val pt_list_inst = thm "pt_list_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
403 |
val pt_optn_inst = thm "pt_option_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
404 |
val pt_noptn_inst = thm "pt_noption_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
405 |
val pt_fun_inst = thm "pt_fun_inst"; |
18068 | 406 |
|
18435 | 407 |
(* for all atom-kind combinations <ak>/<ak'> show that *) |
408 |
(* every <ak> is an instance of pt_<ak'>; the proof for *) |
|
409 |
(* ak!=ak' is by definition; the case ak=ak' uses at_pt_inst. *) |
|
18431 | 410 |
val thy13 = fold (fn ak_name => fn thy => |
411 |
fold (fn ak_name' => fn thy' => |
|
412 |
let |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
413 |
val qu_name = Sign.full_name thy' ak_name'; |
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
414 |
val cls_name = Sign.full_name thy' ("pt_"^ak_name); |
18431 | 415 |
val at_inst = PureThy.get_thm thy' (Name ("at_"^ak_name'^"_inst")); |
416 |
||
19133
7e84a1a3741c
Adapted to Florian's recent changes to the AxClass package.
berghofe
parents:
18759
diff
changeset
|
417 |
val proof1 = EVERY [ClassPackage.intro_classes_tac [], |
18068 | 418 |
rtac ((at_inst RS at_pt_inst) RS pt1) 1, |
419 |
rtac ((at_inst RS at_pt_inst) RS pt2) 1, |
|
420 |
rtac ((at_inst RS at_pt_inst) RS pt3) 1, |
|
421 |
atac 1]; |
|
18431 | 422 |
val simp_s = HOL_basic_ss addsimps |
423 |
PureThy.get_thmss thy' [Name (ak_name^"_prm_"^ak_name'^"_def")]; |
|
19133
7e84a1a3741c
Adapted to Florian's recent changes to the AxClass package.
berghofe
parents:
18759
diff
changeset
|
424 |
val proof2 = EVERY [ClassPackage.intro_classes_tac [], REPEAT (asm_simp_tac simp_s 1)]; |
18431 | 425 |
|
426 |
in |
|
427 |
thy' |
|
19275 | 428 |
|> AxClass.prove_arity (qu_name,[],[cls_name]) |
18431 | 429 |
(if ak_name = ak_name' then proof1 else proof2) |
430 |
end) ak_names thy) ak_names thy12c; |
|
18068 | 431 |
|
18430 | 432 |
(* show that *) |
433 |
(* fun(pt_<ak>,pt_<ak>) *) |
|
18579
002d371401f5
changed the name of the type "nOption" to "noption".
urbanc
parents:
18438
diff
changeset
|
434 |
(* noption(pt_<ak>) *) |
18430 | 435 |
(* option(pt_<ak>) *) |
436 |
(* list(pt_<ak>) *) |
|
437 |
(* *(pt_<ak>,pt_<ak>) *) |
|
18600 | 438 |
(* nprod(pt_<ak>,pt_<ak>) *) |
18430 | 439 |
(* unit *) |
440 |
(* set(pt_<ak>) *) |
|
441 |
(* are instances of pt_<ak> *) |
|
18431 | 442 |
val thy18 = fold (fn ak_name => fn thy => |
18068 | 443 |
let |
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
444 |
val cls_name = Sign.full_name thy ("pt_"^ak_name); |
18068 | 445 |
val at_thm = PureThy.get_thm thy (Name ("at_"^ak_name^"_inst")); |
446 |
val pt_inst = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst")); |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
447 |
|
18430 | 448 |
fun pt_proof thm = |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
449 |
EVERY [ClassPackage.intro_classes_tac [], |
18430 | 450 |
rtac (thm RS pt1) 1, rtac (thm RS pt2) 1, rtac (thm RS pt3) 1, atac 1]; |
451 |
||
452 |
val pt_thm_fun = at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst)); |
|
453 |
val pt_thm_noptn = pt_inst RS pt_noptn_inst; |
|
454 |
val pt_thm_optn = pt_inst RS pt_optn_inst; |
|
455 |
val pt_thm_list = pt_inst RS pt_list_inst; |
|
456 |
val pt_thm_prod = pt_inst RS (pt_inst RS pt_prod_inst); |
|
18600 | 457 |
val pt_thm_nprod = pt_inst RS (pt_inst RS pt_nprod_inst); |
18430 | 458 |
val pt_thm_unit = pt_unit_inst; |
459 |
val pt_thm_set = pt_inst RS pt_set_inst |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
460 |
in |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
461 |
thy |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
462 |
|> AxClass.prove_arity ("fun",[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_fun) |
19494 | 463 |
|> AxClass.prove_arity ("Nominal.noption",[[cls_name]],[cls_name]) (pt_proof pt_thm_noptn) |
19275 | 464 |
|> AxClass.prove_arity ("Datatype.option",[[cls_name]],[cls_name]) (pt_proof pt_thm_optn) |
465 |
|> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (pt_proof pt_thm_list) |
|
466 |
|> AxClass.prove_arity ("*",[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_prod) |
|
19494 | 467 |
|> AxClass.prove_arity ("Nominal.nprod",[[cls_name],[cls_name]],[cls_name]) |
18600 | 468 |
(pt_proof pt_thm_nprod) |
19275 | 469 |
|> AxClass.prove_arity ("Product_Type.unit",[],[cls_name]) (pt_proof pt_thm_unit) |
470 |
|> AxClass.prove_arity ("set",[[cls_name]],[cls_name]) (pt_proof pt_thm_set) |
|
18430 | 471 |
end) ak_names thy13; |
18068 | 472 |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
473 |
(******** fs_<ak> class instances ********) |
18068 | 474 |
(*=========================================*) |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
475 |
(* abbreviations for some lemmas *) |
18431 | 476 |
val fs1 = thm "fs1"; |
477 |
val fs_at_inst = thm "fs_at_inst"; |
|
478 |
val fs_unit_inst = thm "fs_unit_inst"; |
|
479 |
val fs_prod_inst = thm "fs_prod_inst"; |
|
18600 | 480 |
val fs_nprod_inst = thm "fs_nprod_inst"; |
18431 | 481 |
val fs_list_inst = thm "fs_list_inst"; |
482 |
val fs_option_inst = thm "fs_option_inst"; |
|
18437
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
483 |
val dj_supp = thm "dj_supp" |
18068 | 484 |
|
485 |
(* shows that <ak> is an instance of fs_<ak> *) |
|
486 |
(* uses the theorem at_<ak>_inst *) |
|
18431 | 487 |
val thy20 = fold (fn ak_name => fn thy => |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
488 |
fold (fn ak_name' => fn thy' => |
18437
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
489 |
let |
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
490 |
val qu_name = Sign.full_name thy' ak_name'; |
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
491 |
val qu_class = Sign.full_name thy' ("fs_"^ak_name); |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
492 |
val proof = |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
493 |
(if ak_name = ak_name' |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
494 |
then |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
495 |
let val at_thm = PureThy.get_thm thy' (Name ("at_"^ak_name^"_inst")); |
19133
7e84a1a3741c
Adapted to Florian's recent changes to the AxClass package.
berghofe
parents:
18759
diff
changeset
|
496 |
in EVERY [ClassPackage.intro_classes_tac [], |
18437
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
497 |
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
|
498 |
else |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
499 |
let val dj_inst = PureThy.get_thm thy' (Name ("dj_"^ak_name'^"_"^ak_name)); |
22274 | 500 |
val simp_s = HOL_basic_ss addsimps [dj_inst RS dj_supp, finite_emptyI]; |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
501 |
in EVERY [ClassPackage.intro_classes_tac [], asm_simp_tac simp_s 1] end) |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
502 |
in |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
503 |
AxClass.prove_arity (qu_name,[],[qu_class]) proof thy' |
18437
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
504 |
end) ak_names thy) ak_names thy18; |
18068 | 505 |
|
18431 | 506 |
(* shows that *) |
507 |
(* unit *) |
|
508 |
(* *(fs_<ak>,fs_<ak>) *) |
|
18600 | 509 |
(* nprod(fs_<ak>,fs_<ak>) *) |
18431 | 510 |
(* list(fs_<ak>) *) |
511 |
(* option(fs_<ak>) *) |
|
512 |
(* are instances of fs_<ak> *) |
|
18068 | 513 |
|
18431 | 514 |
val thy24 = fold (fn ak_name => fn thy => |
515 |
let |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
516 |
val cls_name = Sign.full_name thy ("fs_"^ak_name); |
18068 | 517 |
val fs_inst = PureThy.get_thm thy (Name ("fs_"^ak_name^"_inst")); |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
518 |
fun fs_proof thm = EVERY [ClassPackage.intro_classes_tac [], rtac (thm RS fs1) 1]; |
18068 | 519 |
|
18600 | 520 |
val fs_thm_unit = fs_unit_inst; |
521 |
val fs_thm_prod = fs_inst RS (fs_inst RS fs_prod_inst); |
|
522 |
val fs_thm_nprod = fs_inst RS (fs_inst RS fs_nprod_inst); |
|
523 |
val fs_thm_list = fs_inst RS fs_list_inst; |
|
524 |
val fs_thm_optn = fs_inst RS fs_option_inst; |
|
18431 | 525 |
in |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
526 |
thy |
19275 | 527 |
|> AxClass.prove_arity ("Product_Type.unit",[],[cls_name]) (fs_proof fs_thm_unit) |
528 |
|> AxClass.prove_arity ("*",[[cls_name],[cls_name]],[cls_name]) (fs_proof fs_thm_prod) |
|
19494 | 529 |
|> AxClass.prove_arity ("Nominal.nprod",[[cls_name],[cls_name]],[cls_name]) |
18600 | 530 |
(fs_proof fs_thm_nprod) |
19275 | 531 |
|> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (fs_proof fs_thm_list) |
532 |
|> AxClass.prove_arity ("Datatype.option",[[cls_name]],[cls_name]) (fs_proof fs_thm_optn) |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
533 |
end) ak_names thy20; |
18431 | 534 |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
535 |
(******** cp_<ak>_<ai> class instances ********) |
18068 | 536 |
(*==============================================*) |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
537 |
(* abbreviations for some lemmas *) |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
538 |
val cp1 = thm "cp1"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
539 |
val cp_unit_inst = thm "cp_unit_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
540 |
val cp_bool_inst = thm "cp_bool_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
541 |
val cp_prod_inst = thm "cp_prod_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
542 |
val cp_list_inst = thm "cp_list_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
543 |
val cp_fun_inst = thm "cp_fun_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
544 |
val cp_option_inst = thm "cp_option_inst"; |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
545 |
val cp_noption_inst = thm "cp_noption_inst"; |
22536
35debf264656
made the type sets instance of the "cp" type-class
urbanc
parents:
22535
diff
changeset
|
546 |
val cp_set_inst = thm "cp_set_inst"; |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
547 |
val pt_perm_compose = thm "pt_perm_compose"; |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
548 |
|
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
549 |
val dj_pp_forget = thm "dj_perm_perm_forget"; |
18068 | 550 |
|
551 |
(* shows that <aj> is an instance of cp_<ak>_<ai> *) |
|
18432 | 552 |
(* for every <ak>/<ai>-combination *) |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
553 |
val thy25 = fold (fn ak_name => fn thy => |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
554 |
fold (fn ak_name' => fn thy' => |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
555 |
fold (fn ak_name'' => fn thy'' => |
18068 | 556 |
let |
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
557 |
val name = Sign.full_name thy'' ak_name; |
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
558 |
val cls_name = Sign.full_name thy'' ("cp_"^ak_name'^"_"^ak_name''); |
18068 | 559 |
val proof = |
560 |
(if (ak_name'=ak_name'') then |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
561 |
(let |
18068 | 562 |
val pt_inst = PureThy.get_thm thy'' (Name ("pt_"^ak_name''^"_inst")); |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
563 |
val at_inst = PureThy.get_thm thy'' (Name ("at_"^ak_name''^"_inst")); |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
564 |
in |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
565 |
EVERY [ClassPackage.intro_classes_tac [], |
18068 | 566 |
rtac (at_inst RS (pt_inst RS pt_perm_compose)) 1] |
567 |
end) |
|
568 |
else |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
569 |
(let |
18068 | 570 |
val dj_inst = PureThy.get_thm thy'' (Name ("dj_"^ak_name''^"_"^ak_name')); |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
571 |
val simp_s = HOL_basic_ss addsimps |
18068 | 572 |
((dj_inst RS dj_pp_forget):: |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
573 |
(PureThy.get_thmss thy'' |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
574 |
[Name (ak_name' ^"_prm_"^ak_name^"_def"), |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
575 |
Name (ak_name''^"_prm_"^ak_name^"_def")])); |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
576 |
in |
19133
7e84a1a3741c
Adapted to Florian's recent changes to the AxClass package.
berghofe
parents:
18759
diff
changeset
|
577 |
EVERY [ClassPackage.intro_classes_tac [], simp_tac simp_s 1] |
18068 | 578 |
end)) |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
579 |
in |
19275 | 580 |
AxClass.prove_arity (name,[],[cls_name]) proof thy'' |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
581 |
end) ak_names thy') ak_names thy) ak_names thy24; |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
582 |
|
18432 | 583 |
(* shows that *) |
584 |
(* units *) |
|
585 |
(* products *) |
|
586 |
(* lists *) |
|
587 |
(* functions *) |
|
588 |
(* options *) |
|
589 |
(* noptions *) |
|
22536
35debf264656
made the type sets instance of the "cp" type-class
urbanc
parents:
22535
diff
changeset
|
590 |
(* sets *) |
18432 | 591 |
(* are instances of cp_<ak>_<ai> for every <ak>/<ai>-combination *) |
592 |
val thy26 = fold (fn ak_name => fn thy => |
|
593 |
fold (fn ak_name' => fn thy' => |
|
594 |
let |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
595 |
val cls_name = Sign.full_name thy' ("cp_"^ak_name^"_"^ak_name'); |
18068 | 596 |
val cp_inst = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst")); |
597 |
val pt_inst = PureThy.get_thm thy' (Name ("pt_"^ak_name^"_inst")); |
|
598 |
val at_inst = PureThy.get_thm thy' (Name ("at_"^ak_name^"_inst")); |
|
18432 | 599 |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
600 |
fun cp_proof thm = EVERY [ClassPackage.intro_classes_tac [],rtac (thm RS cp1) 1]; |
18432 | 601 |
|
602 |
val cp_thm_unit = cp_unit_inst; |
|
603 |
val cp_thm_prod = cp_inst RS (cp_inst RS cp_prod_inst); |
|
604 |
val cp_thm_list = cp_inst RS cp_list_inst; |
|
605 |
val cp_thm_fun = at_inst RS (pt_inst RS (cp_inst RS (cp_inst RS cp_fun_inst))); |
|
606 |
val cp_thm_optn = cp_inst RS cp_option_inst; |
|
607 |
val cp_thm_noptn = cp_inst RS cp_noption_inst; |
|
22536
35debf264656
made the type sets instance of the "cp" type-class
urbanc
parents:
22535
diff
changeset
|
608 |
val cp_thm_set = cp_inst RS cp_set_inst; |
18432 | 609 |
in |
610 |
thy' |
|
19275 | 611 |
|> AxClass.prove_arity ("Product_Type.unit",[],[cls_name]) (cp_proof cp_thm_unit) |
612 |
|> AxClass.prove_arity ("*",[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_prod) |
|
613 |
|> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (cp_proof cp_thm_list) |
|
614 |
|> AxClass.prove_arity ("fun",[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_fun) |
|
615 |
|> AxClass.prove_arity ("Datatype.option",[[cls_name]],[cls_name]) (cp_proof cp_thm_optn) |
|
19494 | 616 |
|> AxClass.prove_arity ("Nominal.noption",[[cls_name]],[cls_name]) (cp_proof cp_thm_noptn) |
22536
35debf264656
made the type sets instance of the "cp" type-class
urbanc
parents:
22535
diff
changeset
|
617 |
|> AxClass.prove_arity ("set",[[cls_name]],[cls_name]) (cp_proof cp_thm_set) |
18432 | 618 |
end) ak_names thy) ak_names thy25; |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
619 |
|
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
620 |
(* show that discrete nominal types are permutation types, finitely *) |
18432 | 621 |
(* supported and have the commutation property *) |
622 |
(* discrete types have a permutation operation defined as pi o x = x; *) |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
623 |
(* which renders the proofs to be simple "simp_all"-proofs. *) |
18432 | 624 |
val thy32 = |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
625 |
let |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
626 |
fun discrete_pt_inst discrete_ty defn = |
18432 | 627 |
fold (fn ak_name => fn thy => |
628 |
let |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
629 |
val qu_class = Sign.full_name thy ("pt_"^ak_name); |
18432 | 630 |
val simp_s = HOL_basic_ss addsimps [defn]; |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
631 |
val proof = EVERY [ClassPackage.intro_classes_tac [], REPEAT (asm_simp_tac simp_s 1)]; |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
632 |
in |
19275 | 633 |
AxClass.prove_arity (discrete_ty,[],[qu_class]) proof thy |
18432 | 634 |
end) ak_names; |
18068 | 635 |
|
18432 | 636 |
fun discrete_fs_inst discrete_ty defn = |
637 |
fold (fn ak_name => fn thy => |
|
638 |
let |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
639 |
val qu_class = Sign.full_name thy ("fs_"^ak_name); |
19494 | 640 |
val supp_def = thm "Nominal.supp_def"; |
22274 | 641 |
val simp_s = HOL_ss addsimps [supp_def,Collect_const,finite_emptyI,defn]; |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
642 |
val proof = EVERY [ClassPackage.intro_classes_tac [], asm_simp_tac simp_s 1]; |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
643 |
in |
19275 | 644 |
AxClass.prove_arity (discrete_ty,[],[qu_class]) proof thy |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
645 |
end) ak_names; |
18068 | 646 |
|
18432 | 647 |
fun discrete_cp_inst discrete_ty defn = |
648 |
fold (fn ak_name' => (fold (fn ak_name => fn thy => |
|
649 |
let |
|
21289
920b7b893d9c
deleted all uses of sign_of as it's now the identity function
urbanc
parents:
20377
diff
changeset
|
650 |
val qu_class = Sign.full_name thy ("cp_"^ak_name^"_"^ak_name'); |
19494 | 651 |
val supp_def = thm "Nominal.supp_def"; |
18432 | 652 |
val simp_s = HOL_ss addsimps [defn]; |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
653 |
val proof = EVERY [ClassPackage.intro_classes_tac [], asm_simp_tac simp_s 1]; |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
654 |
in |
19275 | 655 |
AxClass.prove_arity (discrete_ty,[],[qu_class]) proof thy |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
656 |
end) ak_names)) ak_names; |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
657 |
|
18432 | 658 |
in |
659 |
thy26 |
|
660 |
|> discrete_pt_inst "nat" (thm "perm_nat_def") |
|
661 |
|> discrete_fs_inst "nat" (thm "perm_nat_def") |
|
662 |
|> discrete_cp_inst "nat" (thm "perm_nat_def") |
|
663 |
|> discrete_pt_inst "bool" (thm "perm_bool") |
|
664 |
|> discrete_fs_inst "bool" (thm "perm_bool") |
|
665 |
|> discrete_cp_inst "bool" (thm "perm_bool") |
|
666 |
|> discrete_pt_inst "IntDef.int" (thm "perm_int_def") |
|
667 |
|> discrete_fs_inst "IntDef.int" (thm "perm_int_def") |
|
668 |
|> discrete_cp_inst "IntDef.int" (thm "perm_int_def") |
|
669 |
|> discrete_pt_inst "List.char" (thm "perm_char_def") |
|
670 |
|> discrete_fs_inst "List.char" (thm "perm_char_def") |
|
671 |
|> discrete_cp_inst "List.char" (thm "perm_char_def") |
|
672 |
end; |
|
673 |
||
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
674 |
|
18262 | 675 |
(* abbreviations for some lemmas *) |
676 |
(*===============================*) |
|
19494 | 677 |
val abs_fun_pi = thm "Nominal.abs_fun_pi"; |
678 |
val abs_fun_pi_ineq = thm "Nominal.abs_fun_pi_ineq"; |
|
679 |
val abs_fun_eq = thm "Nominal.abs_fun_eq"; |
|
19562
e56b3c967ae8
added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents:
19548
diff
changeset
|
680 |
val abs_fun_eq' = thm "Nominal.abs_fun_eq'"; |
19494 | 681 |
val dj_perm_forget = thm "Nominal.dj_perm_forget"; |
682 |
val dj_pp_forget = thm "Nominal.dj_perm_perm_forget"; |
|
683 |
val fresh_iff = thm "Nominal.fresh_abs_fun_iff"; |
|
684 |
val fresh_iff_ineq = thm "Nominal.fresh_abs_fun_iff_ineq"; |
|
685 |
val abs_fun_supp = thm "Nominal.abs_fun_supp"; |
|
686 |
val abs_fun_supp_ineq = thm "Nominal.abs_fun_supp_ineq"; |
|
687 |
val pt_swap_bij = thm "Nominal.pt_swap_bij"; |
|
22557
6775c71f1da0
added pt_bij' to the collection perm_swap and also added abs_perm to the collection of equivariance lemmas
urbanc
parents:
22536
diff
changeset
|
688 |
val pt_swap_bij' = thm "Nominal.pt_swap_bij'"; |
19494 | 689 |
val pt_fresh_fresh = thm "Nominal.pt_fresh_fresh"; |
690 |
val pt_bij = thm "Nominal.pt_bij"; |
|
691 |
val pt_perm_compose = thm "Nominal.pt_perm_compose"; |
|
692 |
val pt_perm_compose' = thm "Nominal.pt_perm_compose'"; |
|
693 |
val perm_app = thm "Nominal.pt_fun_app_eq"; |
|
694 |
val at_fresh = thm "Nominal.at_fresh"; |
|
19972
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
695 |
val at_fresh_ineq = thm "Nominal.at_fresh_ineq"; |
19494 | 696 |
val at_calc = thms "Nominal.at_calc"; |
22610 | 697 |
val at_swap_simps = thms "Nominal.at_swap_simps"; |
19494 | 698 |
val at_supp = thm "Nominal.at_supp"; |
699 |
val dj_supp = thm "Nominal.dj_supp"; |
|
700 |
val fresh_left_ineq = thm "Nominal.pt_fresh_left_ineq"; |
|
701 |
val fresh_left = thm "Nominal.pt_fresh_left"; |
|
19548 | 702 |
val fresh_right_ineq = thm "Nominal.pt_fresh_right_ineq"; |
703 |
val fresh_right = thm "Nominal.pt_fresh_right"; |
|
19494 | 704 |
val fresh_bij_ineq = thm "Nominal.pt_fresh_bij_ineq"; |
705 |
val fresh_bij = thm "Nominal.pt_fresh_bij"; |
|
19638
4358b88a9d12
added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents:
19635
diff
changeset
|
706 |
val fresh_aux_ineq = thm "Nominal.pt_fresh_aux_ineq"; |
4358b88a9d12
added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents:
19635
diff
changeset
|
707 |
val fresh_aux = thm "Nominal.pt_fresh_aux"; |
19972
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
708 |
val fresh_eqvt = thm "Nominal.pt_fresh_eqvt"; |
22535
cbee450f88a6
added extended the lemma for equivariance of freshness
urbanc
parents:
22418
diff
changeset
|
709 |
val fresh_eqvt_ineq = thm "Nominal.pt_fresh_eqvt_ineq"; |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
710 |
val set_eqvt = thm "Nominal.pt_set_eqvt"; |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
711 |
val set_diff_eqvt = thm "Nominal.pt_set_diff_eqvt"; |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
712 |
val in_eqvt = thm "Nominal.pt_in_eqvt"; |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
713 |
val eq_eqvt = thm "Nominal.pt_eq_eqvt"; |
19972
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
714 |
val all_eqvt = thm "Nominal.pt_all_eqvt"; |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
715 |
val ex_eqvt = thm "Nominal.pt_ex_eqvt"; |
19494 | 716 |
val pt_pi_rev = thm "Nominal.pt_pi_rev"; |
717 |
val pt_rev_pi = thm "Nominal.pt_rev_pi"; |
|
19972
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
718 |
val at_exists_fresh = thm "Nominal.at_exists_fresh"; |
21377
c29146dc14f1
replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents:
21289
diff
changeset
|
719 |
val at_exists_fresh' = thm "Nominal.at_exists_fresh'"; |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
720 |
|
18068 | 721 |
|
18262 | 722 |
(* Now we collect and instantiate some lemmas w.r.t. all atom *) |
723 |
(* types; this allows for example to use abs_perm (which is a *) |
|
724 |
(* collection of theorems) instead of thm abs_fun_pi with explicit *) |
|
725 |
(* instantiations. *) |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
726 |
val (_, thy33) = |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
727 |
let |
18651
0cb5a8f501aa
added the thms-collection "pt_id" (collection of all pt_<ak>1 lemmas)
urbanc
parents:
18626
diff
changeset
|
728 |
|
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
729 |
(* 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
|
730 |
(* produces a list of theorems of the form [t1 RS thm,..,tn RS thm] *) |
18262 | 731 |
fun instR thm thms = map (fn ti => ti RS thm) thms; |
18068 | 732 |
|
18262 | 733 |
(* takes two theorem lists (hopefully of the same length ;o) *) |
734 |
(* produces a list of theorems of the form *) |
|
735 |
(* [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
|
736 |
fun inst_zip thms1 thms2 = map (fn (t1,t2) => t1 RS t2) (thms1 ~~ thms2); |
18068 | 737 |
|
18262 | 738 |
(* takes a theorem list of the form [l1,...,ln] *) |
739 |
(* and a list of theorem lists of the form *) |
|
740 |
(* [[h11,...,h1m],....,[hk1,....,hkm] *) |
|
741 |
(* produces the list of theorem lists *) |
|
742 |
(* [[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
|
743 |
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
|
744 |
|
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
745 |
(* FIXME: these lists do not need to be created dynamically again *) |
18262 | 746 |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
747 |
|
18068 | 748 |
(* list of all at_inst-theorems *) |
18262 | 749 |
val ats = map (fn ak => PureThy.get_thm thy32 (Name ("at_"^ak^"_inst"))) ak_names |
18068 | 750 |
(* list of all pt_inst-theorems *) |
18262 | 751 |
val pts = map (fn ak => PureThy.get_thm thy32 (Name ("pt_"^ak^"_inst"))) ak_names |
752 |
(* list of all cp_inst-theorems as a collection of lists*) |
|
18068 | 753 |
val cps = |
18262 | 754 |
let fun cps_fun ak1 ak2 = PureThy.get_thm thy32 (Name ("cp_"^ak1^"_"^ak2^"_inst")) |
755 |
in map (fn aki => (map (cps_fun aki) ak_names)) ak_names end; |
|
756 |
(* list of all cp_inst-theorems that have different atom types *) |
|
757 |
val cps' = |
|
758 |
let fun cps'_fun ak1 ak2 = |
|
759 |
if ak1=ak2 then NONE else SOME(PureThy.get_thm thy32 (Name ("cp_"^ak1^"_"^ak2^"_inst"))) |
|
760 |
in map (fn aki => (List.mapPartial I (map (cps'_fun aki) ak_names))) ak_names end; |
|
18068 | 761 |
(* list of all dj_inst-theorems *) |
762 |
val djs = |
|
763 |
let fun djs_fun (ak1,ak2) = |
|
18262 | 764 |
if ak1=ak2 then NONE else SOME(PureThy.get_thm thy32 (Name ("dj_"^ak2^"_"^ak1))) |
765 |
in List.mapPartial I (map djs_fun (Library.product ak_names ak_names)) end; |
|
766 |
(* list of all fs_inst-theorems *) |
|
767 |
val fss = map (fn ak => PureThy.get_thm thy32 (Name ("fs_"^ak^"_inst"))) ak_names |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
768 |
(* list of all at_inst-theorems *) |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
769 |
val fs_axs = map (fn ak => PureThy.get_thm thy32 (Name ("fs_"^ak^"1"))) ak_names |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
770 |
|
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
771 |
fun inst_pt thms = Library.flat (map (fn ti => instR ti pts) thms); |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
772 |
fun inst_at thms = Library.flat (map (fn ti => instR ti ats) thms); |
18262 | 773 |
fun inst_fs thms = Library.flat (map (fn ti => instR ti fss) thms); |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
774 |
fun inst_cp thms cps = Library.flat (inst_mult thms cps); |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
775 |
fun inst_pt_at thms = inst_zip ats (inst_pt thms); |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
776 |
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
|
777 |
fun inst_pt_pt_at_cp thms = inst_cp (inst_zip ats (inst_zip pts (inst_pt thms))) cps; |
18262 | 778 |
fun inst_pt_at_fs thms = inst_zip (inst_fs [fs1]) (inst_zip ats (inst_pt thms)); |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
779 |
fun inst_pt_pt_at_cp thms = |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
780 |
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
|
781 |
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
|
782 |
in i_pt_pt_at_cp end; |
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
783 |
fun inst_pt_pt_at_cp_dj thms = inst_zip djs (inst_pt_pt_at_cp thms); |
18068 | 784 |
in |
18262 | 785 |
thy32 |
18652
3930a060d71b
rolled back the last addition since these lemmas were already
urbanc
parents:
18651
diff
changeset
|
786 |
|> PureThy.add_thmss [(("alpha", inst_pt_at [abs_fun_eq]),[])] |
19562
e56b3c967ae8
added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents:
19548
diff
changeset
|
787 |
||>> PureThy.add_thmss [(("alpha'", inst_pt_at [abs_fun_eq']),[])] |
22557
6775c71f1da0
added pt_bij' to the collection perm_swap and also added abs_perm to the collection of equivariance lemmas
urbanc
parents:
22536
diff
changeset
|
788 |
||>> PureThy.add_thmss [(("perm_swap", inst_pt_at [pt_swap_bij] @ inst_pt_at [pt_swap_bij']),[])] |
22610 | 789 |
||>> PureThy.add_thmss [(("swap_simps", inst_at at_swap_simps),[])] |
19139 | 790 |
||>> PureThy.add_thmss |
791 |
let val thms1 = inst_pt_at [pt_pi_rev]; |
|
792 |
val thms2 = inst_pt_at [pt_rev_pi]; |
|
793 |
in [(("perm_pi_simp",thms1 @ thms2),[])] end |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
794 |
||>> 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
|
795 |
||>> 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
|
796 |
||>> PureThy.add_thmss |
9649e24bc10e
added thms to perm_compose (so far only composition
urbanc
parents:
18435
diff
changeset
|
797 |
let val thms1 = inst_pt_at [pt_perm_compose]; |
9649e24bc10e
added thms to perm_compose (so far only composition
urbanc
parents:
18435
diff
changeset
|
798 |
val thms2 = instR cp1 (Library.flat cps'); |
9649e24bc10e
added thms to perm_compose (so far only composition
urbanc
parents:
18435
diff
changeset
|
799 |
in [(("perm_compose",thms1 @ thms2),[])] end |
19139 | 800 |
||>> PureThy.add_thmss [(("perm_compose'",inst_pt_at [pt_perm_compose']),[])] |
801 |
||>> PureThy.add_thmss [(("perm_app", inst_pt_at [perm_app]),[])] |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
802 |
||>> PureThy.add_thmss [(("supp_atm", (inst_at [at_supp]) @ (inst_dj [dj_supp])),[])] |
19972
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
803 |
||>> PureThy.add_thmss [(("exists_fresh", inst_at [at_exists_fresh]),[])] |
21377
c29146dc14f1
replaced exists_fresh lemma with a version formulated with obtains;
urbanc
parents:
21289
diff
changeset
|
804 |
||>> PureThy.add_thmss [(("exists_fresh'", inst_at [at_exists_fresh']),[])] |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
805 |
||>> PureThy.add_thmss [(("all_eqvt", inst_pt_at [all_eqvt]),[NominalThmDecls.eqvt_add])] |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
806 |
||>> PureThy.add_thmss [(("ex_eqvt", inst_pt_at [ex_eqvt]),[NominalThmDecls.eqvt_add])] |
19972
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
807 |
||>> PureThy.add_thmss |
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
808 |
let val thms1 = inst_at [at_fresh] |
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
809 |
val thms2 = inst_dj [at_fresh_ineq] |
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
810 |
in [(("fresh_atm", thms1 @ thms2),[])] end |
19992 | 811 |
||>> PureThy.add_thmss |
20377
3baf326b2b5f
Removed non-trivial definitions from calc_atm theorem list.
berghofe
parents:
20179
diff
changeset
|
812 |
let val thms1 = filter |
3baf326b2b5f
Removed non-trivial definitions from calc_atm theorem list.
berghofe
parents:
20179
diff
changeset
|
813 |
(fn th => case prop_of th of _ $ _ $ Var _ => true | _ => false) |
3baf326b2b5f
Removed non-trivial definitions from calc_atm theorem list.
berghofe
parents:
20179
diff
changeset
|
814 |
(List.concat (List.concat perm_defs)) |
19993
e0a5783d708f
added simplification rules to the fresh_guess tactic
urbanc
parents:
19992
diff
changeset
|
815 |
in [(("calc_atm", (inst_at at_calc) @ thms1),[])] end |
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
816 |
||>> PureThy.add_thmss |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
817 |
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
|
818 |
and thms2 = inst_pt_pt_at_cp [abs_fun_pi_ineq] |
22557
6775c71f1da0
added pt_bij' to the collection perm_swap and also added abs_perm to the collection of equivariance lemmas
urbanc
parents:
22536
diff
changeset
|
819 |
in [(("abs_perm", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end |
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
820 |
||>> PureThy.add_thmss |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
821 |
let val thms1 = inst_dj [dj_perm_forget] |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
822 |
and thms2 = inst_dj [dj_pp_forget] |
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
823 |
in [(("perm_dj", thms1 @ thms2),[])] end |
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
824 |
||>> PureThy.add_thmss |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
825 |
let val thms1 = inst_pt_at_fs [fresh_iff] |
18626 | 826 |
and thms2 = inst_pt_at [fresh_iff] |
827 |
and thms3 = inst_pt_pt_at_cp_dj [fresh_iff_ineq] |
|
828 |
in [(("abs_fresh", thms1 @ thms2 @ thms3),[])] end |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
829 |
||>> PureThy.add_thmss |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
830 |
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
|
831 |
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
|
832 |
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
|
833 |
in [(("abs_supp", thms1 @ thms2 @ thms3),[])] end |
18396
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
834 |
||>> PureThy.add_thmss |
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
835 |
let val thms1 = inst_pt_at [fresh_left] |
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
836 |
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
|
837 |
in [(("fresh_left", thms1 @ thms2),[])] end |
18426 | 838 |
||>> PureThy.add_thmss |
19548 | 839 |
let val thms1 = inst_pt_at [fresh_right] |
840 |
and thms2 = inst_pt_pt_at_cp [fresh_right_ineq] |
|
841 |
in [(("fresh_right", thms1 @ thms2),[])] end |
|
842 |
||>> PureThy.add_thmss |
|
18426 | 843 |
let val thms1 = inst_pt_at [fresh_bij] |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
844 |
and thms2 = inst_pt_pt_at_cp [fresh_bij_ineq] |
19972
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
845 |
in [(("fresh_bij", thms1 @ thms2),[])] end |
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
846 |
||>> PureThy.add_thmss |
89c5afe4139a
added more infrastructure for the recursion combinator
urbanc
parents:
19638
diff
changeset
|
847 |
let val thms1 = inst_pt_at [fresh_eqvt] |
22535
cbee450f88a6
added extended the lemma for equivariance of freshness
urbanc
parents:
22418
diff
changeset
|
848 |
and thms2 = inst_pt_pt_at_cp_dj [fresh_eqvt_ineq] |
cbee450f88a6
added extended the lemma for equivariance of freshness
urbanc
parents:
22418
diff
changeset
|
849 |
in [(("fresh_eqvt", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
850 |
||>> PureThy.add_thmss |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
851 |
let val thms1 = inst_pt_at [in_eqvt] |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
852 |
in [(("in_eqvt", thms1),[NominalThmDecls.eqvt_add])] end |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
853 |
||>> PureThy.add_thmss |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
854 |
let val thms1 = inst_pt_at [eq_eqvt] |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
855 |
in [(("eq_eqvt", thms1),[NominalThmDecls.eqvt_add])] end |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
856 |
||>> PureThy.add_thmss |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
857 |
let val thms1 = inst_pt [set_eqvt] |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
858 |
in [(("set_eqvt", thms1),[NominalThmDecls.eqvt_add])] end |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
859 |
||>> PureThy.add_thmss |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
860 |
let val thms1 = inst_pt_at [set_diff_eqvt] |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
861 |
in [(("set_diff_eqvt", thms1),[NominalThmDecls.eqvt_add])] end |
19638
4358b88a9d12
added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents:
19635
diff
changeset
|
862 |
||>> PureThy.add_thmss |
4358b88a9d12
added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents:
19635
diff
changeset
|
863 |
let val thms1 = inst_pt_at [fresh_aux] |
4358b88a9d12
added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents:
19635
diff
changeset
|
864 |
and thms2 = inst_pt_pt_at_cp_dj [fresh_aux_ineq] |
4358b88a9d12
added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents:
19635
diff
changeset
|
865 |
in [(("fresh_aux", thms1 @ thms2),[])] end |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
866 |
||>> PureThy.add_thmss [(("fin_supp",fs_axs),[])] |
18068 | 867 |
end; |
868 |
||
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
869 |
in |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
870 |
NominalData.map (fold Symtab.update (full_ak_names ~~ map make_atom_info |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
871 |
(pt_ax_classes ~~ |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
872 |
fs_ax_classes ~~ |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
873 |
map (fn cls => full_ak_names ~~ cls) cp_ax_classes))) thy33 |
18068 | 874 |
end; |
875 |
||
876 |
||
877 |
(* syntax und parsing *) |
|
878 |
structure P = OuterParse and K = OuterKeyword; |
|
879 |
||
880 |
val atom_declP = |
|
881 |
OuterSyntax.command "atom_decl" "Declare new kinds of atoms" K.thy_decl |
|
882 |
(Scan.repeat1 P.name >> (Toplevel.theory o create_nom_typedecls)); |
|
883 |
||
884 |
val _ = OuterSyntax.add_parsers [atom_declP]; |
|
885 |
||
18746
a4ece70964ae
made the change for setup-functions not returning functions
urbanc
parents:
18707
diff
changeset
|
886 |
val setup = NominalData.init; |
18068 | 887 |
|
888 |
end; |