author | wenzelm |
Fri, 27 Mar 2015 17:46:08 +0100 | |
changeset 59826 | 442b09c0f898 |
parent 59802 | 684cfaa12e47 |
child 59936 | b8ffc3dc9e24 |
permissions | -rw-r--r-- |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32952
diff
changeset
|
1 |
(* Title: HOL/Nominal/nominal_atoms.ML |
19494 | 2 |
Author: Christian Urban and Stefan Berghofer, TU Muenchen |
3 |
||
4 |
Declaration of atom types to be used in nominal datatypes. |
|
5 |
*) |
|
18068 | 6 |
|
7 |
signature NOMINAL_ATOMS = |
|
8 |
sig |
|
9 |
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
|
10 |
type atom_info |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
11 |
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
|
12 |
val get_atom_info : theory -> string -> atom_info option |
28372 | 13 |
val the_atom_info : theory -> string -> atom_info |
28729
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
14 |
val fs_class_of : theory -> string -> string |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
15 |
val pt_class_of : theory -> string -> string |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
16 |
val cp_class_of : theory -> string -> string -> string |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
17 |
val at_inst_of : theory -> string -> thm |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
18 |
val pt_inst_of : theory -> string -> thm |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
19 |
val cp_inst_of : theory -> string -> string -> thm |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
20 |
val dj_thm_of : theory -> string -> string -> thm |
18068 | 21 |
val atoms_of : theory -> string list |
22 |
val mk_permT : typ -> typ |
|
23 |
end |
|
24 |
||
25 |
structure NominalAtoms : NOMINAL_ATOMS = |
|
26 |
struct |
|
27 |
||
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
28 |
val finite_emptyI = @{thm "finite.emptyI"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
29 |
val Collect_const = @{thm "Collect_const"}; |
21669 | 30 |
|
24569 | 31 |
val inductive_forall_def = @{thm "induct_forall_def"}; |
32 |
||
21669 | 33 |
|
22846 | 34 |
(* theory data *) |
18068 | 35 |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
36 |
type atom_info = |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
37 |
{pt_class : string, |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
38 |
fs_class : string, |
28729
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
39 |
cp_classes : string Symtab.table, |
28372 | 40 |
at_inst : thm, |
41 |
pt_inst : thm, |
|
28729
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
42 |
cp_inst : thm Symtab.table, |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
43 |
dj_thms : thm Symtab.table}; |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
44 |
|
33522 | 45 |
structure NominalData = Theory_Data |
22846 | 46 |
( |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
47 |
type T = atom_info Symtab.table; |
18068 | 48 |
val empty = Symtab.empty; |
49 |
val extend = I; |
|
33522 | 50 |
fun merge data = Symtab.merge (K true) data; |
22846 | 51 |
); |
18068 | 52 |
|
28729
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
53 |
fun make_atom_info ((((((pt_class, fs_class), cp_classes), at_inst), pt_inst), cp_inst), dj_thms) = |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
54 |
{pt_class = pt_class, |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
55 |
fs_class = fs_class, |
28372 | 56 |
cp_classes = cp_classes, |
57 |
at_inst = at_inst, |
|
58 |
pt_inst = pt_inst, |
|
28729
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
59 |
cp_inst = cp_inst, |
28372 | 60 |
dj_thms = dj_thms}; |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
61 |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
62 |
val get_atom_infos = NominalData.get; |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
63 |
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
|
64 |
|
28729
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
65 |
fun gen_lookup lookup name = case lookup name of |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
66 |
SOME info => info |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
67 |
| NONE => error ("Unknown atom type " ^ quote name); |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
68 |
|
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
69 |
fun the_atom_info thy = gen_lookup (get_atom_info thy); |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
70 |
|
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
71 |
fun gen_lookup' f thy = the_atom_info thy #> f; |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
72 |
fun gen_lookup'' f thy = |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
73 |
gen_lookup' (f #> Symtab.lookup #> gen_lookup) thy; |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
74 |
|
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
75 |
val fs_class_of = gen_lookup' #fs_class; |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
76 |
val pt_class_of = gen_lookup' #pt_class; |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
77 |
val at_inst_of = gen_lookup' #at_inst; |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
78 |
val pt_inst_of = gen_lookup' #pt_inst; |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
79 |
val cp_class_of = gen_lookup'' #cp_classes; |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
80 |
val cp_inst_of = gen_lookup'' #cp_inst; |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
81 |
val dj_thm_of = gen_lookup'' #dj_thms; |
28372 | 82 |
|
18068 | 83 |
fun atoms_of thy = map fst (Symtab.dest (NominalData.get thy)); |
84 |
||
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
85 |
fun mk_permT T = HOLogic.listT (HOLogic.mk_prodT (T, T)); |
18068 | 86 |
|
87 |
fun mk_Cons x xs = |
|
88 |
let val T = fastype_of x |
|
56253 | 89 |
in Const (@{const_name Cons}, T --> HOLogic.listT T --> HOLogic.listT T) $ x $ xs end; |
18068 | 90 |
|
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
91 |
fun add_thms_string args = Global_Theory.add_thms ((map o apfst o apfst) Binding.name args); |
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
92 |
fun add_thmss_string args = Global_Theory.add_thmss ((map o apfst o apfst) Binding.name args); |
29585 | 93 |
|
18068 | 94 |
(* this function sets up all matters related to atom- *) |
95 |
(* kinds; the user specifies a list of atom-kind names *) |
|
96 |
(* atom_decl <ak1> ... <akn> *) |
|
97 |
fun create_nom_typedecls ak_names thy = |
|
98 |
let |
|
24527
888d56a8d9d3
modified proofs so that they are not using claset()
urbanc
parents:
24218
diff
changeset
|
99 |
|
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
100 |
val (_,thy1) = |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
101 |
fold_map (fn ak => fn thy => |
45839
43a5b86bc102
'datatype' specifications allow explicit sort constraints;
wenzelm
parents:
45701
diff
changeset
|
102 |
let val dt = ((Binding.name ak, [], NoSyn), [(Binding.name ak, [@{typ nat}], NoSyn)]) |
58354
04ac60da613e
support (finite values of) codatatypes in Quickcheck
blanchet
parents:
58300
diff
changeset
|
103 |
val (dt_names, thy1) = BNF_LFP_Compat.add_datatype [BNF_LFP_Compat.Kill_Type_Args] [dt] thy; |
31783
cfbe9609ceb1
add_datatypes does not yield particular rules any longer
haftmann
parents:
31781
diff
changeset
|
104 |
|
58354
04ac60da613e
support (finite values of) codatatypes in Quickcheck
blanchet
parents:
58300
diff
changeset
|
105 |
val injects = maps (#inject o BNF_LFP_Compat.the_info thy1 []) dt_names; |
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
106 |
val ak_type = Type (Sign.intern_type thy1 ak,[]) |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
107 |
val ak_sign = Sign.intern_const thy1 ak |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
108 |
|
30595
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents:
30345
diff
changeset
|
109 |
val inj_type = @{typ nat} --> ak_type |
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
wenzelm
parents:
30345
diff
changeset
|
110 |
val inj_on_type = inj_type --> @{typ "nat set"} --> @{typ bool} |
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
111 |
|
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
112 |
(* first statement *) |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
113 |
val stmnt1 = HOLogic.mk_Trueprop |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
114 |
(Const (@{const_name "inj_on"},inj_on_type) $ |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
115 |
Const (ak_sign,inj_type) $ HOLogic.mk_UNIV @{typ nat}) |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
116 |
|
31783
cfbe9609ceb1
add_datatypes does not yield particular rules any longer
haftmann
parents:
31781
diff
changeset
|
117 |
val simp1 = @{thm inj_on_def} :: injects; |
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
118 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
119 |
fun proof1 ctxt = EVERY [simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp1) 1, |
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
120 |
rtac @{thm ballI} 1, |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
121 |
rtac @{thm ballI} 1, |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
122 |
rtac @{thm impI} 1, |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
123 |
atac 1] |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
124 |
|
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
125 |
val (inj_thm,thy2) = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
126 |
add_thms_string [((ak^"_inj",Goal.prove_global thy1 [] [] stmnt1 (proof1 o #context)), [])] thy1 |
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
127 |
|
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
128 |
(* second statement *) |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
129 |
val y = Free ("y",ak_type) |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
130 |
val stmnt2 = HOLogic.mk_Trueprop |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
131 |
(HOLogic.mk_exists ("x",@{typ nat},HOLogic.mk_eq (y,Const (ak_sign,inj_type) $ Bound 0))) |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
132 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
133 |
val proof2 = fn {prems, context = ctxt} => |
59826 | 134 |
Induct_Tacs.case_tac ctxt "y" [] NONE 1 THEN |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
135 |
asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp1) 1 THEN |
27128
d2374ba6c02e
InductTacs.case_tac with proper context and proper declaration of local variable;
wenzelm
parents:
27112
diff
changeset
|
136 |
rtac @{thm exI} 1 THEN |
d2374ba6c02e
InductTacs.case_tac with proper context and proper declaration of local variable;
wenzelm
parents:
27112
diff
changeset
|
137 |
rtac @{thm refl} 1 |
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
138 |
|
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
139 |
(* third statement *) |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
140 |
val (inject_thm,thy3) = |
29585 | 141 |
add_thms_string [((ak^"_injection",Goal.prove_global thy2 [] [] stmnt2 proof2), [])] thy2 |
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
142 |
|
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
143 |
val stmnt3 = HOLogic.mk_Trueprop |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
144 |
(HOLogic.mk_not |
56253 | 145 |
(Const (@{const_name finite}, HOLogic.mk_setT ak_type --> HOLogic.boolT) $ |
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
146 |
HOLogic.mk_UNIV ak_type)) |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
147 |
|
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
148 |
val simp2 = [@{thm image_def},@{thm bex_UNIV}]@inject_thm |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
149 |
val simp3 = [@{thm UNIV_def}] |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
150 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
151 |
fun proof3 ctxt = EVERY [cut_facts_tac inj_thm 1, |
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
152 |
dtac @{thm range_inj_infinite} 1, |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
153 |
asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp2) 1, |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
154 |
simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp3) 1] |
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
155 |
|
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
156 |
val (inf_thm,thy4) = |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
157 |
add_thms_string [((ak^"_infinite",Goal.prove_global thy3 [] [] stmnt3 (proof3 o #context)), [])] thy3 |
24677
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
158 |
in |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
159 |
((inj_thm,inject_thm,inf_thm),thy4) |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
160 |
end) ak_names thy |
c6295d2dce48
changed the representation of atoms to datatypes over nats
urbanc
parents:
24569
diff
changeset
|
161 |
|
18068 | 162 |
(* produces a list consisting of pairs: *) |
163 |
(* fst component is the atom-kind name *) |
|
164 |
(* 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
|
165 |
val full_ak_names = map (Sign.intern_type thy1) ak_names; |
18068 | 166 |
val ak_names_types = ak_names ~~ map (Type o rpair []) full_ak_names; |
167 |
||
168 |
(* declares a swapping function for every atom-kind, it is *) |
|
169 |
(* const swap_<ak> :: <akT> * <akT> => <akT> => <akT> *) |
|
170 |
(* swap_<ak> (a,b) c = (if a=c then b (else if b=c then a else c)) *) |
|
171 |
(* overloades then the general swap-function *) |
|
20179 | 172 |
val (swap_eqs, thy3) = fold_map (fn (ak_name, T) => fn thy => |
18068 | 173 |
let |
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
174 |
val thy' = Sign.add_path "rec" thy; |
18068 | 175 |
val swapT = HOLogic.mk_prodT (T, T) --> T --> T; |
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
176 |
val swap_name = "swap_" ^ ak_name; |
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
177 |
val full_swap_name = Sign.full_bname thy' swap_name; |
18068 | 178 |
val a = Free ("a", T); |
179 |
val b = Free ("b", T); |
|
180 |
val c = Free ("c", T); |
|
181 |
val ab = Free ("ab", HOLogic.mk_prodT (T, T)) |
|
56253 | 182 |
val cif = Const (@{const_name If}, HOLogic.boolT --> T --> T --> T); |
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
183 |
val cswap_akname = Const (full_swap_name, swapT); |
56253 | 184 |
val cswap = Const (@{const_name Nominal.swap}, swapT) |
18068 | 185 |
|
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
186 |
val name = swap_name ^ "_def"; |
18068 | 187 |
val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
188 |
(Free (swap_name, swapT) $ HOLogic.mk_prod (a,b) $ c, |
18068 | 189 |
cif $ HOLogic.mk_eq (a,c) $ b $ (cif $ HOLogic.mk_eq (b,c) $ a $ c))) |
190 |
val def2 = Logic.mk_equals (cswap $ ab $ c, cswap_akname $ ab $ c) |
|
191 |
in |
|
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
192 |
thy' |> |
58225 | 193 |
BNF_LFP_Compat.add_primrec_global |
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
194 |
[(Binding.name swap_name, SOME swapT, NoSyn)] |
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
195 |
[(Attrib.empty_binding, def1)] ||> |
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
196 |
Sign.parent_path ||>> |
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
197 |
Global_Theory.add_defs_unchecked true |
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
198 |
[((Binding.name name, def2), [])] |>> (snd o fst) |
26398
fccb74555d9e
removed redundant axiomatizations of XXX_infinite (fact already proven);
wenzelm
parents:
26343
diff
changeset
|
199 |
end) ak_names_types thy1; |
18068 | 200 |
|
201 |
(* declares a permutation function for every atom-kind acting *) |
|
202 |
(* on such atoms *) |
|
203 |
(* const <ak>_prm_<ak> :: (<akT> * <akT>)list => akT => akT *) |
|
204 |
(* <ak>_prm_<ak> [] a = a *) |
|
205 |
(* <ak>_prm_<ak> (x#xs) a = swap_<ak> x (perm xs a) *) |
|
20179 | 206 |
val (prm_eqs, thy4) = fold_map (fn (ak_name, T) => fn thy => |
18068 | 207 |
let |
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
208 |
val thy' = Sign.add_path "rec" thy; |
18068 | 209 |
val swapT = HOLogic.mk_prodT (T, T) --> T --> T; |
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
210 |
val swap_name = Sign.full_bname thy' ("swap_" ^ ak_name) |
18068 | 211 |
val prmT = mk_permT T --> T --> T; |
212 |
val prm_name = ak_name ^ "_prm_" ^ ak_name; |
|
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
213 |
val prm = Free (prm_name, prmT); |
18068 | 214 |
val x = Free ("x", HOLogic.mk_prodT (T, T)); |
215 |
val xs = Free ("xs", mk_permT T); |
|
216 |
val a = Free ("a", T) ; |
|
217 |
||
56253 | 218 |
val cnil = Const (@{const_name Nil}, mk_permT T); |
18068 | 219 |
|
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
220 |
val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq (prm $ cnil $ a, a)); |
18068 | 221 |
|
222 |
val def2 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
223 |
(prm $ mk_Cons x xs $ a, |
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
224 |
Const (swap_name, swapT) $ x $ (prm $ xs $ a))); |
18068 | 225 |
in |
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
226 |
thy' |> |
58225 | 227 |
BNF_LFP_Compat.add_primrec_global |
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
228 |
[(Binding.name prm_name, SOME prmT, NoSyn)] |
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
229 |
[(Attrib.empty_binding, def1), (Attrib.empty_binding, def2)] ||> |
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
230 |
Sign.parent_path |
20179 | 231 |
end) ak_names_types thy3; |
18068 | 232 |
|
233 |
(* defines permutation functions for all combinations of atom-kinds; *) |
|
234 |
(* there are a trivial cases and non-trivial cases *) |
|
235 |
(* non-trivial case: *) |
|
236 |
(* <ak>_prm_<ak>_def: perm pi a == <ak>_prm_<ak> pi a *) |
|
237 |
(* trivial case with <ak> != <ak'> *) |
|
238 |
(* <ak>_prm<ak'>_def[simp]: perm pi a == a *) |
|
239 |
(* *) |
|
240 |
(* the trivial cases are added to the simplifier, while the non- *) |
|
241 |
(* have their own rules proved below *) |
|
18366 | 242 |
val (perm_defs, thy5) = fold_map (fn (ak_name, T) => fn thy => |
243 |
fold_map (fn (ak_name', T') => fn thy' => |
|
18068 | 244 |
let |
245 |
val perm_def_name = ak_name ^ "_prm_" ^ ak_name'; |
|
246 |
val pi = Free ("pi", mk_permT T); |
|
247 |
val a = Free ("a", T'); |
|
56253 | 248 |
val cperm = Const (@{const_name Nominal.perm}, mk_permT T --> T' --> T'); |
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
249 |
val thy'' = Sign.add_path "rec" thy' |
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
250 |
val cperm_def = Const (Sign.full_bname thy'' perm_def_name, mk_permT T --> T' --> T'); |
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
251 |
val thy''' = Sign.parent_path thy''; |
18068 | 252 |
|
253 |
val name = ak_name ^ "_prm_" ^ ak_name' ^ "_def"; |
|
254 |
val def = Logic.mk_equals |
|
255 |
(cperm $ pi $ a, if ak_name = ak_name' then cperm_def $ pi $ a else a) |
|
256 |
in |
|
41562
90fb3d7474df
Finally removed old primrec package, since Primrec.add_primrec_global
berghofe
parents:
41163
diff
changeset
|
257 |
Global_Theory.add_defs_unchecked true [((Binding.name name, def), [])] thy''' |
18366 | 258 |
end) ak_names_types thy) ak_names_types thy4; |
18068 | 259 |
|
260 |
(* proves that every atom-kind is an instance of at *) |
|
261 |
(* lemma at_<ak>_inst: *) |
|
262 |
(* at TYPE(<ak>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
263 |
val (prm_cons_thms,thy6) = |
29585 | 264 |
thy5 |> add_thms_string (map (fn (ak_name, T) => |
18068 | 265 |
let |
28965 | 266 |
val ak_name_qu = Sign.full_bname thy5 (ak_name); |
18068 | 267 |
val i_type = Type(ak_name_qu,[]); |
56253 | 268 |
val cat = Const (@{const_name Nominal.at}, Term.itselfT i_type --> HOLogic.boolT); |
18068 | 269 |
val at_type = Logic.mk_type i_type; |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
270 |
fun proof ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
271 |
simp_tac (put_simpset HOL_ss ctxt |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
272 |
addsimps maps (Global_Theory.get_thms thy5) |
26337 | 273 |
["at_def", |
274 |
ak_name ^ "_prm_" ^ ak_name ^ "_def", |
|
275 |
ak_name ^ "_prm_" ^ ak_name ^ ".simps", |
|
276 |
"swap_" ^ ak_name ^ "_def", |
|
277 |
"swap_" ^ ak_name ^ ".simps", |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
278 |
ak_name ^ "_infinite"]) 1; |
26337 | 279 |
val name = "at_"^ak_name^ "_inst"; |
18068 | 280 |
val statement = HOLogic.mk_Trueprop (cat $ at_type); |
281 |
in |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
282 |
((name, Goal.prove_global thy5 [] [] statement (proof o #context)), []) |
18068 | 283 |
end) ak_names_types); |
284 |
||
285 |
(* declares a perm-axclass for every atom-kind *) |
|
286 |
(* axclass pt_<ak> *) |
|
287 |
(* pt_<ak>1[simp]: perm [] x = x *) |
|
288 |
(* pt_<ak>2: perm (pi1@pi2) x = perm pi1 (perm pi2 x) *) |
|
289 |
(* 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
|
290 |
val (pt_ax_classes,thy7) = fold_map (fn (ak_name, T) => fn thy => |
18068 | 291 |
let |
26337 | 292 |
val cl_name = "pt_"^ak_name; |
56253 | 293 |
val ty = TFree("'a", @{sort type}); |
18068 | 294 |
val x = Free ("x", ty); |
295 |
val pi1 = Free ("pi1", mk_permT T); |
|
296 |
val pi2 = Free ("pi2", mk_permT T); |
|
56253 | 297 |
val cperm = Const (@{const_name Nominal.perm}, mk_permT T --> ty --> ty); |
298 |
val cnil = Const (@{const_name Nil}, mk_permT T); |
|
299 |
val cappend = Const (@{const_name append}, mk_permT T --> mk_permT T --> mk_permT T); |
|
300 |
val cprm_eq = Const (@{const_name Nominal.prm_eq}, mk_permT T --> mk_permT T --> HOLogic.boolT); |
|
18068 | 301 |
(* nil axiom *) |
302 |
val axiom1 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
303 |
(cperm $ cnil $ x, x)); |
|
304 |
(* append axiom *) |
|
305 |
val axiom2 = HOLogic.mk_Trueprop (HOLogic.mk_eq |
|
306 |
(cperm $ (cappend $ pi1 $ pi2) $ x, cperm $ pi1 $ (cperm $ pi2 $ x))); |
|
307 |
(* perm-eq axiom *) |
|
308 |
val axiom3 = Logic.mk_implies |
|
309 |
(HOLogic.mk_Trueprop (cprm_eq $ pi1 $ pi2), |
|
310 |
HOLogic.mk_Trueprop (HOLogic.mk_eq (cperm $ pi1 $ x, cperm $ pi2 $ x))); |
|
311 |
in |
|
56253 | 312 |
Axclass.define_class (Binding.name cl_name, @{sort type}) [] |
28965 | 313 |
[((Binding.name (cl_name ^ "1"), [Simplifier.simp_add]), [axiom1]), |
314 |
((Binding.name (cl_name ^ "2"), []), [axiom2]), |
|
315 |
((Binding.name (cl_name ^ "3"), []), [axiom3])] thy |
|
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
316 |
end) ak_names_types thy6; |
18068 | 317 |
|
318 |
(* proves that every pt_<ak>-type together with <ak>-type *) |
|
319 |
(* instance of pt *) |
|
320 |
(* lemma pt_<ak>_inst: *) |
|
321 |
(* pt TYPE('x::pt_<ak>) TYPE(<ak>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
322 |
val (prm_inst_thms,thy8) = |
29585 | 323 |
thy7 |> add_thms_string (map (fn (ak_name, T) => |
18068 | 324 |
let |
28965 | 325 |
val ak_name_qu = Sign.full_bname thy7 ak_name; |
326 |
val pt_name_qu = Sign.full_bname thy7 ("pt_"^ak_name); |
|
18068 | 327 |
val i_type1 = TFree("'x",[pt_name_qu]); |
328 |
val i_type2 = Type(ak_name_qu,[]); |
|
56253 | 329 |
val cpt = |
330 |
Const (@{const_name Nominal.pt}, (Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); |
|
18068 | 331 |
val pt_type = Logic.mk_type i_type1; |
332 |
val at_type = Logic.mk_type i_type2; |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
333 |
fun proof ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
334 |
simp_tac (put_simpset HOL_ss ctxt addsimps maps (Global_Theory.get_thms thy7) |
26337 | 335 |
["pt_def", |
336 |
"pt_" ^ ak_name ^ "1", |
|
337 |
"pt_" ^ ak_name ^ "2", |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
338 |
"pt_" ^ ak_name ^ "3"]) 1; |
18068 | 339 |
|
26337 | 340 |
val name = "pt_"^ak_name^ "_inst"; |
18068 | 341 |
val statement = HOLogic.mk_Trueprop (cpt $ pt_type $ at_type); |
342 |
in |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
343 |
((name, Goal.prove_global thy7 [] [] statement (proof o #context)), []) |
18068 | 344 |
end) ak_names_types); |
345 |
||
346 |
(* declares an fs-axclass for every atom-kind *) |
|
347 |
(* axclass fs_<ak> *) |
|
348 |
(* 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
|
349 |
val (fs_ax_classes,thy11) = fold_map (fn (ak_name, T) => fn thy => |
18068 | 350 |
let |
26337 | 351 |
val cl_name = "fs_"^ak_name; |
28965 | 352 |
val pt_name = Sign.full_bname thy ("pt_"^ak_name); |
56253 | 353 |
val ty = TFree("'a",@{sort type}); |
18068 | 354 |
val x = Free ("x", ty); |
56253 | 355 |
val csupp = Const (@{const_name Nominal.supp}, ty --> HOLogic.mk_setT T); |
356 |
val cfinite = Const (@{const_name finite}, HOLogic.mk_setT T --> HOLogic.boolT) |
|
18068 | 357 |
|
22274 | 358 |
val axiom1 = HOLogic.mk_Trueprop (cfinite $ (csupp $ x)); |
18068 | 359 |
|
360 |
in |
|
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51671
diff
changeset
|
361 |
Axclass.define_class (Binding.name cl_name, [pt_name]) [] |
30345 | 362 |
[((Binding.name (cl_name ^ "1"), []), [axiom1])] thy |
18438
b8d867177916
made the changes according to Florian's re-arranging of
urbanc
parents:
18437
diff
changeset
|
363 |
end) ak_names_types thy8; |
26337 | 364 |
|
18068 | 365 |
(* proves that every fs_<ak>-type together with <ak>-type *) |
366 |
(* instance of fs-type *) |
|
367 |
(* lemma abst_<ak>_inst: *) |
|
368 |
(* fs TYPE('x::pt_<ak>) TYPE (<ak>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
369 |
val (fs_inst_thms,thy12) = |
29585 | 370 |
thy11 |> add_thms_string (map (fn (ak_name, T) => |
18068 | 371 |
let |
28965 | 372 |
val ak_name_qu = Sign.full_bname thy11 ak_name; |
373 |
val fs_name_qu = Sign.full_bname thy11 ("fs_"^ak_name); |
|
18068 | 374 |
val i_type1 = TFree("'x",[fs_name_qu]); |
375 |
val i_type2 = Type(ak_name_qu,[]); |
|
56253 | 376 |
val cfs = Const (@{const_name Nominal.fs}, |
18068 | 377 |
(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); |
378 |
val fs_type = Logic.mk_type i_type1; |
|
379 |
val at_type = Logic.mk_type i_type2; |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
380 |
fun proof ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
381 |
simp_tac (put_simpset HOL_ss ctxt addsimps maps (Global_Theory.get_thms thy11) |
26337 | 382 |
["fs_def", |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
383 |
"fs_" ^ ak_name ^ "1"]) 1; |
18068 | 384 |
|
26337 | 385 |
val name = "fs_"^ak_name^ "_inst"; |
18068 | 386 |
val statement = HOLogic.mk_Trueprop (cfs $ fs_type $ at_type); |
387 |
in |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
388 |
((name, Goal.prove_global thy11 [] [] statement (proof o #context)), []) |
18068 | 389 |
end) ak_names_types); |
390 |
||
391 |
(* declares for every atom-kind combination an axclass *) |
|
392 |
(* cp_<ak1>_<ak2> giving a composition property *) |
|
393 |
(* 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
|
394 |
val (cp_ax_classes,thy12b) = fold_map (fn (ak_name, T) => fn thy => |
26337 | 395 |
fold_map (fn (ak_name', T') => fn thy' => |
396 |
let |
|
397 |
val cl_name = "cp_"^ak_name^"_"^ak_name'; |
|
56253 | 398 |
val ty = TFree("'a",@{sort type}); |
18068 | 399 |
val x = Free ("x", ty); |
400 |
val pi1 = Free ("pi1", mk_permT T); |
|
26337 | 401 |
val pi2 = Free ("pi2", mk_permT T'); |
56253 | 402 |
val cperm1 = Const (@{const_name Nominal.perm}, mk_permT T --> ty --> ty); |
403 |
val cperm2 = Const (@{const_name Nominal.perm}, mk_permT T' --> ty --> ty); |
|
404 |
val cperm3 = Const (@{const_name Nominal.perm}, mk_permT T --> mk_permT T' --> mk_permT T'); |
|
18068 | 405 |
|
406 |
val ax1 = HOLogic.mk_Trueprop |
|
26337 | 407 |
(HOLogic.mk_eq (cperm1 $ pi1 $ (cperm2 $ pi2 $ x), |
18068 | 408 |
cperm2 $ (cperm3 $ pi1 $ pi2) $ (cperm1 $ pi1 $ x))); |
26337 | 409 |
in |
56253 | 410 |
Axclass.define_class (Binding.name cl_name, @{sort type}) [] |
28965 | 411 |
[((Binding.name (cl_name ^ "1"), []), [ax1])] thy' |
26337 | 412 |
end) ak_names_types thy) ak_names_types thy12; |
18068 | 413 |
|
414 |
(* proves for every <ak>-combination a cp_<ak1>_<ak2>_inst theorem; *) |
|
415 |
(* lemma cp_<ak1>_<ak2>_inst: *) |
|
416 |
(* 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
|
417 |
val (cp_thms,thy12c) = fold_map (fn (ak_name, T) => fn thy => |
26337 | 418 |
fold_map (fn (ak_name', T') => fn thy' => |
18068 | 419 |
let |
28965 | 420 |
val ak_name_qu = Sign.full_bname thy' (ak_name); |
421 |
val ak_name_qu' = Sign.full_bname thy' (ak_name'); |
|
422 |
val cp_name_qu = Sign.full_bname thy' ("cp_"^ak_name^"_"^ak_name'); |
|
18068 | 423 |
val i_type0 = TFree("'a",[cp_name_qu]); |
424 |
val i_type1 = Type(ak_name_qu,[]); |
|
425 |
val i_type2 = Type(ak_name_qu',[]); |
|
56253 | 426 |
val ccp = Const (@{const_name Nominal.cp}, |
18068 | 427 |
(Term.itselfT i_type0)-->(Term.itselfT i_type1)--> |
428 |
(Term.itselfT i_type2)-->HOLogic.boolT); |
|
429 |
val at_type = Logic.mk_type i_type1; |
|
430 |
val at_type' = Logic.mk_type i_type2; |
|
26337 | 431 |
val cp_type = Logic.mk_type i_type0; |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
432 |
val cp1 = Global_Theory.get_thm thy' ("cp_" ^ ak_name ^ "_" ^ ak_name' ^ "1"); |
18068 | 433 |
|
26337 | 434 |
val name = "cp_"^ak_name^ "_"^ak_name'^"_inst"; |
18068 | 435 |
val statement = HOLogic.mk_Trueprop (ccp $ cp_type $ at_type $ at_type'); |
436 |
||
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
437 |
fun proof ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
438 |
simp_tac (put_simpset HOL_basic_ss ctxt |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
439 |
addsimps maps (Global_Theory.get_thms thy') ["cp_def"]) 1 |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
440 |
THEN EVERY [rtac allI 1, rtac allI 1, rtac allI 1, rtac cp1 1]; |
26337 | 441 |
in |
29585 | 442 |
yield_singleton add_thms_string ((name, |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
443 |
Goal.prove_global thy' [] [] statement (proof o #context)), []) thy' |
26337 | 444 |
end) |
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
445 |
ak_names_types thy) ak_names_types thy12b; |
18068 | 446 |
|
447 |
(* proves for every non-trivial <ak>-combination a disjointness *) |
|
448 |
(* theorem; i.e. <ak1> != <ak2> *) |
|
449 |
(* lemma ds_<ak1>_<ak2>: *) |
|
450 |
(* dj TYPE(<ak1>) TYPE(<ak2>) *) |
|
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
451 |
val (dj_thms, thy12d) = fold_map (fn (ak_name,T) => fn thy => |
26337 | 452 |
fold_map (fn (ak_name',T') => fn thy' => |
18068 | 453 |
(if not (ak_name = ak_name') |
454 |
then |
|
26337 | 455 |
let |
28965 | 456 |
val ak_name_qu = Sign.full_bname thy' ak_name; |
457 |
val ak_name_qu' = Sign.full_bname thy' ak_name'; |
|
18068 | 458 |
val i_type1 = Type(ak_name_qu,[]); |
459 |
val i_type2 = Type(ak_name_qu',[]); |
|
56253 | 460 |
val cdj = Const (@{const_name Nominal.disjoint}, |
18068 | 461 |
(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT); |
462 |
val at_type = Logic.mk_type i_type1; |
|
463 |
val at_type' = Logic.mk_type i_type2; |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
464 |
fun proof ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
465 |
simp_tac (put_simpset HOL_ss ctxt |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
466 |
addsimps maps (Global_Theory.get_thms thy') |
26337 | 467 |
["disjoint_def", |
468 |
ak_name ^ "_prm_" ^ ak_name' ^ "_def", |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
469 |
ak_name' ^ "_prm_" ^ ak_name ^ "_def"]) 1; |
18068 | 470 |
|
26337 | 471 |
val name = "dj_"^ak_name^"_"^ak_name'; |
18068 | 472 |
val statement = HOLogic.mk_Trueprop (cdj $ at_type $ at_type'); |
26337 | 473 |
in |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
474 |
add_thms_string [((name, Goal.prove_global thy' [] [] statement (proof o #context)), [])] thy' |
26337 | 475 |
end |
18068 | 476 |
else |
18381
246807ef6dfb
changed the types in accordance with Florian's changes
urbanc
parents:
18366
diff
changeset
|
477 |
([],thy'))) (* do nothing branch, if ak_name = ak_name' *) |
26337 | 478 |
ak_names_types thy) ak_names_types thy12c; |
18068 | 479 |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
480 |
(******** pt_<ak> class instances ********) |
18068 | 481 |
(*=========================================*) |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
482 |
(* some abbreviations for theorems *) |
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
483 |
val pt1 = @{thm "pt1"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
484 |
val pt2 = @{thm "pt2"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
485 |
val pt3 = @{thm "pt3"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
486 |
val at_pt_inst = @{thm "at_pt_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
487 |
val pt_unit_inst = @{thm "pt_unit_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
488 |
val pt_prod_inst = @{thm "pt_prod_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
489 |
val pt_nprod_inst = @{thm "pt_nprod_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
490 |
val pt_list_inst = @{thm "pt_list_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
491 |
val pt_optn_inst = @{thm "pt_option_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
492 |
val pt_noptn_inst = @{thm "pt_noption_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
493 |
val pt_fun_inst = @{thm "pt_fun_inst"}; |
45979 | 494 |
val pt_set_inst = @{thm "pt_set_inst"}; |
18068 | 495 |
|
18435 | 496 |
(* for all atom-kind combinations <ak>/<ak'> show that *) |
497 |
(* every <ak> is an instance of pt_<ak'>; the proof for *) |
|
498 |
(* ak!=ak' is by definition; the case ak=ak' uses at_pt_inst. *) |
|
18431 | 499 |
val thy13 = fold (fn ak_name => fn thy => |
26337 | 500 |
fold (fn ak_name' => fn thy' => |
18431 | 501 |
let |
28965 | 502 |
val qu_name = Sign.full_bname thy' ak_name'; |
503 |
val cls_name = Sign.full_bname thy' ("pt_"^ak_name); |
|
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
504 |
val at_inst = Global_Theory.get_thm thy' ("at_" ^ ak_name' ^ "_inst"); |
18431 | 505 |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
506 |
fun proof1 ctxt = EVERY [Class.intro_classes_tac ctxt [], |
18068 | 507 |
rtac ((at_inst RS at_pt_inst) RS pt1) 1, |
508 |
rtac ((at_inst RS at_pt_inst) RS pt2) 1, |
|
509 |
rtac ((at_inst RS at_pt_inst) RS pt3) 1, |
|
510 |
atac 1]; |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
511 |
fun proof2 ctxt = |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
512 |
Class.intro_classes_tac ctxt [] THEN |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
513 |
REPEAT (asm_simp_tac |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
514 |
(put_simpset HOL_basic_ss ctxt addsimps |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
515 |
maps (Global_Theory.get_thms thy') [ak_name ^ "_prm_" ^ ak_name' ^ "_def"]) 1); |
18431 | 516 |
in |
517 |
thy' |
|
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51671
diff
changeset
|
518 |
|> Axclass.prove_arity (qu_name,[],[cls_name]) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
519 |
(fn ctxt => if ak_name = ak_name' then proof1 ctxt else proof2 ctxt) |
26484 | 520 |
end) ak_names thy) ak_names thy12d; |
18068 | 521 |
|
18430 | 522 |
(* show that *) |
523 |
(* fun(pt_<ak>,pt_<ak>) *) |
|
18579
002d371401f5
changed the name of the type "nOption" to "noption".
urbanc
parents:
18438
diff
changeset
|
524 |
(* noption(pt_<ak>) *) |
18430 | 525 |
(* option(pt_<ak>) *) |
526 |
(* list(pt_<ak>) *) |
|
527 |
(* *(pt_<ak>,pt_<ak>) *) |
|
18600 | 528 |
(* nprod(pt_<ak>,pt_<ak>) *) |
18430 | 529 |
(* unit *) |
530 |
(* set(pt_<ak>) *) |
|
531 |
(* are instances of pt_<ak> *) |
|
18431 | 532 |
val thy18 = fold (fn ak_name => fn thy => |
18068 | 533 |
let |
28965 | 534 |
val cls_name = Sign.full_bname thy ("pt_"^ak_name); |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
535 |
val at_thm = Global_Theory.get_thm thy ("at_"^ak_name^"_inst"); |
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
536 |
val pt_inst = Global_Theory.get_thm thy ("pt_"^ak_name^"_inst"); |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
537 |
|
51671 | 538 |
fun pt_proof thm ctxt = |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
539 |
EVERY [Class.intro_classes_tac ctxt [], |
18430 | 540 |
rtac (thm RS pt1) 1, rtac (thm RS pt2) 1, rtac (thm RS pt3) 1, atac 1]; |
541 |
||
542 |
val pt_thm_fun = at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst)); |
|
46180
72ee700e1d8f
Corrected pt_set_inst, added missing cp_set_inst, deleted obsolete
berghofe
parents:
45979
diff
changeset
|
543 |
val pt_thm_set = pt_inst RS pt_set_inst; |
18430 | 544 |
val pt_thm_noptn = pt_inst RS pt_noptn_inst; |
545 |
val pt_thm_optn = pt_inst RS pt_optn_inst; |
|
546 |
val pt_thm_list = pt_inst RS pt_list_inst; |
|
547 |
val pt_thm_prod = pt_inst RS (pt_inst RS pt_prod_inst); |
|
18600 | 548 |
val pt_thm_nprod = pt_inst RS (pt_inst RS pt_nprod_inst); |
18430 | 549 |
val pt_thm_unit = pt_unit_inst; |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
550 |
in |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
551 |
thy |
56253 | 552 |
|> Axclass.prove_arity (@{type_name fun},[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_fun) |
553 |
|> Axclass.prove_arity (@{type_name set},[[cls_name]],[cls_name]) (pt_proof pt_thm_set) |
|
554 |
|> Axclass.prove_arity (@{type_name noption},[[cls_name]],[cls_name]) (pt_proof pt_thm_noptn) |
|
555 |
|> Axclass.prove_arity (@{type_name option},[[cls_name]],[cls_name]) (pt_proof pt_thm_optn) |
|
556 |
|> Axclass.prove_arity (@{type_name list},[[cls_name]],[cls_name]) (pt_proof pt_thm_list) |
|
557 |
|> Axclass.prove_arity (@{type_name prod},[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_prod) |
|
558 |
|> Axclass.prove_arity (@{type_name nprod},[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_nprod) |
|
559 |
|> Axclass.prove_arity (@{type_name unit},[],[cls_name]) (pt_proof pt_thm_unit) |
|
18430 | 560 |
end) ak_names thy13; |
18068 | 561 |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
562 |
(******** fs_<ak> class instances ********) |
18068 | 563 |
(*=========================================*) |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
564 |
(* abbreviations for some lemmas *) |
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
565 |
val fs1 = @{thm "fs1"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
566 |
val fs_at_inst = @{thm "fs_at_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
567 |
val fs_unit_inst = @{thm "fs_unit_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
568 |
val fs_prod_inst = @{thm "fs_prod_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
569 |
val fs_nprod_inst = @{thm "fs_nprod_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
570 |
val fs_list_inst = @{thm "fs_list_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
571 |
val fs_option_inst = @{thm "fs_option_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
572 |
val dj_supp = @{thm "dj_supp"}; |
18068 | 573 |
|
574 |
(* shows that <ak> is an instance of fs_<ak> *) |
|
575 |
(* uses the theorem at_<ak>_inst *) |
|
18431 | 576 |
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
|
577 |
fold (fn ak_name' => fn thy' => |
18437
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
578 |
let |
28965 | 579 |
val qu_name = Sign.full_bname thy' ak_name'; |
580 |
val qu_class = Sign.full_bname thy' ("fs_"^ak_name); |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
581 |
fun proof ctxt = |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
582 |
(if ak_name = ak_name' |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
583 |
then |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
584 |
let val at_thm = Global_Theory.get_thm thy' ("at_"^ak_name^"_inst"); |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
585 |
in EVERY [Class.intro_classes_tac ctxt [], |
18437
ee0283e5dfe3
added proofs to show that every atom-kind combination
urbanc
parents:
18436
diff
changeset
|
586 |
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
|
587 |
else |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
588 |
let val dj_inst = Global_Theory.get_thm thy' ("dj_"^ak_name'^"_"^ak_name); |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
589 |
val simp_s = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
590 |
put_simpset HOL_basic_ss ctxt addsimps [dj_inst RS dj_supp, finite_emptyI]; |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
591 |
in EVERY [Class.intro_classes_tac ctxt [], asm_simp_tac simp_s 1] end) |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
592 |
in |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
593 |
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
|
594 |
end) ak_names thy) ak_names thy18; |
18068 | 595 |
|
18431 | 596 |
(* shows that *) |
597 |
(* unit *) |
|
598 |
(* *(fs_<ak>,fs_<ak>) *) |
|
18600 | 599 |
(* nprod(fs_<ak>,fs_<ak>) *) |
18431 | 600 |
(* list(fs_<ak>) *) |
601 |
(* option(fs_<ak>) *) |
|
602 |
(* are instances of fs_<ak> *) |
|
18068 | 603 |
|
18431 | 604 |
val thy24 = fold (fn ak_name => fn thy => |
605 |
let |
|
28965 | 606 |
val cls_name = Sign.full_bname thy ("fs_"^ak_name); |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
607 |
val fs_inst = Global_Theory.get_thm thy ("fs_"^ak_name^"_inst"); |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
608 |
fun fs_proof thm ctxt = EVERY [Class.intro_classes_tac ctxt [], rtac (thm RS fs1) 1]; |
18068 | 609 |
|
18600 | 610 |
val fs_thm_unit = fs_unit_inst; |
611 |
val fs_thm_prod = fs_inst RS (fs_inst RS fs_prod_inst); |
|
612 |
val fs_thm_nprod = fs_inst RS (fs_inst RS fs_nprod_inst); |
|
613 |
val fs_thm_list = fs_inst RS fs_list_inst; |
|
614 |
val fs_thm_optn = fs_inst RS fs_option_inst; |
|
18431 | 615 |
in |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
616 |
thy |
56253 | 617 |
|> Axclass.prove_arity (@{type_name unit},[],[cls_name]) (fs_proof fs_thm_unit) |
618 |
|> Axclass.prove_arity (@{type_name prod},[[cls_name],[cls_name]],[cls_name]) (fs_proof fs_thm_prod) |
|
619 |
|> Axclass.prove_arity (@{type_name nprod},[[cls_name],[cls_name]],[cls_name]) (fs_proof fs_thm_nprod) |
|
620 |
|> Axclass.prove_arity (@{type_name list},[[cls_name]],[cls_name]) (fs_proof fs_thm_list) |
|
621 |
|> Axclass.prove_arity (@{type_name 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
|
622 |
end) ak_names thy20; |
18431 | 623 |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
624 |
(******** cp_<ak>_<ai> class instances ********) |
18068 | 625 |
(*==============================================*) |
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
626 |
(* abbreviations for some lemmas *) |
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
627 |
val cp1 = @{thm "cp1"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
628 |
val cp_unit_inst = @{thm "cp_unit_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
629 |
val cp_bool_inst = @{thm "cp_bool_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
630 |
val cp_prod_inst = @{thm "cp_prod_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
631 |
val cp_list_inst = @{thm "cp_list_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
632 |
val cp_fun_inst = @{thm "cp_fun_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
633 |
val cp_option_inst = @{thm "cp_option_inst"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
634 |
val cp_noption_inst = @{thm "cp_noption_inst"}; |
46180
72ee700e1d8f
Corrected pt_set_inst, added missing cp_set_inst, deleted obsolete
berghofe
parents:
45979
diff
changeset
|
635 |
val cp_set_inst = @{thm "cp_set_inst"}; |
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
636 |
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
|
637 |
|
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
638 |
val dj_pp_forget = @{thm "dj_perm_perm_forget"}; |
18068 | 639 |
|
640 |
(* shows that <aj> is an instance of cp_<ak>_<ai> *) |
|
18432 | 641 |
(* for every <ak>/<ai>-combination *) |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
642 |
val thy25 = fold (fn ak_name => fn thy => |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
643 |
fold (fn ak_name' => fn thy' => |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
644 |
fold (fn ak_name'' => fn thy'' => |
18068 | 645 |
let |
28965 | 646 |
val name = Sign.full_bname thy'' ak_name; |
647 |
val cls_name = Sign.full_bname thy'' ("cp_"^ak_name'^"_"^ak_name''); |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
648 |
fun proof ctxt = |
18068 | 649 |
(if (ak_name'=ak_name'') then |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
650 |
(let |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
651 |
val pt_inst = Global_Theory.get_thm thy'' ("pt_"^ak_name''^"_inst"); |
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
652 |
val at_inst = Global_Theory.get_thm thy'' ("at_"^ak_name''^"_inst"); |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
653 |
in |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
654 |
EVERY [Class.intro_classes_tac ctxt [], |
18068 | 655 |
rtac (at_inst RS (pt_inst RS pt_perm_compose)) 1] |
656 |
end) |
|
26337 | 657 |
else |
658 |
(let |
|
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
659 |
val dj_inst = Global_Theory.get_thm thy'' ("dj_"^ak_name''^"_"^ak_name'); |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
660 |
val simp_s = put_simpset HOL_basic_ss ctxt addsimps |
18068 | 661 |
((dj_inst RS dj_pp_forget):: |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
662 |
(maps (Global_Theory.get_thms thy'') |
26337 | 663 |
[ak_name' ^"_prm_"^ak_name^"_def", |
664 |
ak_name''^"_prm_"^ak_name^"_def"])); |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
665 |
in |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
666 |
EVERY [Class.intro_classes_tac ctxt [], simp_tac simp_s 1] |
18068 | 667 |
end)) |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
668 |
in |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
669 |
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
|
670 |
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
|
671 |
|
18432 | 672 |
(* shows that *) |
673 |
(* units *) |
|
674 |
(* products *) |
|
675 |
(* lists *) |
|
676 |
(* functions *) |
|
677 |
(* options *) |
|
678 |
(* noptions *) |
|
22536
35debf264656
made the type sets instance of the "cp" type-class
urbanc
parents:
22535
diff
changeset
|
679 |
(* sets *) |
18432 | 680 |
(* are instances of cp_<ak>_<ai> for every <ak>/<ai>-combination *) |
681 |
val thy26 = fold (fn ak_name => fn thy => |
|
26337 | 682 |
fold (fn ak_name' => fn thy' => |
18432 | 683 |
let |
28965 | 684 |
val cls_name = Sign.full_bname thy' ("cp_"^ak_name^"_"^ak_name'); |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
685 |
val cp_inst = Global_Theory.get_thm thy' ("cp_"^ak_name^"_"^ak_name'^"_inst"); |
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
686 |
val pt_inst = Global_Theory.get_thm thy' ("pt_"^ak_name^"_inst"); |
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
687 |
val at_inst = Global_Theory.get_thm thy' ("at_"^ak_name^"_inst"); |
18432 | 688 |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
689 |
fun cp_proof thm ctxt = EVERY [Class.intro_classes_tac ctxt [], rtac (thm RS cp1) 1]; |
26337 | 690 |
|
18432 | 691 |
val cp_thm_unit = cp_unit_inst; |
692 |
val cp_thm_prod = cp_inst RS (cp_inst RS cp_prod_inst); |
|
693 |
val cp_thm_list = cp_inst RS cp_list_inst; |
|
694 |
val cp_thm_fun = at_inst RS (pt_inst RS (cp_inst RS (cp_inst RS cp_fun_inst))); |
|
695 |
val cp_thm_optn = cp_inst RS cp_option_inst; |
|
696 |
val cp_thm_noptn = cp_inst RS cp_noption_inst; |
|
46180
72ee700e1d8f
Corrected pt_set_inst, added missing cp_set_inst, deleted obsolete
berghofe
parents:
45979
diff
changeset
|
697 |
val cp_thm_set = cp_inst RS cp_set_inst; |
18432 | 698 |
in |
699 |
thy' |
|
56253 | 700 |
|> Axclass.prove_arity (@{type_name unit},[],[cls_name]) (cp_proof cp_thm_unit) |
51685
385ef6706252
more standard module name Axclass (according to file name);
wenzelm
parents:
51671
diff
changeset
|
701 |
|> Axclass.prove_arity (@{type_name Product_Type.prod}, [[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_prod) |
56253 | 702 |
|> Axclass.prove_arity (@{type_name list},[[cls_name]],[cls_name]) (cp_proof cp_thm_list) |
703 |
|> Axclass.prove_arity (@{type_name fun},[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_fun) |
|
704 |
|> Axclass.prove_arity (@{type_name option},[[cls_name]],[cls_name]) (cp_proof cp_thm_optn) |
|
705 |
|> Axclass.prove_arity (@{type_name noption},[[cls_name]],[cls_name]) (cp_proof cp_thm_noptn) |
|
706 |
|> Axclass.prove_arity (@{type_name set},[[cls_name]],[cls_name]) (cp_proof cp_thm_set) |
|
18432 | 707 |
end) ak_names thy) ak_names thy25; |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
708 |
|
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
709 |
(* show that discrete nominal types are permutation types, finitely *) |
18432 | 710 |
(* supported and have the commutation property *) |
711 |
(* 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
|
712 |
(* which renders the proofs to be simple "simp_all"-proofs. *) |
18432 | 713 |
val thy32 = |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
714 |
let |
26337 | 715 |
fun discrete_pt_inst discrete_ty defn = |
716 |
fold (fn ak_name => fn thy => |
|
717 |
let |
|
28965 | 718 |
val qu_class = Sign.full_bname thy ("pt_"^ak_name); |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
719 |
fun proof ctxt = |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
720 |
Class.intro_classes_tac ctxt [] THEN |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
721 |
REPEAT (asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Simpdata.mk_eq defn]) 1); |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
722 |
in |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
723 |
Axclass.prove_arity (discrete_ty, [], [qu_class]) proof thy |
18432 | 724 |
end) ak_names; |
18068 | 725 |
|
18432 | 726 |
fun discrete_fs_inst discrete_ty defn = |
26337 | 727 |
fold (fn ak_name => fn thy => |
728 |
let |
|
28965 | 729 |
val qu_class = Sign.full_bname thy ("fs_"^ak_name); |
44684
8dde3352d5c4
assert Pure equations for theorem references; tuned
haftmann
parents:
41562
diff
changeset
|
730 |
val supp_def = Simpdata.mk_eq @{thm "Nominal.supp_def"}; |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
731 |
fun proof ctxt = |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
732 |
Class.intro_classes_tac ctxt [] THEN |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
733 |
asm_simp_tac (put_simpset HOL_ss ctxt |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
734 |
addsimps [supp_def, Collect_const, finite_emptyI, Simpdata.mk_eq defn]) 1; |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
735 |
in |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
736 |
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
|
737 |
end) ak_names; |
18068 | 738 |
|
18432 | 739 |
fun discrete_cp_inst discrete_ty defn = |
26337 | 740 |
fold (fn ak_name' => (fold (fn ak_name => fn thy => |
741 |
let |
|
28965 | 742 |
val qu_class = Sign.full_bname thy ("cp_"^ak_name^"_"^ak_name'); |
44684
8dde3352d5c4
assert Pure equations for theorem references; tuned
haftmann
parents:
41562
diff
changeset
|
743 |
val supp_def = Simpdata.mk_eq @{thm "Nominal.supp_def"}; |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
744 |
fun proof ctxt = |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58354
diff
changeset
|
745 |
Class.intro_classes_tac ctxt [] THEN |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
746 |
asm_simp_tac (put_simpset HOL_ss ctxt addsimps [Simpdata.mk_eq defn]) 1; |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
747 |
in |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
51685
diff
changeset
|
748 |
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
|
749 |
end) ak_names)) ak_names; |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
750 |
|
18432 | 751 |
in |
752 |
thy26 |
|
44684
8dde3352d5c4
assert Pure equations for theorem references; tuned
haftmann
parents:
41562
diff
changeset
|
753 |
|> discrete_pt_inst @{type_name nat} @{thm perm_nat_def} |
8dde3352d5c4
assert Pure equations for theorem references; tuned
haftmann
parents:
41562
diff
changeset
|
754 |
|> discrete_fs_inst @{type_name nat} @{thm perm_nat_def} |
8dde3352d5c4
assert Pure equations for theorem references; tuned
haftmann
parents:
41562
diff
changeset
|
755 |
|> discrete_cp_inst @{type_name nat} @{thm perm_nat_def} |
44689 | 756 |
|> discrete_pt_inst @{type_name bool} @{thm perm_bool_def} |
757 |
|> discrete_fs_inst @{type_name bool} @{thm perm_bool_def} |
|
758 |
|> discrete_cp_inst @{type_name bool} @{thm perm_bool_def} |
|
44684
8dde3352d5c4
assert Pure equations for theorem references; tuned
haftmann
parents:
41562
diff
changeset
|
759 |
|> discrete_pt_inst @{type_name int} @{thm perm_int_def} |
8dde3352d5c4
assert Pure equations for theorem references; tuned
haftmann
parents:
41562
diff
changeset
|
760 |
|> discrete_fs_inst @{type_name int} @{thm perm_int_def} |
8dde3352d5c4
assert Pure equations for theorem references; tuned
haftmann
parents:
41562
diff
changeset
|
761 |
|> discrete_cp_inst @{type_name int} @{thm perm_int_def} |
8dde3352d5c4
assert Pure equations for theorem references; tuned
haftmann
parents:
41562
diff
changeset
|
762 |
|> discrete_pt_inst @{type_name char} @{thm perm_char_def} |
8dde3352d5c4
assert Pure equations for theorem references; tuned
haftmann
parents:
41562
diff
changeset
|
763 |
|> discrete_fs_inst @{type_name char} @{thm perm_char_def} |
8dde3352d5c4
assert Pure equations for theorem references; tuned
haftmann
parents:
41562
diff
changeset
|
764 |
|> discrete_cp_inst @{type_name char} @{thm perm_char_def} |
18432 | 765 |
end; |
766 |
||
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
767 |
|
18262 | 768 |
(* abbreviations for some lemmas *) |
769 |
(*===============================*) |
|
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
770 |
val abs_fun_pi = @{thm "Nominal.abs_fun_pi"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
771 |
val abs_fun_pi_ineq = @{thm "Nominal.abs_fun_pi_ineq"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
772 |
val abs_fun_eq = @{thm "Nominal.abs_fun_eq"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
773 |
val abs_fun_eq' = @{thm "Nominal.abs_fun_eq'"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
774 |
val abs_fun_fresh = @{thm "Nominal.abs_fun_fresh"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
775 |
val abs_fun_fresh' = @{thm "Nominal.abs_fun_fresh'"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
776 |
val dj_perm_forget = @{thm "Nominal.dj_perm_forget"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
777 |
val dj_pp_forget = @{thm "Nominal.dj_perm_perm_forget"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
778 |
val fresh_iff = @{thm "Nominal.fresh_abs_fun_iff"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
779 |
val fresh_iff_ineq = @{thm "Nominal.fresh_abs_fun_iff_ineq"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
780 |
val abs_fun_supp = @{thm "Nominal.abs_fun_supp"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
781 |
val abs_fun_supp_ineq = @{thm "Nominal.abs_fun_supp_ineq"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
782 |
val pt_swap_bij = @{thm "Nominal.pt_swap_bij"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
783 |
val pt_swap_bij' = @{thm "Nominal.pt_swap_bij'"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
784 |
val pt_fresh_fresh = @{thm "Nominal.pt_fresh_fresh"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
785 |
val pt_bij = @{thm "Nominal.pt_bij"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
786 |
val pt_perm_compose = @{thm "Nominal.pt_perm_compose"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
787 |
val pt_perm_compose' = @{thm "Nominal.pt_perm_compose'"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
788 |
val perm_app = @{thm "Nominal.pt_fun_app_eq"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
789 |
val at_fresh = @{thm "Nominal.at_fresh"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
790 |
val at_fresh_ineq = @{thm "Nominal.at_fresh_ineq"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
791 |
val at_calc = @{thms "Nominal.at_calc"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
792 |
val at_swap_simps = @{thms "Nominal.at_swap_simps"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
793 |
val at_supp = @{thm "Nominal.at_supp"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
794 |
val dj_supp = @{thm "Nominal.dj_supp"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
795 |
val fresh_left_ineq = @{thm "Nominal.pt_fresh_left_ineq"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
796 |
val fresh_left = @{thm "Nominal.pt_fresh_left"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
797 |
val fresh_right_ineq = @{thm "Nominal.pt_fresh_right_ineq"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
798 |
val fresh_right = @{thm "Nominal.pt_fresh_right"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
799 |
val fresh_bij_ineq = @{thm "Nominal.pt_fresh_bij_ineq"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
800 |
val fresh_bij = @{thm "Nominal.pt_fresh_bij"}; |
26773 | 801 |
val fresh_star_bij_ineq = @{thms "Nominal.pt_fresh_star_bij_ineq"}; |
802 |
val fresh_star_bij = @{thms "Nominal.pt_fresh_star_bij"}; |
|
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
803 |
val fresh_eqvt = @{thm "Nominal.pt_fresh_eqvt"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
804 |
val fresh_eqvt_ineq = @{thm "Nominal.pt_fresh_eqvt_ineq"}; |
30086 | 805 |
val fresh_star_eqvt = @{thms "Nominal.pt_fresh_star_eqvt"}; |
806 |
val fresh_star_eqvt_ineq= @{thms "Nominal.pt_fresh_star_eqvt_ineq"}; |
|
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
807 |
val set_diff_eqvt = @{thm "Nominal.pt_set_diff_eqvt"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
808 |
val in_eqvt = @{thm "Nominal.pt_in_eqvt"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
809 |
val eq_eqvt = @{thm "Nominal.pt_eq_eqvt"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
810 |
val all_eqvt = @{thm "Nominal.pt_all_eqvt"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
811 |
val ex_eqvt = @{thm "Nominal.pt_ex_eqvt"}; |
28011 | 812 |
val ex1_eqvt = @{thm "Nominal.pt_ex1_eqvt"}; |
813 |
val the_eqvt = @{thm "Nominal.pt_the_eqvt"}; |
|
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
814 |
val pt_pi_rev = @{thm "Nominal.pt_pi_rev"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
815 |
val pt_rev_pi = @{thm "Nominal.pt_rev_pi"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
816 |
val at_exists_fresh = @{thm "Nominal.at_exists_fresh"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
817 |
val at_exists_fresh' = @{thm "Nominal.at_exists_fresh'"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
818 |
val fresh_perm_app_ineq = @{thm "Nominal.pt_fresh_perm_app_ineq"}; |
26337 | 819 |
val fresh_perm_app = @{thm "Nominal.pt_fresh_perm_app"}; |
23894
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
820 |
val fresh_aux = @{thm "Nominal.pt_fresh_aux"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
821 |
val pt_perm_supp_ineq = @{thm "Nominal.pt_perm_supp_ineq"}; |
1a4167d761ac
tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents:
23158
diff
changeset
|
822 |
val pt_perm_supp = @{thm "Nominal.pt_perm_supp"}; |
26090 | 823 |
val subseteq_eqvt = @{thm "Nominal.pt_subseteq_eqvt"}; |
22786 | 824 |
|
18262 | 825 |
(* Now we collect and instantiate some lemmas w.r.t. all atom *) |
826 |
(* types; this allows for example to use abs_perm (which is a *) |
|
827 |
(* collection of theorems) instead of thm abs_fun_pi with explicit *) |
|
828 |
(* instantiations. *) |
|
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
829 |
val (_, thy33) = |
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
830 |
let |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
831 |
val ctxt32 = Proof_Context.init_global thy32; |
18651
0cb5a8f501aa
added the thms-collection "pt_id" (collection of all pt_<ak>1 lemmas)
urbanc
parents:
18626
diff
changeset
|
832 |
|
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
833 |
(* 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
|
834 |
(* produces a list of theorems of the form [t1 RS thm,..,tn RS thm] *) |
18262 | 835 |
fun instR thm thms = map (fn ti => ti RS thm) thms; |
18068 | 836 |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32952
diff
changeset
|
837 |
(* takes a theorem thm and a list of theorems [(t1,m1),..,(tn,mn)] *) |
26773 | 838 |
(* produces a list of theorems of the form [[t1,m1] MRS thm,..,[tn,mn] MRS thm] *) |
839 |
fun instRR thm thms = map (fn (ti,mi) => [ti,mi] MRS thm) thms; |
|
840 |
||
18262 | 841 |
(* takes two theorem lists (hopefully of the same length ;o) *) |
842 |
(* produces a list of theorems of the form *) |
|
843 |
(* [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
|
844 |
fun inst_zip thms1 thms2 = map (fn (t1,t2) => t1 RS t2) (thms1 ~~ thms2); |
18068 | 845 |
|
18262 | 846 |
(* takes a theorem list of the form [l1,...,ln] *) |
847 |
(* and a list of theorem lists of the form *) |
|
848 |
(* [[h11,...,h1m],....,[hk1,....,hkm] *) |
|
849 |
(* produces the list of theorem lists *) |
|
850 |
(* [[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
|
851 |
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
|
852 |
|
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
853 |
(* FIXME: these lists do not need to be created dynamically again *) |
18262 | 854 |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
855 |
|
18068 | 856 |
(* list of all at_inst-theorems *) |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
857 |
val ats = map (fn ak => Global_Theory.get_thm thy32 ("at_"^ak^"_inst")) ak_names |
18068 | 858 |
(* list of all pt_inst-theorems *) |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
859 |
val pts = map (fn ak => Global_Theory.get_thm thy32 ("pt_"^ak^"_inst")) ak_names |
18262 | 860 |
(* list of all cp_inst-theorems as a collection of lists*) |
18068 | 861 |
val cps = |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
862 |
let fun cps_fun ak1 ak2 = Global_Theory.get_thm thy32 ("cp_"^ak1^"_"^ak2^"_inst") |
26337 | 863 |
in map (fn aki => (map (cps_fun aki) ak_names)) ak_names end; |
18262 | 864 |
(* list of all cp_inst-theorems that have different atom types *) |
865 |
val cps' = |
|
26337 | 866 |
let fun cps'_fun ak1 ak2 = |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
867 |
if ak1=ak2 then NONE else SOME (Global_Theory.get_thm thy32 ("cp_"^ak1^"_"^ak2^"_inst")) |
32952 | 868 |
in map (fn aki => (map_filter I (map (cps'_fun aki) ak_names))) ak_names end; |
18068 | 869 |
(* list of all dj_inst-theorems *) |
870 |
val djs = |
|
26337 | 871 |
let fun djs_fun ak1 ak2 = |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
872 |
if ak1=ak2 then NONE else SOME(Global_Theory.get_thm thy32 ("dj_"^ak2^"_"^ak1)) |
26337 | 873 |
in map_filter I (map_product djs_fun ak_names ak_names) end; |
18262 | 874 |
(* list of all fs_inst-theorems *) |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
875 |
val fss = map (fn ak => Global_Theory.get_thm thy32 ("fs_"^ak^"_inst")) ak_names |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
876 |
(* list of all at_inst-theorems *) |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37678
diff
changeset
|
877 |
val fs_axs = map (fn ak => Global_Theory.get_thm thy32 ("fs_"^ak^"1")) ak_names |
20097
be2d96bbf39c
replaced mk_listT by HOLogic.listT; trivial whitespace/comment changes
webertj
parents:
20046
diff
changeset
|
878 |
|
25538 | 879 |
fun inst_pt thms = maps (fn ti => instR ti pts) thms; |
880 |
fun inst_at thms = maps (fn ti => instR ti ats) thms; |
|
881 |
fun inst_fs thms = maps (fn ti => instR ti fss) thms; |
|
882 |
fun inst_cp thms cps = flat (inst_mult thms cps); |
|
26773 | 883 |
fun inst_pt_at thms = maps (fn ti => instRR ti (pts ~~ ats)) thms; |
25538 | 884 |
fun inst_dj thms = maps (fn ti => instR ti djs) thms; |
26337 | 885 |
fun inst_pt_pt_at_cp thms = inst_cp (inst_zip ats (inst_zip pts (inst_pt thms))) cps; |
18262 | 886 |
fun inst_pt_at_fs thms = inst_zip (inst_fs [fs1]) (inst_zip ats (inst_pt thms)); |
26337 | 887 |
fun inst_pt_pt_at_cp thms = |
888 |
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
|
889 |
val i_pt_pt_at_cp = inst_cp i_pt_pt_at cps'; |
26337 | 890 |
in i_pt_pt_at_cp end; |
18396
b3e7da94b51f
added a fresh_left lemma that contains all instantiation
urbanc
parents:
18381
diff
changeset
|
891 |
fun inst_pt_pt_at_cp_dj thms = inst_zip djs (inst_pt_pt_at_cp thms); |
18068 | 892 |
in |
18262 | 893 |
thy32 |
29585 | 894 |
|> add_thmss_string [(("alpha", inst_pt_at [abs_fun_eq]),[])] |
895 |
||>> add_thmss_string [(("alpha'", inst_pt_at [abs_fun_eq']),[])] |
|
896 |
||>> add_thmss_string [(("alpha_fresh", inst_pt_at [abs_fun_fresh]),[])] |
|
897 |
||>> add_thmss_string [(("alpha_fresh'", inst_pt_at [abs_fun_fresh']),[])] |
|
898 |
||>> add_thmss_string [(("perm_swap", inst_pt_at [pt_swap_bij] @ inst_pt_at [pt_swap_bij']),[])] |
|
899 |
||>> add_thmss_string |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32952
diff
changeset
|
900 |
let val thms1 = inst_at at_swap_simps |
27399
1fb3d1219c12
added facts to lemma swap_simps and tuned lemma calc_atms
urbanc
parents:
27216
diff
changeset
|
901 |
and thms2 = inst_dj [dj_perm_forget] |
1fb3d1219c12
added facts to lemma swap_simps and tuned lemma calc_atms
urbanc
parents:
27216
diff
changeset
|
902 |
in [(("swap_simps", thms1 @ thms2),[])] end |
29585 | 903 |
||>> add_thmss_string |
26337 | 904 |
let val thms1 = inst_pt_at [pt_pi_rev]; |
905 |
val thms2 = inst_pt_at [pt_rev_pi]; |
|
19139 | 906 |
in [(("perm_pi_simp",thms1 @ thms2),[])] end |
29585 | 907 |
||>> add_thmss_string [(("perm_fresh_fresh", inst_pt_at [pt_fresh_fresh]),[])] |
908 |
||>> add_thmss_string [(("perm_bij", inst_pt_at [pt_bij]),[])] |
|
909 |
||>> add_thmss_string |
|
26337 | 910 |
let val thms1 = inst_pt_at [pt_perm_compose]; |
911 |
val thms2 = instR cp1 (Library.flat cps'); |
|
18436
9649e24bc10e
added thms to perm_compose (so far only composition
urbanc
parents:
18435
diff
changeset
|
912 |
in [(("perm_compose",thms1 @ thms2),[])] end |
29585 | 913 |
||>> add_thmss_string [(("perm_compose'",inst_pt_at [pt_perm_compose']),[])] |
914 |
||>> add_thmss_string [(("perm_app", inst_pt_at [perm_app]),[])] |
|
915 |
||>> add_thmss_string [(("supp_atm", (inst_at [at_supp]) @ (inst_dj [dj_supp])),[])] |
|
916 |
||>> add_thmss_string [(("exists_fresh", inst_at [at_exists_fresh]),[])] |
|
917 |
||>> add_thmss_string [(("exists_fresh'", inst_at [at_exists_fresh']),[])] |
|
918 |
||>> add_thmss_string |
|
24569 | 919 |
let |
920 |
val thms1 = inst_pt_at [all_eqvt]; |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
51717
diff
changeset
|
921 |
val thms2 = map (fold_rule ctxt32 [inductive_forall_def]) thms1 |
24569 | 922 |
in |
923 |
[(("all_eqvt", thms1 @ thms2), [NominalThmDecls.eqvt_force_add])] |
|
924 |
end |
|
29585 | 925 |
||>> add_thmss_string [(("ex_eqvt", inst_pt_at [ex_eqvt]),[NominalThmDecls.eqvt_force_add])] |
926 |
||>> add_thmss_string [(("ex1_eqvt", inst_pt_at [ex1_eqvt]),[NominalThmDecls.eqvt_force_add])] |
|
927 |
||>> add_thmss_string [(("the_eqvt", inst_pt_at [the_eqvt]),[NominalThmDecls.eqvt_force_add])] |
|
928 |
||>> add_thmss_string |
|
26337 | 929 |
let val thms1 = inst_at [at_fresh] |
930 |
val thms2 = inst_dj [at_fresh_ineq] |
|
931 |
in [(("fresh_atm", thms1 @ thms2),[])] end |
|
29585 | 932 |
||>> add_thmss_string |
27399
1fb3d1219c12
added facts to lemma swap_simps and tuned lemma calc_atms
urbanc
parents:
27216
diff
changeset
|
933 |
let val thms1 = inst_at at_calc |
1fb3d1219c12
added facts to lemma swap_simps and tuned lemma calc_atms
urbanc
parents:
27216
diff
changeset
|
934 |
and thms2 = inst_dj [dj_perm_forget] |
1fb3d1219c12
added facts to lemma swap_simps and tuned lemma calc_atms
urbanc
parents:
27216
diff
changeset
|
935 |
in [(("calc_atm", thms1 @ thms2),[])] end |
29585 | 936 |
||>> add_thmss_string |
26337 | 937 |
let val thms1 = inst_pt_at [abs_fun_pi] |
938 |
and thms2 = inst_pt_pt_at_cp [abs_fun_pi_ineq] |
|
939 |
in [(("abs_perm", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end |
|
29585 | 940 |
||>> add_thmss_string |
26337 | 941 |
let val thms1 = inst_dj [dj_perm_forget] |
942 |
and thms2 = inst_dj [dj_pp_forget] |
|
18279
f7a18e2b10fc
made some of the theorem look-ups static (by using
urbanc
parents:
18262
diff
changeset
|
943 |
in [(("perm_dj", thms1 @ thms2),[])] end |
29585 | 944 |
||>> add_thmss_string |
26337 | 945 |
let val thms1 = inst_pt_at_fs [fresh_iff] |
18626 | 946 |
and thms2 = inst_pt_at [fresh_iff] |
26337 | 947 |
and thms3 = inst_pt_pt_at_cp_dj [fresh_iff_ineq] |
948 |
in [(("abs_fresh", thms1 @ thms2 @ thms3),[])] end |
|
29585 | 949 |
||>> add_thmss_string |
26337 | 950 |
let val thms1 = inst_pt_at [abs_fun_supp] |
951 |
and thms2 = inst_pt_at_fs [abs_fun_supp] |
|
952 |
and thms3 = inst_pt_pt_at_cp_dj [abs_fun_supp_ineq] |
|
953 |
in [(("abs_supp", thms1 @ thms2 @ thms3),[])] end |
|
29585 | 954 |
||>> add_thmss_string |
26337 | 955 |
let val thms1 = inst_pt_at [fresh_left] |
956 |
and thms2 = inst_pt_pt_at_cp [fresh_left_ineq] |
|
957 |
in [(("fresh_left", thms1 @ thms2),[])] end |
|
29585 | 958 |
||>> add_thmss_string |
26337 | 959 |
let val thms1 = inst_pt_at [fresh_right] |
960 |
and thms2 = inst_pt_pt_at_cp [fresh_right_ineq] |
|
961 |
in [(("fresh_right", thms1 @ thms2),[])] end |
|
29585 | 962 |
||>> add_thmss_string |
26337 | 963 |
let val thms1 = inst_pt_at [fresh_bij] |
964 |
and thms2 = inst_pt_pt_at_cp [fresh_bij_ineq] |
|
965 |
in [(("fresh_bij", thms1 @ thms2),[])] end |
|
29585 | 966 |
||>> add_thmss_string |
26773 | 967 |
let val thms1 = inst_pt_at fresh_star_bij |
30086 | 968 |
and thms2 = maps (fn ti => inst_pt_pt_at_cp [ti]) fresh_star_bij_ineq |
26773 | 969 |
in [(("fresh_star_bij", thms1 @ thms2),[])] end |
29585 | 970 |
||>> add_thmss_string |
26337 | 971 |
let val thms1 = inst_pt_at [fresh_eqvt] |
22535
cbee450f88a6
added extended the lemma for equivariance of freshness
urbanc
parents:
22418
diff
changeset
|
972 |
and thms2 = inst_pt_pt_at_cp_dj [fresh_eqvt_ineq] |
26337 | 973 |
in [(("fresh_eqvt", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end |
29585 | 974 |
||>> add_thmss_string |
30086 | 975 |
let val thms1 = inst_pt_at fresh_star_eqvt |
976 |
and thms2 = maps (fn ti => inst_pt_pt_at_cp_dj [ti]) fresh_star_eqvt_ineq |
|
977 |
in [(("fresh_star_eqvt", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end |
|
978 |
||>> add_thmss_string |
|
26337 | 979 |
let val thms1 = inst_pt_at [in_eqvt] |
980 |
in [(("in_eqvt", thms1),[NominalThmDecls.eqvt_add])] end |
|
29585 | 981 |
||>> add_thmss_string |
26337 | 982 |
let val thms1 = inst_pt_at [eq_eqvt] |
983 |
in [(("eq_eqvt", thms1),[NominalThmDecls.eqvt_add])] end |
|
29585 | 984 |
||>> add_thmss_string |
26337 | 985 |
let val thms1 = inst_pt_at [set_diff_eqvt] |
986 |
in [(("set_diff_eqvt", thms1),[NominalThmDecls.eqvt_add])] end |
|
29585 | 987 |
||>> add_thmss_string |
26337 | 988 |
let val thms1 = inst_pt_at [subseteq_eqvt] |
989 |
in [(("subseteq_eqvt", thms1),[NominalThmDecls.eqvt_add])] end |
|
29585 | 990 |
||>> add_thmss_string |
26337 | 991 |
let val thms1 = inst_pt_at [fresh_aux] |
992 |
and thms2 = inst_pt_pt_at_cp_dj [fresh_perm_app_ineq] |
|
993 |
in [(("fresh_aux", thms1 @ thms2),[])] end |
|
29585 | 994 |
||>> add_thmss_string |
26337 | 995 |
let val thms1 = inst_pt_at [fresh_perm_app] |
996 |
and thms2 = inst_pt_pt_at_cp_dj [fresh_perm_app_ineq] |
|
997 |
in [(("fresh_perm_app", thms1 @ thms2),[])] end |
|
29585 | 998 |
||>> add_thmss_string |
26337 | 999 |
let val thms1 = inst_pt_at [pt_perm_supp] |
1000 |
and thms2 = inst_pt_pt_at_cp [pt_perm_supp_ineq] |
|
1001 |
in [(("supp_eqvt", thms1 @ thms2),[NominalThmDecls.eqvt_add])] end |
|
29585 | 1002 |
||>> add_thmss_string [(("fin_supp",fs_axs),[])] |
26337 | 1003 |
end; |
18068 | 1004 |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
1005 |
in |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
1006 |
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
|
1007 |
(pt_ax_classes ~~ |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22274
diff
changeset
|
1008 |
fs_ax_classes ~~ |
28729
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
1009 |
map (fn cls => Symtab.make (full_ak_names ~~ cls)) cp_ax_classes ~~ |
28372 | 1010 |
prm_cons_thms ~~ prm_inst_thms ~~ |
28729
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
1011 |
map (fn ths => Symtab.make (full_ak_names ~~ ths)) cp_thms ~~ |
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
1012 |
map (fn thss => Symtab.make |
32952 | 1013 |
(map_filter (fn (s, [th]) => SOME (s, th) | _ => NONE) |
28729
cfd169f1dae2
Some more functions for accessing information about atoms.
berghofe
parents:
28372
diff
changeset
|
1014 |
(full_ak_names ~~ thss))) dj_thms))) thy33 |
18068 | 1015 |
end; |
1016 |
||
1017 |
||
1018 |
(* syntax und parsing *) |
|
1019 |
||
24867 | 1020 |
val _ = |
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46180
diff
changeset
|
1021 |
Outer_Syntax.command @{command_spec "atom_decl"} "declare new kinds of atoms" |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
33522
diff
changeset
|
1022 |
(Scan.repeat1 Parse.name >> (Toplevel.theory o create_nom_typedecls)); |
18068 | 1023 |
|
1024 |
end; |