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