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