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