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