src/HOL/Nominal/nominal_atoms.ML
author urbanc
Tue, 04 Jul 2006 15:57:19 +0200
changeset 19992 bb383dcec3d8
parent 19972 89c5afe4139a
child 19993 e0a5783d708f
permissions -rw-r--r--
made calc_atm stronger by including some relative obvious simplification rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
     1
(*  Title:      HOL/Nominal/nominal_atoms.ML
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
     2
    ID:         $Id$
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
     3
    Author:     Christian Urban and Stefan Berghofer, TU Muenchen
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
     4
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
     5
Declaration of atom types to be used in nominal datatypes.
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
     6
*)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
     7
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
     8
signature NOMINAL_ATOMS =
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
     9
sig
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    10
  val create_nom_typedecls : string list -> theory -> theory
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    11
  val atoms_of : theory -> string list
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    12
  val mk_permT : typ -> typ
18746
a4ece70964ae made the change for setup-functions not returning functions
urbanc
parents: 18707
diff changeset
    13
  val setup : theory -> theory
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    14
end
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    15
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    16
structure NominalAtoms : NOMINAL_ATOMS =
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    17
struct
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    18
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    19
(* data kind 'HOL/nominal' *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    20
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    21
structure NominalArgs =
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    22
struct
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    23
  val name = "HOL/nominal";
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    24
  type T = unit Symtab.table;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    25
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    26
  val empty = Symtab.empty;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    27
  val copy = I;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    28
  val extend = I;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    29
  fun merge _ x = Symtab.merge (K true) x;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    30
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    31
  fun print sg tab = ();
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    32
end;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    33
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    34
structure NominalData = TheoryDataFun(NominalArgs);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    35
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    36
fun atoms_of thy = map fst (Symtab.dest (NominalData.get thy));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    37
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    38
(* FIXME: add to hologic.ML ? *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    39
fun mk_listT T = Type ("List.list", [T]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    40
fun mk_permT T = mk_listT (HOLogic.mk_prodT (T, T));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    41
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    42
fun mk_Cons x xs =
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    43
  let val T = fastype_of x
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    44
  in Const ("List.list.Cons", T --> mk_listT T --> mk_listT T) $ x $ xs end;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    45
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    46
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    47
(* this function sets up all matters related to atom-  *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    48
(* kinds; the user specifies a list of atom-kind names *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    49
(* atom_decl <ak1> ... <akn>                           *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    50
fun create_nom_typedecls ak_names thy =
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    51
  let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    52
    (* declares a type-decl for every atom-kind: *) 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    53
    (* that is typedecl <ak>                     *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    54
    val thy1 = TypedefPackage.add_typedecls (map (fn x => (x,[],NoSyn)) ak_names) thy;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    55
    
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    56
    (* produces a list consisting of pairs:         *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    57
    (*  fst component is the atom-kind name         *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    58
    (*  snd component is its type                   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    59
    val full_ak_names = map (Sign.intern_type (sign_of thy1)) ak_names;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    60
    val ak_names_types = ak_names ~~ map (Type o rpair []) full_ak_names;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    61
     
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    62
    (* adds for every atom-kind an axiom             *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    63
    (* <ak>_infinite: infinite (UNIV::<ak_type> set) *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
    64
    val (inf_axs,thy2) = PureThy.add_axioms_i (map (fn (ak_name, T) =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    65
      let 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    66
	val name = ak_name ^ "_infinite"
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    67
        val axiom = HOLogic.mk_Trueprop (HOLogic.mk_not
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    68
                    (HOLogic.mk_mem (HOLogic.mk_UNIV T,
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    69
                     Const ("Finite_Set.Finites", HOLogic.mk_setT (HOLogic.mk_setT T)))))
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    70
      in
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    71
	((name, axiom), []) 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    72
      end) ak_names_types) thy1;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    73
    
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    74
    (* declares a swapping function for every atom-kind, it is         *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    75
    (* const swap_<ak> :: <akT> * <akT> => <akT> => <akT>              *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    76
    (* swap_<ak> (a,b) c = (if a=c then b (else if b=c then a else c)) *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    77
    (* overloades then the general swap-function                       *) 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    78
    val (thy3, swap_eqs) = foldl_map (fn (thy, (ak_name, T)) =>
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    79
      let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    80
        val swapT = HOLogic.mk_prodT (T, T) --> T --> T;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    81
        val swap_name = Sign.full_name (sign_of thy) ("swap_" ^ ak_name);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    82
        val a = Free ("a", T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    83
        val b = Free ("b", T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    84
        val c = Free ("c", T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    85
        val ab = Free ("ab", HOLogic.mk_prodT (T, T))
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    86
        val cif = Const ("HOL.If", HOLogic.boolT --> T --> T --> T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    87
        val cswap_akname = Const (swap_name, swapT);
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
    88
        val cswap = Const ("Nominal.swap", swapT)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    89
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    90
        val name = "swap_"^ak_name^"_def";
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    91
        val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    92
		   (cswap_akname $ HOLogic.mk_prod (a,b) $ c,
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    93
                    cif $ HOLogic.mk_eq (a,c) $ b $ (cif $ HOLogic.mk_eq (b,c) $ a $ c)))
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    94
        val def2 = Logic.mk_equals (cswap $ ab $ c, cswap_akname $ ab $ c)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    95
      in
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    96
        thy |> Theory.add_consts_i [("swap_" ^ ak_name, swapT, NoSyn)] 
19635
f7aa7d174343 unchecked definitions;
wenzelm
parents: 19562
diff changeset
    97
            |> (#2 o PureThy.add_defs_unchecked_i true [((name, def2),[])])
f7aa7d174343 unchecked definitions;
wenzelm
parents: 19562
diff changeset
    98
            |> PrimrecPackage.add_primrec_unchecked_i "" [(("", def1),[])]            
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
    99
      end) (thy2, ak_names_types);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   100
    
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   101
    (* declares a permutation function for every atom-kind acting  *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   102
    (* on such atoms                                               *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   103
    (* const <ak>_prm_<ak> :: (<akT> * <akT>)list => akT => akT    *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   104
    (* <ak>_prm_<ak> []     a = a                                  *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   105
    (* <ak>_prm_<ak> (x#xs) a = swap_<ak> x (perm xs a)            *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   106
    val (thy4, prm_eqs) = foldl_map (fn (thy, (ak_name, T)) =>
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   107
      let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   108
        val swapT = HOLogic.mk_prodT (T, T) --> T --> T;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   109
        val swap_name = Sign.full_name (sign_of thy) ("swap_" ^ ak_name)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   110
        val prmT = mk_permT T --> T --> T;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   111
        val prm_name = ak_name ^ "_prm_" ^ ak_name;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   112
        val qu_prm_name = Sign.full_name (sign_of thy) prm_name;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   113
        val x  = Free ("x", HOLogic.mk_prodT (T, T));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   114
        val xs = Free ("xs", mk_permT T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   115
        val a  = Free ("a", T) ;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   116
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   117
        val cnil  = Const ("List.list.Nil", mk_permT T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   118
        
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   119
        val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (qu_prm_name, prmT) $ cnil $ a, a));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   120
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   121
        val def2 = HOLogic.mk_Trueprop (HOLogic.mk_eq
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   122
                   (Const (qu_prm_name, prmT) $ mk_Cons x xs $ a,
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   123
                    Const (swap_name, swapT) $ x $ (Const (qu_prm_name, prmT) $ xs $ a)));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   124
      in
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   125
        thy |> Theory.add_consts_i [(prm_name, mk_permT T --> T --> T, NoSyn)] 
19635
f7aa7d174343 unchecked definitions;
wenzelm
parents: 19562
diff changeset
   126
            |> PrimrecPackage.add_primrec_unchecked_i "" [(("", def1), []),(("", def2), [])]
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   127
      end) (thy3, ak_names_types);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   128
    
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   129
    (* defines permutation functions for all combinations of atom-kinds; *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   130
    (* there are a trivial cases and non-trivial cases                   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   131
    (* non-trivial case:                                                 *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   132
    (* <ak>_prm_<ak>_def:  perm pi a == <ak>_prm_<ak> pi a               *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   133
    (* trivial case with <ak> != <ak'>                                   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   134
    (* <ak>_prm<ak'>_def[simp]:  perm pi a == a                          *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   135
    (*                                                                   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   136
    (* the trivial cases are added to the simplifier, while the non-     *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   137
    (* have their own rules proved below                                 *)  
18366
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18355
diff changeset
   138
    val (perm_defs, thy5) = fold_map (fn (ak_name, T) => fn thy =>
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18355
diff changeset
   139
      fold_map (fn (ak_name', T') => fn thy' =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   140
        let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   141
          val perm_def_name = ak_name ^ "_prm_" ^ ak_name';
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   142
          val pi = Free ("pi", mk_permT T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   143
          val a  = Free ("a", T');
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   144
          val cperm = Const ("Nominal.perm", mk_permT T --> T' --> T');
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   145
          val cperm_def = Const (Sign.full_name (sign_of thy') perm_def_name, mk_permT T --> T' --> T');
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   146
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   147
          val name = ak_name ^ "_prm_" ^ ak_name' ^ "_def";
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   148
          val def = Logic.mk_equals
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   149
                    (cperm $ pi $ a, if ak_name = ak_name' then cperm_def $ pi $ a else a)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   150
        in
19635
f7aa7d174343 unchecked definitions;
wenzelm
parents: 19562
diff changeset
   151
          PureThy.add_defs_unchecked_i true [((name, def),[])] thy'
18366
78b4f225b640 Adapted to new type of PureThy.add_defs_i.
berghofe
parents: 18355
diff changeset
   152
        end) ak_names_types thy) ak_names_types thy4;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   153
    
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   154
    (* proves that every atom-kind is an instance of at *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   155
    (* lemma at_<ak>_inst:                              *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   156
    (* at TYPE(<ak>)                                    *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   157
    val (prm_cons_thms,thy6) = 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   158
      thy5 |> PureThy.add_thms (map (fn (ak_name, T) =>
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   159
      let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   160
        val ak_name_qu = Sign.full_name (sign_of thy5) (ak_name);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   161
        val i_type = Type(ak_name_qu,[]);
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   162
	val cat = Const ("Nominal.at",(Term.itselfT i_type)  --> HOLogic.boolT);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   163
        val at_type = Logic.mk_type i_type;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   164
        val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy5
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   165
                                  [Name "at_def",
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   166
                                   Name (ak_name ^ "_prm_" ^ ak_name ^ "_def"),
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   167
                                   Name (ak_name ^ "_prm_" ^ ak_name ^ ".simps"),
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   168
                                   Name ("swap_" ^ ak_name ^ "_def"),
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   169
                                   Name ("swap_" ^ ak_name ^ ".simps"),
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   170
                                   Name (ak_name ^ "_infinite")]
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   171
	    
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   172
	val name = "at_"^ak_name^ "_inst";
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   173
        val statement = HOLogic.mk_Trueprop (cat $ at_type);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   174
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   175
        val proof = fn _ => auto_tac (claset(),simp_s);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   176
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   177
      in 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   178
        ((name, standard (Goal.prove thy5 [] [] statement proof)), []) 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   179
      end) ak_names_types);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   180
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   181
    (* declares a perm-axclass for every atom-kind               *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   182
    (* axclass pt_<ak>                                           *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   183
    (* pt_<ak>1[simp]: perm [] x = x                             *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   184
    (* pt_<ak>2:       perm (pi1@pi2) x = perm pi1 (perm pi2 x)  *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   185
    (* pt_<ak>3:       pi1 ~ pi2 ==> perm pi1 x = perm pi2 x     *)
18438
b8d867177916 made the changes according to Florian's re-arranging of
urbanc
parents: 18437
diff changeset
   186
     val (pt_ax_classes,thy7) =  fold_map (fn (ak_name, T) => fn thy =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   187
      let 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   188
	  val cl_name = "pt_"^ak_name;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   189
          val ty = TFree("'a",["HOL.type"]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   190
          val x   = Free ("x", ty);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   191
          val pi1 = Free ("pi1", mk_permT T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   192
          val pi2 = Free ("pi2", mk_permT T);
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   193
          val cperm = Const ("Nominal.perm", mk_permT T --> ty --> ty);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   194
          val cnil  = Const ("List.list.Nil", mk_permT T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   195
          val cappend = Const ("List.op @",mk_permT T --> mk_permT T --> mk_permT T);
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   196
          val cprm_eq = Const ("Nominal.prm_eq",mk_permT T --> mk_permT T --> HOLogic.boolT);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   197
          (* nil axiom *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   198
          val axiom1 = HOLogic.mk_Trueprop (HOLogic.mk_eq 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   199
                       (cperm $ cnil $ x, x));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   200
          (* append axiom *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   201
          val axiom2 = HOLogic.mk_Trueprop (HOLogic.mk_eq
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   202
                       (cperm $ (cappend $ pi1 $ pi2) $ x, cperm $ pi1 $ (cperm $ pi2 $ x)));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   203
          (* perm-eq axiom *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   204
          val axiom3 = Logic.mk_implies
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   205
                       (HOLogic.mk_Trueprop (cprm_eq $ pi1 $ pi2),
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   206
                        HOLogic.mk_Trueprop (HOLogic.mk_eq (cperm $ pi1 $ x, cperm $ pi2 $ x)));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   207
      in
19509
351e1b1ea251 AxClass.define_class_i;
wenzelm
parents: 19494
diff changeset
   208
          AxClass.define_class_i (cl_name, ["HOL.type"]) []
19488
8bd2c840458e Adapted to new interface of add_axclass_i.
berghofe
parents: 19477
diff changeset
   209
                [((cl_name ^ "1", [Simplifier.simp_add]), [axiom1]),
8bd2c840458e Adapted to new interface of add_axclass_i.
berghofe
parents: 19477
diff changeset
   210
                 ((cl_name ^ "2", []), [axiom2]),                           
8bd2c840458e Adapted to new interface of add_axclass_i.
berghofe
parents: 19477
diff changeset
   211
                 ((cl_name ^ "3", []), [axiom3])] thy                          
18438
b8d867177916 made the changes according to Florian's re-arranging of
urbanc
parents: 18437
diff changeset
   212
      end) ak_names_types thy6;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   213
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   214
    (* proves that every pt_<ak>-type together with <ak>-type *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   215
    (* instance of pt                                         *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   216
    (* lemma pt_<ak>_inst:                                    *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   217
    (* pt TYPE('x::pt_<ak>) TYPE(<ak>)                        *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   218
    val (prm_inst_thms,thy8) = 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   219
      thy7 |> PureThy.add_thms (map (fn (ak_name, T) =>
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   220
      let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   221
        val ak_name_qu = Sign.full_name (sign_of thy7) (ak_name);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   222
        val pt_name_qu = Sign.full_name (sign_of thy7) ("pt_"^ak_name);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   223
        val i_type1 = TFree("'x",[pt_name_qu]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   224
        val i_type2 = Type(ak_name_qu,[]);
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   225
	val cpt = Const ("Nominal.pt",(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   226
        val pt_type = Logic.mk_type i_type1;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   227
        val at_type = Logic.mk_type i_type2;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   228
        val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy7
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   229
                                  [Name "pt_def",
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   230
                                   Name ("pt_" ^ ak_name ^ "1"),
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   231
                                   Name ("pt_" ^ ak_name ^ "2"),
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   232
                                   Name ("pt_" ^ ak_name ^ "3")];
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   233
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   234
	val name = "pt_"^ak_name^ "_inst";
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   235
        val statement = HOLogic.mk_Trueprop (cpt $ pt_type $ at_type);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   236
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   237
        val proof = fn _ => auto_tac (claset(),simp_s);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   238
      in 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   239
        ((name, standard (Goal.prove thy7 [] [] statement proof)), []) 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   240
      end) ak_names_types);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   241
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   242
     (* declares an fs-axclass for every atom-kind       *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   243
     (* axclass fs_<ak>                                  *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   244
     (* fs_<ak>1: finite ((supp x)::<ak> set)            *)
18438
b8d867177916 made the changes according to Florian's re-arranging of
urbanc
parents: 18437
diff changeset
   245
     val (fs_ax_classes,thy11) =  fold_map (fn (ak_name, T) => fn thy =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   246
       let 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   247
	  val cl_name = "fs_"^ak_name;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   248
	  val pt_name = Sign.full_name (sign_of thy) ("pt_"^ak_name);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   249
          val ty = TFree("'a",["HOL.type"]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   250
          val x   = Free ("x", ty);
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   251
          val csupp    = Const ("Nominal.supp", ty --> HOLogic.mk_setT T);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   252
          val cfinites = Const ("Finite_Set.Finites", HOLogic.mk_setT (HOLogic.mk_setT T))
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   253
          
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   254
          val axiom1   = HOLogic.mk_Trueprop (HOLogic.mk_mem (csupp $ x, cfinites));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   255
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   256
       in  
19509
351e1b1ea251 AxClass.define_class_i;
wenzelm
parents: 19494
diff changeset
   257
        AxClass.define_class_i (cl_name, [pt_name]) [] [((cl_name ^ "1", []), [axiom1])] thy            
18438
b8d867177916 made the changes according to Florian's re-arranging of
urbanc
parents: 18437
diff changeset
   258
       end) ak_names_types thy8; 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   259
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   260
     (* proves that every fs_<ak>-type together with <ak>-type   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   261
     (* instance of fs-type                                      *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   262
     (* lemma abst_<ak>_inst:                                    *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   263
     (* fs TYPE('x::pt_<ak>) TYPE (<ak>)                         *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   264
     val (fs_inst_thms,thy12) = 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   265
       thy11 |> PureThy.add_thms (map (fn (ak_name, T) =>
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   266
       let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   267
         val ak_name_qu = Sign.full_name (sign_of thy11) (ak_name);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   268
         val fs_name_qu = Sign.full_name (sign_of thy11) ("fs_"^ak_name);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   269
         val i_type1 = TFree("'x",[fs_name_qu]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   270
         val i_type2 = Type(ak_name_qu,[]);
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   271
 	 val cfs = Const ("Nominal.fs", 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   272
                                 (Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   273
         val fs_type = Logic.mk_type i_type1;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   274
         val at_type = Logic.mk_type i_type2;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   275
	 val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy11
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   276
                                   [Name "fs_def",
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   277
                                    Name ("fs_" ^ ak_name ^ "1")];
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   278
    
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   279
	 val name = "fs_"^ak_name^ "_inst";
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   280
         val statement = HOLogic.mk_Trueprop (cfs $ fs_type $ at_type);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   281
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   282
         val proof = fn _ => auto_tac (claset(),simp_s);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   283
       in 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   284
         ((name, standard (Goal.prove thy11 [] [] statement proof)), []) 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   285
       end) ak_names_types);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   286
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   287
       (* declares for every atom-kind combination an axclass            *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   288
       (* cp_<ak1>_<ak2> giving a composition property                   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   289
       (* cp_<ak1>_<ak2>1: pi1 o pi2 o x = (pi1 o pi2) o (pi1 o x)       *)
18438
b8d867177916 made the changes according to Florian's re-arranging of
urbanc
parents: 18437
diff changeset
   290
        val (_,thy12b) = fold_map (fn (ak_name, T) => fn thy =>
b8d867177916 made the changes according to Florian's re-arranging of
urbanc
parents: 18437
diff changeset
   291
	 fold_map (fn (ak_name', T') => fn thy' =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   292
	     let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   293
	       val cl_name = "cp_"^ak_name^"_"^ak_name';
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   294
	       val ty = TFree("'a",["HOL.type"]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   295
               val x   = Free ("x", ty);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   296
               val pi1 = Free ("pi1", mk_permT T);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   297
	       val pi2 = Free ("pi2", mk_permT T');                  
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   298
	       val cperm1 = Const ("Nominal.perm", mk_permT T  --> ty --> ty);
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   299
               val cperm2 = Const ("Nominal.perm", mk_permT T' --> ty --> ty);
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   300
               val cperm3 = Const ("Nominal.perm", mk_permT T  --> mk_permT T' --> mk_permT T');
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   301
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   302
               val ax1   = HOLogic.mk_Trueprop 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   303
			   (HOLogic.mk_eq (cperm1 $ pi1 $ (cperm2 $ pi2 $ x), 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   304
                                           cperm2 $ (cperm3 $ pi1 $ pi2) $ (cperm1 $ pi1 $ x)));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   305
	       in  
19509
351e1b1ea251 AxClass.define_class_i;
wenzelm
parents: 19494
diff changeset
   306
		 AxClass.define_class_i (cl_name, ["HOL.type"]) [] [((cl_name ^ "1", []), [ax1])] thy'  
18438
b8d867177916 made the changes according to Florian's re-arranging of
urbanc
parents: 18437
diff changeset
   307
	       end) ak_names_types thy) ak_names_types thy12;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   308
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   309
        (* proves for every <ak>-combination a cp_<ak1>_<ak2>_inst theorem;     *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   310
        (* lemma cp_<ak1>_<ak2>_inst:                                           *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   311
        (* cp TYPE('a::cp_<ak1>_<ak2>) TYPE(<ak1>) TYPE(<ak2>)                  *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   312
        val (cp_thms,thy12c) = fold_map (fn (ak_name, T) => fn thy =>
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   313
	 fold_map (fn (ak_name', T') => fn thy' =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   314
           let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   315
             val ak_name_qu  = Sign.full_name (sign_of thy') (ak_name);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   316
	     val ak_name_qu' = Sign.full_name (sign_of thy') (ak_name');
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   317
             val cp_name_qu  = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name');
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   318
             val i_type0 = TFree("'a",[cp_name_qu]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   319
             val i_type1 = Type(ak_name_qu,[]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   320
             val i_type2 = Type(ak_name_qu',[]);
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   321
	     val ccp = Const ("Nominal.cp",
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   322
                             (Term.itselfT i_type0)-->(Term.itselfT i_type1)-->
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   323
                                                      (Term.itselfT i_type2)-->HOLogic.boolT);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   324
             val at_type  = Logic.mk_type i_type1;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   325
             val at_type' = Logic.mk_type i_type2;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   326
	     val cp_type  = Logic.mk_type i_type0;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   327
             val simp_s   = HOL_basic_ss addsimps PureThy.get_thmss thy' [(Name "cp_def")];
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   328
	     val cp1      = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"1"));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   329
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   330
	     val name = "cp_"^ak_name^ "_"^ak_name'^"_inst";
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   331
             val statement = HOLogic.mk_Trueprop (ccp $ cp_type $ at_type $ at_type');
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   332
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   333
             val proof = fn _ => EVERY [auto_tac (claset(),simp_s), rtac cp1 1];
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   334
	   in
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   335
	     PureThy.add_thms [((name, standard (Goal.prove thy' [] [] statement proof)), [])] thy'
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   336
	   end) 
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   337
           ak_names_types thy) ak_names_types thy12b;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   338
       
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   339
        (* proves for every non-trivial <ak>-combination a disjointness   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   340
        (* theorem; i.e. <ak1> != <ak2>                                   *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   341
        (* lemma ds_<ak1>_<ak2>:                                          *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   342
        (* dj TYPE(<ak1>) TYPE(<ak2>)                                     *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   343
        val (dj_thms, thy12d) = fold_map (fn (ak_name,T) => fn thy =>
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   344
	  fold_map (fn (ak_name',T') => fn thy' =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   345
          (if not (ak_name = ak_name') 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   346
           then 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   347
	       let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   348
		 val ak_name_qu  = Sign.full_name (sign_of thy') (ak_name);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   349
	         val ak_name_qu' = Sign.full_name (sign_of thy') (ak_name');
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   350
                 val i_type1 = Type(ak_name_qu,[]);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   351
                 val i_type2 = Type(ak_name_qu',[]);
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   352
	         val cdj = Const ("Nominal.disjoint",
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   353
                           (Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   354
                 val at_type  = Logic.mk_type i_type1;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   355
                 val at_type' = Logic.mk_type i_type2;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   356
                 val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy' 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   357
					   [Name "disjoint_def",
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   358
                                            Name (ak_name^"_prm_"^ak_name'^"_def"),
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   359
                                            Name (ak_name'^"_prm_"^ak_name^"_def")];
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   360
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   361
	         val name = "dj_"^ak_name^"_"^ak_name';
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   362
                 val statement = HOLogic.mk_Trueprop (cdj $ at_type $ at_type');
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   363
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   364
                 val proof = fn _ => auto_tac (claset(),simp_s);
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   365
	       in
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   366
		PureThy.add_thms [((name, standard (Goal.prove thy' [] [] statement proof)), [])] thy'
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   367
	       end
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   368
           else 
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   369
            ([],thy')))  (* do nothing branch, if ak_name = ak_name' *) 
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   370
	    ak_names_types thy) ak_names_types thy12c;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   371
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   372
     (*<<<<<<<  pt_<ak> class instances  >>>>>>>*)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   373
     (*=========================================*)
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   374
     (* some abbreviations for theorems *)
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   375
      val pt1           = thm "pt1";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   376
      val pt2           = thm "pt2";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   377
      val pt3           = thm "pt3";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   378
      val at_pt_inst    = thm "at_pt_inst";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   379
      val pt_set_inst   = thm "pt_set_inst"; 
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   380
      val pt_unit_inst  = thm "pt_unit_inst";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   381
      val pt_prod_inst  = thm "pt_prod_inst"; 
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   382
      val pt_nprod_inst = thm "pt_nprod_inst"; 
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   383
      val pt_list_inst  = thm "pt_list_inst";   
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   384
      val pt_optn_inst  = thm "pt_option_inst";   
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   385
      val pt_noptn_inst = thm "pt_noption_inst";   
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   386
      val pt_fun_inst   = thm "pt_fun_inst";     
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   387
18435
318d2c271040 tuned one comment
urbanc
parents: 18432
diff changeset
   388
     (* for all atom-kind combinations <ak>/<ak'> show that        *)
318d2c271040 tuned one comment
urbanc
parents: 18432
diff changeset
   389
     (* every <ak> is an instance of pt_<ak'>; the proof for       *)
318d2c271040 tuned one comment
urbanc
parents: 18432
diff changeset
   390
     (* ak!=ak' is by definition; the case ak=ak' uses at_pt_inst. *)
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   391
     val thy13 = fold (fn ak_name => fn thy =>
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   392
	fold (fn ak_name' => fn thy' =>
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   393
         let
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   394
           val qu_name =  Sign.full_name (sign_of thy') ak_name';
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   395
           val cls_name = Sign.full_name (sign_of thy') ("pt_"^ak_name);
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   396
           val at_inst  = PureThy.get_thm thy' (Name ("at_"^ak_name'^"_inst")); 
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   397
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   398
           val proof1 = EVERY [ClassPackage.intro_classes_tac [],
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   399
                                 rtac ((at_inst RS at_pt_inst) RS pt1) 1,
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   400
                                 rtac ((at_inst RS at_pt_inst) RS pt2) 1,
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   401
                                 rtac ((at_inst RS at_pt_inst) RS pt3) 1,
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   402
                                 atac 1];
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   403
           val simp_s = HOL_basic_ss addsimps 
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   404
                        PureThy.get_thmss thy' [Name (ak_name^"_prm_"^ak_name'^"_def")];  
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   405
           val proof2 = EVERY [ClassPackage.intro_classes_tac [], REPEAT (asm_simp_tac simp_s 1)];
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   406
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   407
         in
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   408
           thy'
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   409
           |> AxClass.prove_arity (qu_name,[],[cls_name])
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   410
              (if ak_name = ak_name' then proof1 else proof2)
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   411
         end) ak_names thy) ak_names thy12c;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   412
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   413
     (* show that                       *)
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   414
     (*      fun(pt_<ak>,pt_<ak>)       *)
18579
002d371401f5 changed the name of the type "nOption" to "noption".
urbanc
parents: 18438
diff changeset
   415
     (*      noption(pt_<ak>)           *)
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   416
     (*      option(pt_<ak>)            *)
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   417
     (*      list(pt_<ak>)              *)
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   418
     (*      *(pt_<ak>,pt_<ak>)         *)
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   419
     (*      nprod(pt_<ak>,pt_<ak>)     *)
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   420
     (*      unit                       *)
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   421
     (*      set(pt_<ak>)               *)
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   422
     (* are instances of pt_<ak>        *)
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   423
     val thy18 = fold (fn ak_name => fn thy =>
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   424
       let
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   425
          val cls_name = Sign.full_name (sign_of thy) ("pt_"^ak_name);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   426
          val at_thm   = PureThy.get_thm thy (Name ("at_"^ak_name^"_inst"));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   427
          val pt_inst  = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst"));
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   428
          
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   429
          fun pt_proof thm = 
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   430
	      EVERY [ClassPackage.intro_classes_tac [],
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   431
                     rtac (thm RS pt1) 1, rtac (thm RS pt2) 1, rtac (thm RS pt3) 1, atac 1];
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   432
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   433
          val pt_thm_fun   = at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst));
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   434
          val pt_thm_noptn = pt_inst RS pt_noptn_inst; 
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   435
          val pt_thm_optn  = pt_inst RS pt_optn_inst; 
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   436
          val pt_thm_list  = pt_inst RS pt_list_inst;
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   437
          val pt_thm_prod  = pt_inst RS (pt_inst RS pt_prod_inst);
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   438
          val pt_thm_nprod = pt_inst RS (pt_inst RS pt_nprod_inst);
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   439
          val pt_thm_unit  = pt_unit_inst;
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   440
          val pt_thm_set   = pt_inst RS pt_set_inst
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   441
       in 
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   442
	thy
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   443
	|> AxClass.prove_arity ("fun",[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_fun)
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   444
        |> AxClass.prove_arity ("Nominal.noption",[[cls_name]],[cls_name]) (pt_proof pt_thm_noptn) 
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   445
        |> AxClass.prove_arity ("Datatype.option",[[cls_name]],[cls_name]) (pt_proof pt_thm_optn)
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   446
        |> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (pt_proof pt_thm_list)
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   447
        |> AxClass.prove_arity ("*",[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_prod)
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   448
        |> AxClass.prove_arity ("Nominal.nprod",[[cls_name],[cls_name]],[cls_name]) 
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   449
                                    (pt_proof pt_thm_nprod)
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   450
        |> AxClass.prove_arity ("Product_Type.unit",[],[cls_name]) (pt_proof pt_thm_unit)
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   451
        |> AxClass.prove_arity ("set",[[cls_name]],[cls_name]) (pt_proof pt_thm_set)
18430
46c18c0b52c1 improved the code for showing that a type is
urbanc
parents: 18426
diff changeset
   452
     end) ak_names thy13; 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   453
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   454
       (*<<<<<<<  fs_<ak> class instances  >>>>>>>*)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   455
       (*=========================================*)
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   456
       (* abbreviations for some lemmas *)
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   457
       val fs1            = thm "fs1";
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   458
       val fs_at_inst     = thm "fs_at_inst";
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   459
       val fs_unit_inst   = thm "fs_unit_inst";
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   460
       val fs_prod_inst   = thm "fs_prod_inst";
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   461
       val fs_nprod_inst  = thm "fs_nprod_inst";
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   462
       val fs_list_inst   = thm "fs_list_inst";
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   463
       val fs_option_inst = thm "fs_option_inst";
18437
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   464
       val dj_supp        = thm "dj_supp"
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   465
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   466
       (* shows that <ak> is an instance of fs_<ak>     *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   467
       (* uses the theorem at_<ak>_inst                 *)
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   468
       val thy20 = fold (fn ak_name => fn thy =>
18437
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   469
	fold (fn ak_name' => fn thy' => 
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   470
        let
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   471
           val qu_name =  Sign.full_name (sign_of thy') ak_name';
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   472
           val qu_class = Sign.full_name (sign_of thy') ("fs_"^ak_name);
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   473
           val proof = 
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   474
	       (if ak_name = ak_name'
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   475
	        then
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   476
	          let val at_thm = PureThy.get_thm thy' (Name ("at_"^ak_name^"_inst"));
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   477
                  in  EVERY [ClassPackage.intro_classes_tac [],
18437
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   478
                             rtac ((at_thm RS fs_at_inst) RS fs1) 1] end
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   479
                else
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   480
		  let val dj_inst = PureThy.get_thm thy' (Name ("dj_"^ak_name'^"_"^ak_name));
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   481
                      val simp_s = HOL_basic_ss addsimps [dj_inst RS dj_supp, Finites.emptyI]; 
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   482
                  in EVERY [ClassPackage.intro_classes_tac [], asm_simp_tac simp_s 1] end)      
18437
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   483
        in 
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   484
	 AxClass.prove_arity (qu_name,[],[qu_class]) proof thy' 
18437
ee0283e5dfe3 added proofs to show that every atom-kind combination
urbanc
parents: 18436
diff changeset
   485
        end) ak_names thy) ak_names thy18;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   486
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   487
       (* shows that                  *)
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   488
       (*    unit                     *)
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   489
       (*    *(fs_<ak>,fs_<ak>)       *)
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   490
       (*    nprod(fs_<ak>,fs_<ak>)   *)
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   491
       (*    list(fs_<ak>)            *)
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   492
       (*    option(fs_<ak>)          *) 
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   493
       (* are instances of fs_<ak>    *)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   494
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   495
       val thy24 = fold (fn ak_name => fn thy => 
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   496
        let
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   497
          val cls_name = Sign.full_name (sign_of thy) ("fs_"^ak_name);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   498
          val fs_inst  = PureThy.get_thm thy (Name ("fs_"^ak_name^"_inst"));
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   499
          fun fs_proof thm = EVERY [ClassPackage.intro_classes_tac [], rtac (thm RS fs1) 1];      
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   500
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   501
          val fs_thm_unit  = fs_unit_inst;
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   502
          val fs_thm_prod  = fs_inst RS (fs_inst RS fs_prod_inst);
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   503
          val fs_thm_nprod = fs_inst RS (fs_inst RS fs_nprod_inst);
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   504
          val fs_thm_list  = fs_inst RS fs_list_inst;
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   505
          val fs_thm_optn  = fs_inst RS fs_option_inst;
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   506
        in 
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   507
         thy 
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   508
         |> AxClass.prove_arity ("Product_Type.unit",[],[cls_name]) (fs_proof fs_thm_unit) 
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   509
         |> AxClass.prove_arity ("*",[[cls_name],[cls_name]],[cls_name]) (fs_proof fs_thm_prod) 
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   510
         |> AxClass.prove_arity ("Nominal.nprod",[[cls_name],[cls_name]],[cls_name]) 
18600
20ad06db427b added private datatype nprod (copy of prod) to be
urbanc
parents: 18579
diff changeset
   511
                                     (fs_proof fs_thm_nprod) 
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   512
         |> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (fs_proof fs_thm_list)
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   513
         |> AxClass.prove_arity ("Datatype.option",[[cls_name]],[cls_name]) (fs_proof fs_thm_optn)
18431
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   514
        end) ak_names thy20; 
a59c79a3544c improved the finite-support proof
urbanc
parents: 18430
diff changeset
   515
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   516
       (*<<<<<<<  cp_<ak>_<ai> class instances  >>>>>>>*)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   517
       (*==============================================*)
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   518
       (* abbreviations for some lemmas *)
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   519
       val cp1             = thm "cp1";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   520
       val cp_unit_inst    = thm "cp_unit_inst";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   521
       val cp_bool_inst    = thm "cp_bool_inst";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   522
       val cp_prod_inst    = thm "cp_prod_inst";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   523
       val cp_list_inst    = thm "cp_list_inst";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   524
       val cp_fun_inst     = thm "cp_fun_inst";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   525
       val cp_option_inst  = thm "cp_option_inst";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   526
       val cp_noption_inst = thm "cp_noption_inst";
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   527
       val pt_perm_compose = thm "pt_perm_compose";
19477
a95176d0f0dd isar-keywords.el
urbanc
parents: 19275
diff changeset
   528
       
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   529
       val dj_pp_forget    = thm "dj_perm_perm_forget";
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   530
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   531
       (* shows that <aj> is an instance of cp_<ak>_<ai>  *)
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   532
       (* for every  <ak>/<ai>-combination                *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   533
       val thy25 = fold (fn ak_name => fn thy => 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   534
	 fold (fn ak_name' => fn thy' => 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   535
          fold (fn ak_name'' => fn thy'' => 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   536
            let
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   537
              val name =  Sign.full_name (sign_of thy'') ak_name;
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   538
              val cls_name = Sign.full_name (sign_of thy'') ("cp_"^ak_name'^"_"^ak_name'');
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   539
              val proof =
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   540
                (if (ak_name'=ak_name'') then 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   541
		  (let
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   542
                    val pt_inst  = PureThy.get_thm thy'' (Name ("pt_"^ak_name''^"_inst"));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   543
		    val at_inst  = PureThy.get_thm thy'' (Name ("at_"^ak_name''^"_inst"));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   544
                  in 
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   545
		   EVERY [ClassPackage.intro_classes_tac [], 
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   546
                          rtac (at_inst RS (pt_inst RS pt_perm_compose)) 1]
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   547
                  end)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   548
		else
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   549
		  (let 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   550
                     val dj_inst  = PureThy.get_thm thy'' (Name ("dj_"^ak_name''^"_"^ak_name'));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   551
		     val simp_s = HOL_basic_ss addsimps 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   552
                                        ((dj_inst RS dj_pp_forget)::
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   553
                                         (PureThy.get_thmss thy'' 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   554
					   [Name (ak_name' ^"_prm_"^ak_name^"_def"),
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   555
                                            Name (ak_name''^"_prm_"^ak_name^"_def")]));  
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   556
		  in 
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   557
                    EVERY [ClassPackage.intro_classes_tac [], simp_tac simp_s 1]
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   558
                  end))
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   559
	      in
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   560
                AxClass.prove_arity (name,[],[cls_name]) proof thy''
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   561
	      end) ak_names thy') ak_names thy) ak_names thy24;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   562
      
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   563
       (* shows that                                                    *) 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   564
       (*      units                                                    *) 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   565
       (*      products                                                 *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   566
       (*      lists                                                    *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   567
       (*      functions                                                *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   568
       (*      options                                                  *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   569
       (*      noptions                                                 *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   570
       (* are instances of cp_<ak>_<ai> for every <ak>/<ai>-combination *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   571
       val thy26 = fold (fn ak_name => fn thy =>
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   572
	fold (fn ak_name' => fn thy' =>
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   573
        let
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   574
            val cls_name = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name');
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   575
            val cp_inst  = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst"));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   576
            val pt_inst  = PureThy.get_thm thy' (Name ("pt_"^ak_name^"_inst"));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   577
            val at_inst  = PureThy.get_thm thy' (Name ("at_"^ak_name^"_inst"));
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   578
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   579
            fun cp_proof thm  = EVERY [ClassPackage.intro_classes_tac [],rtac (thm RS cp1) 1];     
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   580
	  
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   581
            val cp_thm_unit = cp_unit_inst;
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   582
            val cp_thm_prod = cp_inst RS (cp_inst RS cp_prod_inst);
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   583
            val cp_thm_list = cp_inst RS cp_list_inst;
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   584
            val cp_thm_fun  = at_inst RS (pt_inst RS (cp_inst RS (cp_inst RS cp_fun_inst)));
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   585
            val cp_thm_optn = cp_inst RS cp_option_inst;
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   586
            val cp_thm_noptn = cp_inst RS cp_noption_inst;
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   587
        in
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   588
         thy'
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   589
         |> AxClass.prove_arity ("Product_Type.unit",[],[cls_name]) (cp_proof cp_thm_unit)
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   590
	 |> AxClass.prove_arity ("*",[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_prod)
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   591
         |> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (cp_proof cp_thm_list)
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   592
         |> AxClass.prove_arity ("fun",[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_fun)
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   593
         |> AxClass.prove_arity ("Datatype.option",[[cls_name]],[cls_name]) (cp_proof cp_thm_optn)
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   594
         |> AxClass.prove_arity ("Nominal.noption",[[cls_name]],[cls_name]) (cp_proof cp_thm_noptn)
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   595
        end) ak_names thy) ak_names thy25;
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   596
       
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   597
     (* show that discrete nominal types are permutation types, finitely     *) 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   598
     (* supported and have the commutation property                          *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   599
     (* discrete types have a permutation operation defined as pi o x = x;   *)
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   600
     (* which renders the proofs to be simple "simp_all"-proofs.             *)            
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   601
     val thy32 =
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   602
        let 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   603
	  fun discrete_pt_inst discrete_ty defn = 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   604
	     fold (fn ak_name => fn thy =>
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   605
	     let
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   606
	       val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name);
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   607
	       val simp_s = HOL_basic_ss addsimps [defn];
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   608
               val proof = EVERY [ClassPackage.intro_classes_tac [], REPEAT (asm_simp_tac simp_s 1)];      
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   609
             in  
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   610
	       AxClass.prove_arity (discrete_ty,[],[qu_class]) proof thy
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   611
             end) ak_names;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   612
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   613
          fun discrete_fs_inst discrete_ty defn = 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   614
	     fold (fn ak_name => fn thy =>
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   615
	     let
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   616
	       val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name);
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   617
	       val supp_def = thm "Nominal.supp_def";
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   618
               val simp_s = HOL_ss addsimps [supp_def,Collect_const,Finites.emptyI,defn];
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   619
               val proof = EVERY [ClassPackage.intro_classes_tac [], asm_simp_tac simp_s 1];      
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   620
             in  
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   621
	       AxClass.prove_arity (discrete_ty,[],[qu_class]) proof thy
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   622
             end) ak_names;  
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   623
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   624
          fun discrete_cp_inst discrete_ty defn = 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   625
	     fold (fn ak_name' => (fold (fn ak_name => fn thy =>
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   626
	     let
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   627
	       val qu_class = Sign.full_name (sign_of thy) ("cp_"^ak_name^"_"^ak_name');
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   628
	       val supp_def = thm "Nominal.supp_def";
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   629
               val simp_s = HOL_ss addsimps [defn];
19133
7e84a1a3741c Adapted to Florian's recent changes to the AxClass package.
berghofe
parents: 18759
diff changeset
   630
               val proof = EVERY [ClassPackage.intro_classes_tac [], asm_simp_tac simp_s 1];      
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   631
             in  
19275
3d10de7a8ca7 add_inst_arity_i renamed to prove_arity.
berghofe
parents: 19165
diff changeset
   632
	       AxClass.prove_arity (discrete_ty,[],[qu_class]) proof thy
18432
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   633
             end) ak_names)) ak_names;  
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   634
          
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   635
        in
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   636
         thy26
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   637
         |> discrete_pt_inst "nat"  (thm "perm_nat_def")
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   638
         |> discrete_fs_inst "nat"  (thm "perm_nat_def") 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   639
         |> discrete_cp_inst "nat"  (thm "perm_nat_def") 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   640
         |> discrete_pt_inst "bool" (thm "perm_bool")
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   641
         |> discrete_fs_inst "bool" (thm "perm_bool")
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   642
         |> discrete_cp_inst "bool" (thm "perm_bool")
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   643
         |> discrete_pt_inst "IntDef.int" (thm "perm_int_def")
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   644
         |> discrete_fs_inst "IntDef.int" (thm "perm_int_def") 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   645
         |> discrete_cp_inst "IntDef.int" (thm "perm_int_def") 
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   646
         |> discrete_pt_inst "List.char" (thm "perm_char_def")
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   647
         |> discrete_fs_inst "List.char" (thm "perm_char_def")
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   648
         |> discrete_cp_inst "List.char" (thm "perm_char_def")
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   649
        end;
0b596274ba4f more cleaning up - this time of the cp-instance
urbanc
parents: 18431
diff changeset
   650
19562
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19548
diff changeset
   651
 
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   652
       (* abbreviations for some lemmas *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   653
       (*===============================*)
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   654
       val abs_fun_pi          = thm "Nominal.abs_fun_pi";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   655
       val abs_fun_pi_ineq     = thm "Nominal.abs_fun_pi_ineq";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   656
       val abs_fun_eq          = thm "Nominal.abs_fun_eq";
19562
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19548
diff changeset
   657
       val abs_fun_eq'         = thm "Nominal.abs_fun_eq'";
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   658
       val dj_perm_forget      = thm "Nominal.dj_perm_forget";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   659
       val dj_pp_forget        = thm "Nominal.dj_perm_perm_forget";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   660
       val fresh_iff           = thm "Nominal.fresh_abs_fun_iff";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   661
       val fresh_iff_ineq      = thm "Nominal.fresh_abs_fun_iff_ineq";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   662
       val abs_fun_supp        = thm "Nominal.abs_fun_supp";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   663
       val abs_fun_supp_ineq   = thm "Nominal.abs_fun_supp_ineq";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   664
       val pt_swap_bij         = thm "Nominal.pt_swap_bij";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   665
       val pt_fresh_fresh      = thm "Nominal.pt_fresh_fresh";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   666
       val pt_bij              = thm "Nominal.pt_bij";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   667
       val pt_perm_compose     = thm "Nominal.pt_perm_compose";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   668
       val pt_perm_compose'    = thm "Nominal.pt_perm_compose'";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   669
       val perm_app            = thm "Nominal.pt_fun_app_eq";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   670
       val at_fresh            = thm "Nominal.at_fresh";
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   671
       val at_fresh_ineq       = thm "Nominal.at_fresh_ineq";
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   672
       val at_calc             = thms "Nominal.at_calc";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   673
       val at_supp             = thm "Nominal.at_supp";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   674
       val dj_supp             = thm "Nominal.dj_supp";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   675
       val fresh_left_ineq     = thm "Nominal.pt_fresh_left_ineq";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   676
       val fresh_left          = thm "Nominal.pt_fresh_left";
19548
c0a896828194 added lemma fresh_right, which is useful
urbanc
parents: 19509
diff changeset
   677
       val fresh_right_ineq    = thm "Nominal.pt_fresh_right_ineq";
c0a896828194 added lemma fresh_right, which is useful
urbanc
parents: 19509
diff changeset
   678
       val fresh_right         = thm "Nominal.pt_fresh_right";
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   679
       val fresh_bij_ineq      = thm "Nominal.pt_fresh_bij_ineq";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   680
       val fresh_bij           = thm "Nominal.pt_fresh_bij";
19638
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19635
diff changeset
   681
       val fresh_aux_ineq      = thm "Nominal.pt_fresh_aux_ineq";
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19635
diff changeset
   682
       val fresh_aux           = thm "Nominal.pt_fresh_aux";
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   683
       val fresh_eqvt          = thm "Nominal.pt_fresh_eqvt";
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   684
       val all_eqvt            = thm "Nominal.pt_all_eqvt";
19494
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   685
       val pt_pi_rev           = thm "Nominal.pt_pi_rev";
2e909d5309f4 Renamed "nominal" theory to "Nominal".
berghofe
parents: 19488
diff changeset
   686
       val pt_rev_pi           = thm "Nominal.pt_rev_pi";
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   687
       val at_exists_fresh     = thm "Nominal.at_exists_fresh";
19638
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19635
diff changeset
   688
  
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   689
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   690
       (* Now we collect and instantiate some lemmas w.r.t. all atom      *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   691
       (* types; this allows for example to use abs_perm (which is a      *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   692
       (* collection of theorems) instead of thm abs_fun_pi with explicit *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   693
       (* instantiations.                                                 *)
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   694
       val (_,thy33) = 
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   695
	 let 
18651
0cb5a8f501aa added the thms-collection "pt_id" (collection of all pt_<ak>1 lemmas)
urbanc
parents: 18626
diff changeset
   696
             
0cb5a8f501aa added the thms-collection "pt_id" (collection of all pt_<ak>1 lemmas)
urbanc
parents: 18626
diff changeset
   697
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   698
             (* takes a theorem thm and a list of theorems [t1,..,tn]            *)
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   699
             (* produces a list of theorems of the form [t1 RS thm,..,tn RS thm] *) 
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   700
             fun instR thm thms = map (fn ti => ti RS thm) thms;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   701
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   702
             (* takes two theorem lists (hopefully of the same length ;o)                *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   703
             (* produces a list of theorems of the form                                  *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   704
             (* [t1 RS m1,..,tn RS mn] where [t1,..,tn] is thms1 and [m1,..,mn] is thms2 *) 
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   705
             fun inst_zip thms1 thms2 = map (fn (t1,t2) => t1 RS t2) (thms1 ~~ thms2);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   706
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   707
             (* takes a theorem list of the form [l1,...,ln]              *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   708
             (* and a list of theorem lists of the form                   *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   709
             (* [[h11,...,h1m],....,[hk1,....,hkm]                        *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   710
             (* produces the list of theorem lists                        *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   711
             (* [[l1 RS h11,...,l1 RS h1m],...,[ln RS hk1,...,ln RS hkm]] *)
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   712
             fun inst_mult thms thmss = map (fn (t,ts) => instR t ts) (thms ~~ thmss);
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   713
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   714
             (* FIXME: these lists do not need to be created dynamically again *)
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   715
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   716
             (* list of all at_inst-theorems *)
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   717
             val ats = map (fn ak => PureThy.get_thm thy32 (Name ("at_"^ak^"_inst"))) ak_names
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   718
             (* list of all pt_inst-theorems *)
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   719
             val pts = map (fn ak => PureThy.get_thm thy32 (Name ("pt_"^ak^"_inst"))) ak_names
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   720
             (* list of all cp_inst-theorems as a collection of lists*)
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   721
             val cps = 
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   722
		 let fun cps_fun ak1 ak2 = PureThy.get_thm thy32 (Name ("cp_"^ak1^"_"^ak2^"_inst"))
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   723
		 in map (fn aki => (map (cps_fun aki) ak_names)) ak_names end; 
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   724
             (* list of all cp_inst-theorems that have different atom types *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   725
             val cps' = 
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   726
		let fun cps'_fun ak1 ak2 = 
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   727
		if ak1=ak2 then NONE else SOME(PureThy.get_thm thy32 (Name ("cp_"^ak1^"_"^ak2^"_inst")))
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   728
		in map (fn aki => (List.mapPartial I (map (cps'_fun aki) ak_names))) ak_names end;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   729
             (* list of all dj_inst-theorems *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   730
             val djs = 
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   731
	       let fun djs_fun (ak1,ak2) = 
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   732
		     if ak1=ak2 then NONE else SOME(PureThy.get_thm thy32 (Name ("dj_"^ak2^"_"^ak1)))
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   733
	       in List.mapPartial I (map djs_fun (Library.product ak_names ak_names)) end;
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   734
             (* list of all fs_inst-theorems *)
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   735
             val fss = map (fn ak => PureThy.get_thm thy32 (Name ("fs_"^ak^"_inst"))) ak_names
18651
0cb5a8f501aa added the thms-collection "pt_id" (collection of all pt_<ak>1 lemmas)
urbanc
parents: 18626
diff changeset
   736
              
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   737
             fun inst_pt thms = Library.flat (map (fn ti => instR ti pts) thms); 
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   738
             fun inst_at thms = Library.flat (map (fn ti => instR ti ats) thms);               
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   739
             fun inst_fs thms = Library.flat (map (fn ti => instR ti fss) thms);
18436
9649e24bc10e added thms to perm_compose (so far only composition
urbanc
parents: 18435
diff changeset
   740
             fun inst_cp thms cps = Library.flat (inst_mult thms cps); 
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   741
	     fun inst_pt_at thms = inst_zip ats (inst_pt thms);			
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   742
             fun inst_dj thms = Library.flat (map (fn ti => instR ti djs) thms);  
18436
9649e24bc10e added thms to perm_compose (so far only composition
urbanc
parents: 18435
diff changeset
   743
	     fun inst_pt_pt_at_cp thms = inst_cp (inst_zip ats (inst_zip pts (inst_pt thms))) cps;
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   744
             fun inst_pt_at_fs thms = inst_zip (inst_fs [fs1]) (inst_zip ats (inst_pt thms));
18396
b3e7da94b51f added a fresh_left lemma that contains all instantiation
urbanc
parents: 18381
diff changeset
   745
	     fun inst_pt_pt_at_cp thms = 
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   746
		 let val i_pt_pt_at = inst_zip ats (inst_zip pts (inst_pt thms));
18436
9649e24bc10e added thms to perm_compose (so far only composition
urbanc
parents: 18435
diff changeset
   747
                     val i_pt_pt_at_cp = inst_cp i_pt_pt_at cps';
18396
b3e7da94b51f added a fresh_left lemma that contains all instantiation
urbanc
parents: 18381
diff changeset
   748
		 in i_pt_pt_at_cp end;
b3e7da94b51f added a fresh_left lemma that contains all instantiation
urbanc
parents: 18381
diff changeset
   749
             fun inst_pt_pt_at_cp_dj thms = inst_zip djs (inst_pt_pt_at_cp thms);
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   750
           in
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   751
            thy32 
18652
3930a060d71b rolled back the last addition since these lemmas were already
urbanc
parents: 18651
diff changeset
   752
	    |>   PureThy.add_thmss [(("alpha", inst_pt_at [abs_fun_eq]),[])]
19562
e56b3c967ae8 added the lemma abs_fun_eq' to the nominal theory,
urbanc
parents: 19548
diff changeset
   753
            ||>> PureThy.add_thmss [(("alpha'", inst_pt_at [abs_fun_eq']),[])]
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   754
            ||>> PureThy.add_thmss [(("perm_swap", inst_pt_at [pt_swap_bij]),[])]
19139
af69e41eab5d improved the decision-procedure for permutations;
urbanc
parents: 19134
diff changeset
   755
            ||>> PureThy.add_thmss 
af69e41eab5d improved the decision-procedure for permutations;
urbanc
parents: 19134
diff changeset
   756
	      let val thms1 = inst_pt_at [pt_pi_rev];
af69e41eab5d improved the decision-procedure for permutations;
urbanc
parents: 19134
diff changeset
   757
		  val thms2 = inst_pt_at [pt_rev_pi];
af69e41eab5d improved the decision-procedure for permutations;
urbanc
parents: 19134
diff changeset
   758
              in [(("perm_pi_simp",thms1 @ thms2),[])] end
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   759
            ||>> PureThy.add_thmss [(("perm_fresh_fresh", inst_pt_at [pt_fresh_fresh]),[])]
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   760
            ||>> PureThy.add_thmss [(("perm_bij", inst_pt_at [pt_bij]),[])]
18436
9649e24bc10e added thms to perm_compose (so far only composition
urbanc
parents: 18435
diff changeset
   761
            ||>> PureThy.add_thmss 
9649e24bc10e added thms to perm_compose (so far only composition
urbanc
parents: 18435
diff changeset
   762
	      let val thms1 = inst_pt_at [pt_perm_compose];
9649e24bc10e added thms to perm_compose (so far only composition
urbanc
parents: 18435
diff changeset
   763
		  val thms2 = instR cp1 (Library.flat cps');
9649e24bc10e added thms to perm_compose (so far only composition
urbanc
parents: 18435
diff changeset
   764
              in [(("perm_compose",thms1 @ thms2),[])] end
19139
af69e41eab5d improved the decision-procedure for permutations;
urbanc
parents: 19134
diff changeset
   765
            ||>> PureThy.add_thmss [(("perm_compose'",inst_pt_at [pt_perm_compose']),[])] 
af69e41eab5d improved the decision-procedure for permutations;
urbanc
parents: 19134
diff changeset
   766
            ||>> PureThy.add_thmss [(("perm_app", inst_pt_at [perm_app]),[])]
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   767
            ||>> PureThy.add_thmss [(("supp_atm", (inst_at [at_supp]) @ (inst_dj [dj_supp])),[])]
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   768
            ||>> PureThy.add_thmss [(("exists_fresh", inst_at [at_exists_fresh]),[])]
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   769
            ||>> PureThy.add_thmss [(("all_eqvt", inst_pt_at [all_eqvt]),[])]
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   770
            ||>> PureThy.add_thmss 
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   771
	      let val thms1 = inst_at [at_fresh]
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   772
		  val thms2 = inst_dj [at_fresh_ineq]
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   773
	      in [(("fresh_atm", thms1 @ thms2),[])] end
19992
bb383dcec3d8 made calc_atm stronger by including some relative
urbanc
parents: 19972
diff changeset
   774
            ||>> PureThy.add_thmss
bb383dcec3d8 made calc_atm stronger by including some relative
urbanc
parents: 19972
diff changeset
   775
	      let val thms1 = List.concat (List.concat perm_defs)
bb383dcec3d8 made calc_atm stronger by including some relative
urbanc
parents: 19972
diff changeset
   776
                  val thms2 = List.concat prm_eqs
bb383dcec3d8 made calc_atm stronger by including some relative
urbanc
parents: 19972
diff changeset
   777
                  val thms3 = List.concat swap_eqs
bb383dcec3d8 made calc_atm stronger by including some relative
urbanc
parents: 19972
diff changeset
   778
              in [(("calc_atm", (inst_at at_calc) @ thms1 @ thms2 @ thms3 ),[])] end
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   779
            ||>> PureThy.add_thmss
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   780
	      let val thms1 = inst_pt_at [abs_fun_pi]
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   781
		  and thms2 = inst_pt_pt_at_cp [abs_fun_pi_ineq]
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   782
	      in [(("abs_perm", thms1 @ thms2),[])] end
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   783
            ||>> PureThy.add_thmss
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   784
	      let val thms1 = inst_dj [dj_perm_forget]
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   785
		  and thms2 = inst_dj [dj_pp_forget]
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   786
              in [(("perm_dj", thms1 @ thms2),[])] end
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   787
            ||>> PureThy.add_thmss
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   788
	      let val thms1 = inst_pt_at_fs [fresh_iff]
18626
b6596f579e40 added some lemmas to the collection "abs_fresh"
urbanc
parents: 18600
diff changeset
   789
                  and thms2 = inst_pt_at [fresh_iff]
b6596f579e40 added some lemmas to the collection "abs_fresh"
urbanc
parents: 18600
diff changeset
   790
		  and thms3 = inst_pt_pt_at_cp_dj [fresh_iff_ineq]
b6596f579e40 added some lemmas to the collection "abs_fresh"
urbanc
parents: 18600
diff changeset
   791
	    in [(("abs_fresh", thms1 @ thms2 @ thms3),[])] end
18381
246807ef6dfb changed the types in accordance with Florian's changes
urbanc
parents: 18366
diff changeset
   792
	    ||>> PureThy.add_thmss
18279
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   793
	      let val thms1 = inst_pt_at [abs_fun_supp]
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   794
		  and thms2 = inst_pt_at_fs [abs_fun_supp]
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   795
		  and thms3 = inst_pt_pt_at_cp_dj [abs_fun_supp_ineq]
f7a18e2b10fc made some of the theorem look-ups static (by using
urbanc
parents: 18262
diff changeset
   796
	      in [(("abs_supp", thms1 @ thms2 @ thms3),[])] end
18396
b3e7da94b51f added a fresh_left lemma that contains all instantiation
urbanc
parents: 18381
diff changeset
   797
            ||>> PureThy.add_thmss
b3e7da94b51f added a fresh_left lemma that contains all instantiation
urbanc
parents: 18381
diff changeset
   798
	      let val thms1 = inst_pt_at [fresh_left]
b3e7da94b51f added a fresh_left lemma that contains all instantiation
urbanc
parents: 18381
diff changeset
   799
		  and thms2 = inst_pt_pt_at_cp [fresh_left_ineq]
b3e7da94b51f added a fresh_left lemma that contains all instantiation
urbanc
parents: 18381
diff changeset
   800
	      in [(("fresh_left", thms1 @ thms2),[])] end
18426
d2303e8654a2 added container-lemma fresh_eqvt
urbanc
parents: 18396
diff changeset
   801
            ||>> PureThy.add_thmss
19548
c0a896828194 added lemma fresh_right, which is useful
urbanc
parents: 19509
diff changeset
   802
	      let val thms1 = inst_pt_at [fresh_right]
c0a896828194 added lemma fresh_right, which is useful
urbanc
parents: 19509
diff changeset
   803
		  and thms2 = inst_pt_pt_at_cp [fresh_right_ineq]
c0a896828194 added lemma fresh_right, which is useful
urbanc
parents: 19509
diff changeset
   804
	      in [(("fresh_right", thms1 @ thms2),[])] end
c0a896828194 added lemma fresh_right, which is useful
urbanc
parents: 19509
diff changeset
   805
            ||>> PureThy.add_thmss
18426
d2303e8654a2 added container-lemma fresh_eqvt
urbanc
parents: 18396
diff changeset
   806
	      let val thms1 = inst_pt_at [fresh_bij]
d2303e8654a2 added container-lemma fresh_eqvt
urbanc
parents: 18396
diff changeset
   807
		  and thms2 = inst_pt_pt_at_cp [fresh_bij_ineq]
19972
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   808
	      in [(("fresh_bij", thms1 @ thms2),[])] end
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   809
            ||>> PureThy.add_thmss
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   810
	      let val thms1 = inst_pt_at [fresh_eqvt]
89c5afe4139a added more infrastructure for the recursion combinator
urbanc
parents: 19638
diff changeset
   811
	      in [(("fresh_eqvt", thms1),[])] end
19638
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19635
diff changeset
   812
            ||>> PureThy.add_thmss
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19635
diff changeset
   813
	      let val thms1 = inst_pt_at [fresh_aux]
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19635
diff changeset
   814
		  and thms2 = inst_pt_pt_at_cp_dj [fresh_aux_ineq]
4358b88a9d12 added the lemmas pt_fresh_aux and pt_fresh_aux_ineq
urbanc
parents: 19635
diff changeset
   815
	      in [(("fresh_aux", thms1 @ thms2),[])] end
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   816
	   end;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   817
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   818
    in NominalData.put (fold Symtab.update (map (rpair ()) full_ak_names)
18262
d0fcd4d684f5 finished cleaning up the parts that collect
urbanc
parents: 18100
diff changeset
   819
      (NominalData.get thy11)) thy33
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   820
    end;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   821
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   822
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   823
(* syntax und parsing *)
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   824
structure P = OuterParse and K = OuterKeyword;
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   825
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   826
val atom_declP =
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   827
  OuterSyntax.command "atom_decl" "Declare new kinds of atoms" K.thy_decl
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   828
    (Scan.repeat1 P.name >> (Toplevel.theory o create_nom_typedecls));
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   829
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   830
val _ = OuterSyntax.add_parsers [atom_declP];
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   831
18746
a4ece70964ae made the change for setup-functions not returning functions
urbanc
parents: 18707
diff changeset
   832
val setup = NominalData.init;
18068
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   833
e8c3d371594e Moved atom stuff to new file nominal_atoms.ML
berghofe
parents:
diff changeset
   834
end;