src/HOL/Nominal/nominal_package.ML
author urbanc
Wed, 02 Nov 2005 15:31:12 +0100
changeset 18067 8b9848d150ba
parent 18066 d1e47ee13070
child 18068 e8c3d371594e
permissions -rw-r--r--
- completed the list of thms for supp_atm - cleaned up the way how thms are collected under single names
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     1
(* $Id$ *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     2
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     3
signature NOMINAL_PACKAGE =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     4
sig
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     5
  val create_nom_typedecls : string list -> theory -> theory
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     6
  val add_nominal_datatype : bool -> string list -> (string list * bstring * mixfix *
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     7
    (bstring * string list * mixfix) list) list -> theory -> theory *
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     8
      {distinct : thm list list,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
     9
       inject : thm list list,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    10
       exhaustion : thm list,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    11
       rec_thms : thm list,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    12
       case_thms : thm list list,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    13
       split_thms : (thm * thm) list,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    14
       induction : thm,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    15
       size : thm list,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    16
       simps : thm list}
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    17
  val setup : (theory -> theory) list
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    18
end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    19
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    20
structure NominalPackage (*: NOMINAL_PACKAGE *) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    21
struct
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    22
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    23
open DatatypeAux;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    24
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    25
(* data kind 'HOL/nominal' *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    26
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    27
structure NominalArgs =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    28
struct
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    29
  val name = "HOL/nominal";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    30
  type T = unit Symtab.table;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    31
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    32
  val empty = Symtab.empty;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    33
  val copy = I;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    34
  val extend = I;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    35
  fun merge _ x = Symtab.merge (K true) x;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    36
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    37
  fun print sg tab = ();
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    38
end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    39
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    40
structure NominalData = TheoryDataFun(NominalArgs);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    41
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    42
fun atoms_of thy = map fst (Symtab.dest (NominalData.get thy));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    43
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    44
(* FIXME: add to hologic.ML ? *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    45
fun mk_listT T = Type ("List.list", [T]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    46
fun mk_permT T = mk_listT (HOLogic.mk_prodT (T, T));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    47
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    48
fun mk_Cons x xs =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    49
  let val T = fastype_of x
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    50
  in Const ("List.list.Cons", T --> mk_listT T --> mk_listT T) $ x $ xs end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    51
18067
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
    52
(* FIXME: should be a library function *)
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
    53
fun cprod ([], ys) = []
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
    54
  | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    55
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    56
(* this function sets up all matters related to atom-  *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    57
(* kinds; the user specifies a list of atom-kind names *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    58
(* atom_decl <ak1> ... <akn>                           *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    59
fun create_nom_typedecls ak_names thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    60
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    61
    (* declares a type-decl for every atom-kind: *) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    62
    (* that is typedecl <ak>                     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    63
    val thy1 = TypedefPackage.add_typedecls (map (fn x => (x,[],NoSyn)) ak_names) thy;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    64
    
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    65
    (* produces a list consisting of pairs:         *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    66
    (*  fst component is the atom-kind name         *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    67
    (*  snd component is its type                   *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    68
    val full_ak_names = map (Sign.intern_type (sign_of thy1)) ak_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    69
    val ak_names_types = ak_names ~~ map (Type o rpair []) full_ak_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    70
     
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    71
    (* adds for every atom-kind an axiom             *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    72
    (* <ak>_infinite: infinite (UNIV::<ak_type> set) *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    73
    val (thy2,inf_axs) = PureThy.add_axioms_i (map (fn (ak_name, T) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    74
      let 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    75
	val name = ak_name ^ "_infinite"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    76
        val axiom = HOLogic.mk_Trueprop (HOLogic.mk_not
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    77
                    (HOLogic.mk_mem (HOLogic.mk_UNIV T,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    78
                     Const ("Finite_Set.Finites", HOLogic.mk_setT (HOLogic.mk_setT T)))))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    79
      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    80
	((name, axiom), []) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    81
      end) ak_names_types) thy1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    82
    
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    83
    (* declares a swapping function for every atom-kind, it is         *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    84
    (* const swap_<ak> :: <akT> * <akT> => <akT> => <akT>              *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    85
    (* swap_<ak> (a,b) c = (if a=c then b (else if b=c then a else c)) *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    86
    (* overloades then the general swap-function                       *) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    87
    val (thy3, swap_eqs) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    88
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    89
        val swapT = HOLogic.mk_prodT (T, T) --> T --> T;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    90
        val swap_name = Sign.full_name (sign_of thy) ("swap_" ^ ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    91
        val a = Free ("a", T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    92
        val b = Free ("b", T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    93
        val c = Free ("c", T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    94
        val ab = Free ("ab", HOLogic.mk_prodT (T, T))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    95
        val cif = Const ("HOL.If", HOLogic.boolT --> T --> T --> T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    96
        val cswap_akname = Const (swap_name, swapT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    97
        val cswap = Const ("nominal.swap", swapT)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    98
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
    99
        val name = "swap_"^ak_name^"_def";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   100
        val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   101
		   (cswap_akname $ HOLogic.mk_prod (a,b) $ c,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   102
                    cif $ HOLogic.mk_eq (a,c) $ b $ (cif $ HOLogic.mk_eq (b,c) $ a $ c)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   103
        val def2 = Logic.mk_equals (cswap $ ab $ c, cswap_akname $ ab $ c)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   104
      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   105
        thy |> Theory.add_consts_i [("swap_" ^ ak_name, swapT, NoSyn)] 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   106
            |> (#1 o PureThy.add_defs_i true [((name, def2),[])])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   107
            |> PrimrecPackage.add_primrec_i "" [(("", def1),[])]            
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   108
      end) (thy2, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   109
    
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   110
    (* declares a permutation function for every atom-kind acting  *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   111
    (* on such atoms                                               *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   112
    (* const <ak>_prm_<ak> :: (<akT> * <akT>)list => akT => akT    *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   113
    (* <ak>_prm_<ak> []     a = a                                  *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   114
    (* <ak>_prm_<ak> (x#xs) a = swap_<ak> x (perm xs a)            *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   115
    val (thy4, prm_eqs) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   116
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   117
        val swapT = HOLogic.mk_prodT (T, T) --> T --> T;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   118
        val swap_name = Sign.full_name (sign_of thy) ("swap_" ^ ak_name)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   119
        val prmT = mk_permT T --> T --> T;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   120
        val prm_name = ak_name ^ "_prm_" ^ ak_name;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   121
        val qu_prm_name = Sign.full_name (sign_of thy) prm_name;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   122
        val x  = Free ("x", HOLogic.mk_prodT (T, T));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   123
        val xs = Free ("xs", mk_permT T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   124
        val a  = Free ("a", T) ;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   125
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   126
        val cnil  = Const ("List.list.Nil", mk_permT T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   127
        
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   128
        val def1 = HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (qu_prm_name, prmT) $ cnil $ a, a));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   129
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   130
        val def2 = HOLogic.mk_Trueprop (HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   131
                   (Const (qu_prm_name, prmT) $ mk_Cons x xs $ a,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   132
                    Const (swap_name, swapT) $ x $ (Const (qu_prm_name, prmT) $ xs $ a)));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   133
      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   134
        thy |> Theory.add_consts_i [(prm_name, mk_permT T --> T --> T, NoSyn)] 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   135
            |> PrimrecPackage.add_primrec_i "" [(("", def1), []),(("", def2), [])]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   136
      end) (thy3, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   137
    
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   138
    (* defines permutation functions for all combinations of atom-kinds; *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   139
    (* there are a trivial cases and non-trivial cases                   *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   140
    (* non-trivial case:                                                 *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   141
    (* <ak>_prm_<ak>_def:  perm pi a == <ak>_prm_<ak> pi a               *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   142
    (* trivial case with <ak> != <ak'>                                   *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   143
    (* <ak>_prm<ak'>_def[simp]:  perm pi a == a                          *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   144
    (*                                                                   *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   145
    (* the trivial cases are added to the simplifier, while the non-     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   146
    (* have their own rules proved below                                 *)  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   147
    val (thy5, perm_defs) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   148
      foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   149
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   150
          val perm_def_name = ak_name ^ "_prm_" ^ ak_name';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   151
          val pi = Free ("pi", mk_permT T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   152
          val a  = Free ("a", T');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   153
          val cperm = Const ("nominal.perm", mk_permT T --> T' --> T');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   154
          val cperm_def = Const (Sign.full_name (sign_of thy') perm_def_name, mk_permT T --> T' --> T');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   155
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   156
          val name = ak_name ^ "_prm_" ^ ak_name' ^ "_def";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   157
          val def = Logic.mk_equals
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   158
                    (cperm $ pi $ a, if ak_name = ak_name' then cperm_def $ pi $ a else a)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   159
        in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   160
          thy' |> PureThy.add_defs_i true [((name, def),[])] 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   161
        end) (thy, ak_names_types)) (thy4, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   162
    
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   163
    (* proves that every atom-kind is an instance of at *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   164
    (* lemma at_<ak>_inst:                              *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   165
    (* at TYPE(<ak>)                                    *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   166
    val (thy6, prm_cons_thms) = 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   167
      thy5 |> PureThy.add_thms (map (fn (ak_name, T) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   168
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   169
        val ak_name_qu = Sign.full_name (sign_of thy5) (ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   170
        val i_type = Type(ak_name_qu,[]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   171
	val cat = Const ("nominal.at",(Term.itselfT i_type)  --> HOLogic.boolT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   172
        val at_type = Logic.mk_type i_type;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   173
        val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy5
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   174
                                  [Name "at_def",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   175
                                   Name (ak_name ^ "_prm_" ^ ak_name ^ "_def"),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   176
                                   Name (ak_name ^ "_prm_" ^ ak_name ^ ".simps"),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   177
                                   Name ("swap_" ^ ak_name ^ "_def"),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   178
                                   Name ("swap_" ^ ak_name ^ ".simps"),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   179
                                   Name (ak_name ^ "_infinite")]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   180
	    
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   181
	val name = "at_"^ak_name^ "_inst";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   182
        val statement = HOLogic.mk_Trueprop (cat $ at_type);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   183
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   184
        val proof = fn _ => auto_tac (claset(),simp_s);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   185
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   186
      in 
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   187
        ((name, standard (Goal.prove thy5 [] [] statement proof)), []) 
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   188
      end) ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   189
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   190
    (* declares a perm-axclass for every atom-kind               *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   191
    (* axclass pt_<ak>                                           *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   192
    (* pt_<ak>1[simp]: perm [] x = x                             *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   193
    (* pt_<ak>2:       perm (pi1@pi2) x = perm pi1 (perm pi2 x)  *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   194
    (* pt_<ak>3:       pi1 ~ pi2 ==> perm pi1 x = perm pi2 x     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   195
     val (thy7, pt_ax_classes) =  foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   196
      let 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   197
	  val cl_name = "pt_"^ak_name;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   198
          val ty = TFree("'a",["HOL.type"]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   199
          val x   = Free ("x", ty);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   200
          val pi1 = Free ("pi1", mk_permT T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   201
          val pi2 = Free ("pi2", mk_permT T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   202
          val cperm = Const ("nominal.perm", mk_permT T --> ty --> ty);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   203
          val cnil  = Const ("List.list.Nil", mk_permT T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   204
          val cappend = Const ("List.op @",mk_permT T --> mk_permT T --> mk_permT T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   205
          val cprm_eq = Const ("nominal.prm_eq",mk_permT T --> mk_permT T --> HOLogic.boolT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   206
          (* nil axiom *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   207
          val axiom1 = HOLogic.mk_Trueprop (HOLogic.mk_eq 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   208
                       (cperm $ cnil $ x, x));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   209
          (* append axiom *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   210
          val axiom2 = HOLogic.mk_Trueprop (HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   211
                       (cperm $ (cappend $ pi1 $ pi2) $ x, cperm $ pi1 $ (cperm $ pi2 $ x)));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   212
          (* perm-eq axiom *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   213
          val axiom3 = Logic.mk_implies
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   214
                       (HOLogic.mk_Trueprop (cprm_eq $ pi1 $ pi2),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   215
                        HOLogic.mk_Trueprop (HOLogic.mk_eq (cperm $ pi1 $ x, cperm $ pi2 $ x)));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   216
      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   217
        thy |> AxClass.add_axclass_i (cl_name, ["HOL.type"])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   218
                [((cl_name^"1", axiom1),[Simplifier.simp_add_global]), 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   219
                 ((cl_name^"2", axiom2),[]),                           
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   220
                 ((cl_name^"3", axiom3),[])]                          
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   221
      end) (thy6,ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   222
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   223
    (* proves that every pt_<ak>-type together with <ak>-type *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   224
    (* instance of pt                                         *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   225
    (* lemma pt_<ak>_inst:                                    *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   226
    (* pt TYPE('x::pt_<ak>) TYPE(<ak>)                        *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   227
    val (thy8, prm_inst_thms) = 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   228
      thy7 |> PureThy.add_thms (map (fn (ak_name, T) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   229
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   230
        val ak_name_qu = Sign.full_name (sign_of thy7) (ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   231
        val pt_name_qu = Sign.full_name (sign_of thy7) ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   232
        val i_type1 = TFree("'x",[pt_name_qu]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   233
        val i_type2 = Type(ak_name_qu,[]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   234
	val cpt = Const ("nominal.pt",(Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   235
        val pt_type = Logic.mk_type i_type1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   236
        val at_type = Logic.mk_type i_type2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   237
        val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy7
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   238
                                  [Name "pt_def",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   239
                                   Name ("pt_" ^ ak_name ^ "1"),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   240
                                   Name ("pt_" ^ ak_name ^ "2"),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   241
                                   Name ("pt_" ^ ak_name ^ "3")];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   242
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   243
	val name = "pt_"^ak_name^ "_inst";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   244
        val statement = HOLogic.mk_Trueprop (cpt $ pt_type $ at_type);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   245
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   246
        val proof = fn _ => auto_tac (claset(),simp_s);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   247
      in 
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   248
        ((name, standard (Goal.prove thy7 [] [] statement proof)), []) 
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   249
      end) ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   250
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   251
     (* declares an fs-axclass for every atom-kind       *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   252
     (* axclass fs_<ak>                                  *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   253
     (* fs_<ak>1: finite ((supp x)::<ak> set)            *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   254
     val (thy11, fs_ax_classes) =  foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   255
       let 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   256
	  val cl_name = "fs_"^ak_name;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   257
	  val pt_name = Sign.full_name (sign_of thy) ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   258
          val ty = TFree("'a",["HOL.type"]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   259
          val x   = Free ("x", ty);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   260
          val csupp    = Const ("nominal.supp", ty --> HOLogic.mk_setT T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   261
          val cfinites = Const ("Finite_Set.Finites", HOLogic.mk_setT (HOLogic.mk_setT T))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   262
          
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   263
          val axiom1   = HOLogic.mk_Trueprop (HOLogic.mk_mem (csupp $ x, cfinites));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   264
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   265
       in  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   266
        thy |> AxClass.add_axclass_i (cl_name, [pt_name]) [((cl_name^"1", axiom1),[])]            
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   267
       end) (thy8,ak_names_types); 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   268
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   269
     (* proves that every fs_<ak>-type together with <ak>-type   *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   270
     (* instance of fs-type                                      *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   271
     (* lemma abst_<ak>_inst:                                    *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   272
     (* fs TYPE('x::pt_<ak>) TYPE (<ak>)                         *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   273
     val (thy12, fs_inst_thms) = 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   274
       thy11 |> PureThy.add_thms (map (fn (ak_name, T) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   275
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   276
         val ak_name_qu = Sign.full_name (sign_of thy11) (ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   277
         val fs_name_qu = Sign.full_name (sign_of thy11) ("fs_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   278
         val i_type1 = TFree("'x",[fs_name_qu]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   279
         val i_type2 = Type(ak_name_qu,[]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   280
 	 val cfs = Const ("nominal.fs", 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   281
                                 (Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   282
         val fs_type = Logic.mk_type i_type1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   283
         val at_type = Logic.mk_type i_type2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   284
	 val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy11
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   285
                                   [Name "fs_def",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   286
                                    Name ("fs_" ^ ak_name ^ "1")];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   287
    
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   288
	 val name = "fs_"^ak_name^ "_inst";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   289
         val statement = HOLogic.mk_Trueprop (cfs $ fs_type $ at_type);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   290
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   291
         val proof = fn _ => auto_tac (claset(),simp_s);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   292
       in 
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   293
         ((name, standard (Goal.prove thy11 [] [] statement proof)), []) 
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   294
       end) ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   295
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   296
       (* declares for every atom-kind combination an axclass            *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   297
       (* cp_<ak1>_<ak2> giving a composition property                   *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   298
       (* cp_<ak1>_<ak2>1: pi1 o pi2 o x = (pi1 o pi2) o (pi1 o x)       *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   299
        val (thy12b,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   300
	 foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   301
	     let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   302
	       val cl_name = "cp_"^ak_name^"_"^ak_name';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   303
	       val ty = TFree("'a",["HOL.type"]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   304
               val x   = Free ("x", ty);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   305
               val pi1 = Free ("pi1", mk_permT T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   306
	       val pi2 = Free ("pi2", mk_permT T');                  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   307
	       val cperm1 = Const ("nominal.perm", mk_permT T  --> ty --> ty);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   308
               val cperm2 = Const ("nominal.perm", mk_permT T' --> ty --> ty);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   309
               val cperm3 = Const ("nominal.perm", mk_permT T  --> mk_permT T' --> mk_permT T');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   310
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   311
               val ax1   = HOLogic.mk_Trueprop 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   312
			   (HOLogic.mk_eq (cperm1 $ pi1 $ (cperm2 $ pi2 $ x), 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   313
                                           cperm2 $ (cperm3 $ pi1 $ pi2) $ (cperm1 $ pi1 $ x)));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   314
	       in  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   315
	       (fst (AxClass.add_axclass_i (cl_name, ["HOL.type"]) [((cl_name^"1", ax1),[])] thy'),())  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   316
	       end) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   317
	   (thy, ak_names_types)) (thy12, ak_names_types)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   318
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   319
        (* proves for every <ak>-combination a cp_<ak1>_<ak2>_inst theorem;     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   320
        (* lemma cp_<ak1>_<ak2>_inst:                                           *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   321
        (* cp TYPE('a::cp_<ak1>_<ak2>) TYPE(<ak1>) TYPE(<ak2>)                  *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   322
        val (thy12c, cp_thms) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   323
	 foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   324
           let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   325
             val ak_name_qu  = Sign.full_name (sign_of thy') (ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   326
	     val ak_name_qu' = Sign.full_name (sign_of thy') (ak_name');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   327
             val cp_name_qu  = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   328
             val i_type0 = TFree("'a",[cp_name_qu]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   329
             val i_type1 = Type(ak_name_qu,[]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   330
             val i_type2 = Type(ak_name_qu',[]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   331
	     val ccp = Const ("nominal.cp",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   332
                             (Term.itselfT i_type0)-->(Term.itselfT i_type1)-->
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   333
                                                      (Term.itselfT i_type2)-->HOLogic.boolT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   334
             val at_type  = Logic.mk_type i_type1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   335
             val at_type' = Logic.mk_type i_type2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   336
	     val cp_type  = Logic.mk_type i_type0;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   337
             val simp_s   = HOL_basic_ss addsimps PureThy.get_thmss thy' [(Name "cp_def")];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   338
	     val cp1      = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"1"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   339
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   340
	     val name = "cp_"^ak_name^ "_"^ak_name'^"_inst";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   341
             val statement = HOLogic.mk_Trueprop (ccp $ cp_type $ at_type $ at_type');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   342
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   343
             val proof = fn _ => EVERY [auto_tac (claset(),simp_s), rtac cp1 1];
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   344
	   in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   345
	     thy' |> PureThy.add_thms 
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   346
                    [((name, standard (Goal.prove thy' [] [] statement proof)), [])]
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   347
	   end) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   348
	   (thy, ak_names_types)) (thy12b, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   349
       
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   350
        (* proves for every non-trivial <ak>-combination a disjointness   *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   351
        (* theorem; i.e. <ak1> != <ak2>                                   *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   352
        (* lemma ds_<ak1>_<ak2>:                                          *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   353
        (* dj TYPE(<ak1>) TYPE(<ak2>)                                     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   354
        val (thy12d, dj_thms) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   355
	  foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   356
          (if not (ak_name = ak_name') 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   357
           then 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   358
	       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   359
		 val ak_name_qu  = Sign.full_name (sign_of thy') (ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   360
	         val ak_name_qu' = Sign.full_name (sign_of thy') (ak_name');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   361
                 val i_type1 = Type(ak_name_qu,[]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   362
                 val i_type2 = Type(ak_name_qu',[]);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   363
	         val cdj = Const ("nominal.disjoint",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   364
                           (Term.itselfT i_type1)-->(Term.itselfT i_type2)-->HOLogic.boolT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   365
                 val at_type  = Logic.mk_type i_type1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   366
                 val at_type' = Logic.mk_type i_type2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   367
                 val simp_s = HOL_basic_ss addsimps PureThy.get_thmss thy' 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   368
					   [Name "disjoint_def",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   369
                                            Name (ak_name^"_prm_"^ak_name'^"_def"),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   370
                                            Name (ak_name'^"_prm_"^ak_name^"_def")];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   371
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   372
	         val name = "dj_"^ak_name^"_"^ak_name';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   373
                 val statement = HOLogic.mk_Trueprop (cdj $ at_type $ at_type');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   374
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   375
                 val proof = fn _ => auto_tac (claset(),simp_s);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   376
	       in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   377
		   thy' |> PureThy.add_thms 
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
   378
                        [((name, standard (Goal.prove thy' [] [] statement proof)), []) ]
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   379
	       end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   380
           else 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   381
            (thy',[])))  (* do nothing branch, if ak_name = ak_name' *) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   382
	   (thy, ak_names_types)) (thy12c, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   383
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   384
     (*<<<<<<<  pt_<ak> class instances  >>>>>>>*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   385
     (*=========================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   386
     
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   387
     (* some frequently used theorems *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   388
      val pt1 = PureThy.get_thm thy12c (Name "pt1");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   389
      val pt2 = PureThy.get_thm thy12c (Name "pt2");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   390
      val pt3 = PureThy.get_thm thy12c (Name "pt3");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   391
      val at_pt_inst    = PureThy.get_thm thy12c (Name "at_pt_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   392
      val pt_bool_inst  = PureThy.get_thm thy12c (Name "pt_bool_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   393
      val pt_set_inst   = PureThy.get_thm thy12c (Name "pt_set_inst"); 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   394
      val pt_unit_inst  = PureThy.get_thm thy12c (Name "pt_unit_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   395
      val pt_prod_inst  = PureThy.get_thm thy12c (Name "pt_prod_inst"); 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   396
      val pt_list_inst  = PureThy.get_thm thy12c (Name "pt_list_inst");   
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   397
      val pt_optn_inst  = PureThy.get_thm thy12c (Name "pt_option_inst");   
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   398
      val pt_noptn_inst = PureThy.get_thm thy12c (Name "pt_noption_inst");   
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   399
      val pt_fun_inst   = PureThy.get_thm thy12c (Name "pt_fun_inst");     
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   400
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   401
     (* for all atom-kind combination shows that         *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   402
     (* every <ak> is an instance of pt_<ai>             *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   403
     val (thy13,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   404
	 foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   405
          (if ak_name = ak_name'
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   406
	   then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   407
	     let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   408
	      val qu_name =  Sign.full_name (sign_of thy') ak_name;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   409
              val qu_class = Sign.full_name (sign_of thy') ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   410
              val at_inst  = PureThy.get_thm thy' (Name ("at_"^ak_name ^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   411
              val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   412
                                 rtac ((at_inst RS at_pt_inst) RS pt1) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   413
                                 rtac ((at_inst RS at_pt_inst) RS pt2) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   414
                                 rtac ((at_inst RS at_pt_inst) RS pt3) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   415
                                 atac 1];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   416
             in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   417
	      (AxClass.add_inst_arity_i (qu_name,[],[qu_class]) proof thy',()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   418
             end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   419
           else 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   420
             let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   421
	      val qu_name' = Sign.full_name (sign_of thy') ak_name';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   422
              val qu_class = Sign.full_name (sign_of thy') ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   423
              val simp_s = HOL_basic_ss addsimps 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   424
                           PureThy.get_thmss thy' [Name (ak_name^"_prm_"^ak_name'^"_def")];  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   425
              val proof = EVERY [AxClass.intro_classes_tac [], auto_tac (claset(),simp_s)];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   426
             in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   427
	      (AxClass.add_inst_arity_i (qu_name',[],[qu_class]) proof thy',()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   428
             end)) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   429
	     (thy, ak_names_types)) (thy12c, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   430
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   431
     (* shows that bool is an instance of pt_<ak>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   432
     (* uses the theorem pt_bool_inst                 *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   433
     val (thy14,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   434
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   435
          val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   436
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   437
                             rtac (pt_bool_inst RS pt1) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   438
                             rtac (pt_bool_inst RS pt2) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   439
                             rtac (pt_bool_inst RS pt3) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   440
                             atac 1];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   441
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   442
	 (AxClass.add_inst_arity_i ("bool",[],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   443
       end) (thy13,ak_names_types); 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   444
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   445
     (* shows that set(pt_<ak>) is an instance of pt_<ak>          *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   446
     (* unfolds the permutation definition and applies pt_<ak>i    *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   447
     val (thy15,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   448
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   449
          val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   450
          val pt_inst  = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst"));  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   451
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   452
                             rtac ((pt_inst RS pt_set_inst) RS pt1) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   453
                             rtac ((pt_inst RS pt_set_inst) RS pt2) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   454
                             rtac ((pt_inst RS pt_set_inst) RS pt3) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   455
                             atac 1];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   456
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   457
	 (AxClass.add_inst_arity_i ("set",[[qu_class]],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   458
       end) (thy14,ak_names_types); 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   459
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   460
     (* shows that unit is an instance of pt_<ak>          *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   461
     val (thy16,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   462
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   463
          val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   464
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   465
                             rtac (pt_unit_inst RS pt1) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   466
                             rtac (pt_unit_inst RS pt2) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   467
                             rtac (pt_unit_inst RS pt3) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   468
                             atac 1];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   469
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   470
	 (AxClass.add_inst_arity_i ("Product_Type.unit",[],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   471
       end) (thy15,ak_names_types); 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   472
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   473
     (* shows that *(pt_<ak>,pt_<ak>) is an instance of pt_<ak> *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   474
     (* uses the theorem pt_prod_inst and pt_<ak>_inst          *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   475
     val (thy17,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   476
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   477
          val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   478
          val pt_inst  = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst"));  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   479
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   480
                             rtac ((pt_inst RS (pt_inst RS pt_prod_inst)) RS pt1) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   481
                             rtac ((pt_inst RS (pt_inst RS pt_prod_inst)) RS pt2) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   482
                             rtac ((pt_inst RS (pt_inst RS pt_prod_inst)) RS pt3) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   483
                             atac 1];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   484
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   485
          (AxClass.add_inst_arity_i ("*",[[qu_class],[qu_class]],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   486
       end) (thy16,ak_names_types); 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   487
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   488
     (* shows that list(pt_<ak>) is an instance of pt_<ak>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   489
     (* uses the theorem pt_list_inst and pt_<ak>_inst         *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   490
     val (thy18,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   491
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   492
          val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   493
          val pt_inst  = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   494
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   495
                             rtac ((pt_inst RS pt_list_inst) RS pt1) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   496
                             rtac ((pt_inst RS pt_list_inst) RS pt2) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   497
                             rtac ((pt_inst RS pt_list_inst) RS pt3) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   498
                             atac 1];      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   499
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   500
	 (AxClass.add_inst_arity_i ("List.list",[[qu_class]],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   501
       end) (thy17,ak_names_types); 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   502
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   503
     (* shows that option(pt_<ak>) is an instance of pt_<ak>   *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   504
     (* uses the theorem pt_option_inst and pt_<ak>_inst       *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   505
     val (thy18a,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   506
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   507
          val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   508
          val pt_inst  = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   509
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   510
                             rtac ((pt_inst RS pt_optn_inst) RS pt1) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   511
                             rtac ((pt_inst RS pt_optn_inst) RS pt2) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   512
                             rtac ((pt_inst RS pt_optn_inst) RS pt3) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   513
                             atac 1];      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   514
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   515
	 (AxClass.add_inst_arity_i ("Datatype.option",[[qu_class]],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   516
       end) (thy18,ak_names_types); 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   517
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   518
     (* shows that nOption(pt_<ak>) is an instance of pt_<ak>   *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   519
     (* uses the theorem pt_option_inst and pt_<ak>_inst       *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   520
     val (thy18b,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   521
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   522
          val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   523
          val pt_inst  = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   524
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   525
                             rtac ((pt_inst RS pt_noptn_inst) RS pt1) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   526
                             rtac ((pt_inst RS pt_noptn_inst) RS pt2) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   527
                             rtac ((pt_inst RS pt_noptn_inst) RS pt3) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   528
                             atac 1];      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   529
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   530
	 (AxClass.add_inst_arity_i ("nominal.nOption",[[qu_class]],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   531
       end) (thy18a,ak_names_types); 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   532
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   533
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   534
     (* shows that fun(pt_<ak>,pt_<ak>) is an instance of pt_<ak>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   535
     (* uses the theorem pt_list_inst and pt_<ak>_inst                *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   536
     val (thy19,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   537
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   538
          val qu_class = Sign.full_name (sign_of thy) ("pt_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   539
          val at_thm   = PureThy.get_thm thy (Name ("at_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   540
          val pt_inst  = PureThy.get_thm thy (Name ("pt_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   541
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   542
                             rtac ((at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst))) RS pt1) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   543
                             rtac ((at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst))) RS pt2) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   544
                             rtac ((at_thm RS (pt_inst RS (pt_inst RS pt_fun_inst))) RS pt3) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   545
                             atac 1];      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   546
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   547
	 (AxClass.add_inst_arity_i ("fun",[[qu_class],[qu_class]],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   548
       end) (thy18b,ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   549
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   550
       (*<<<<<<<  fs_<ak> class instances  >>>>>>>*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   551
       (*=========================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   552
       val fs1          = PureThy.get_thm thy19 (Name "fs1");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   553
       val fs_at_inst   = PureThy.get_thm thy19 (Name "fs_at_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   554
       val fs_unit_inst = PureThy.get_thm thy19 (Name "fs_unit_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   555
       val fs_bool_inst = PureThy.get_thm thy19 (Name "fs_bool_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   556
       val fs_prod_inst = PureThy.get_thm thy19 (Name "fs_prod_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   557
       val fs_list_inst = PureThy.get_thm thy19 (Name "fs_list_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   558
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   559
       (* shows that <ak> is an instance of fs_<ak>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   560
       (* uses the theorem at_<ak>_inst                 *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   561
       val (thy20,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   562
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   563
          val qu_name =  Sign.full_name (sign_of thy) ak_name;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   564
          val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   565
          val at_thm   = PureThy.get_thm thy (Name ("at_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   566
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   567
                             rtac ((at_thm RS fs_at_inst) RS fs1) 1];      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   568
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   569
	 (AxClass.add_inst_arity_i (qu_name,[],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   570
       end) (thy19,ak_names_types);  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   571
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   572
       (* shows that unit is an instance of fs_<ak>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   573
       (* uses the theorem fs_unit_inst                 *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   574
       val (thy21,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   575
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   576
          val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   577
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   578
                             rtac (fs_unit_inst RS fs1) 1];      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   579
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   580
	 (AxClass.add_inst_arity_i ("Product_Type.unit",[],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   581
       end) (thy20,ak_names_types);  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   582
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   583
       (* shows that bool is an instance of fs_<ak>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   584
       (* uses the theorem fs_bool_inst                 *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   585
       val (thy22,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   586
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   587
          val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   588
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   589
                             rtac (fs_bool_inst RS fs1) 1];      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   590
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   591
	 (AxClass.add_inst_arity_i ("bool",[],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   592
       end) (thy21,ak_names_types);  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   593
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   594
       (* shows that *(fs_<ak>,fs_<ak>) is an instance of fs_<ak>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   595
       (* uses the theorem fs_prod_inst                               *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   596
       val (thy23,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   597
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   598
          val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   599
          val fs_inst  = PureThy.get_thm thy (Name ("fs_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   600
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   601
                             rtac ((fs_inst RS (fs_inst RS fs_prod_inst)) RS fs1) 1];      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   602
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   603
	 (AxClass.add_inst_arity_i ("*",[[qu_class],[qu_class]],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   604
       end) (thy22,ak_names_types);  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   605
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   606
       (* shows that list(fs_<ak>) is an instance of fs_<ak>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   607
       (* uses the theorem fs_list_inst                          *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   608
       val (thy24,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   609
       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   610
          val qu_class = Sign.full_name (sign_of thy) ("fs_"^ak_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   611
          val fs_inst  = PureThy.get_thm thy (Name ("fs_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   612
          val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   613
                              rtac ((fs_inst RS fs_list_inst) RS fs1) 1];      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   614
       in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   615
	 (AxClass.add_inst_arity_i ("List.list",[[qu_class]],[qu_class]) proof thy,()) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   616
       end) (thy23,ak_names_types);  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   617
	   
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   618
       (*<<<<<<<  cp_<ak>_<ai> class instances  >>>>>>>*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   619
       (*==============================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   620
       val cp1             = PureThy.get_thm thy24 (Name "cp1");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   621
       val cp_unit_inst    = PureThy.get_thm thy24 (Name "cp_unit_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   622
       val cp_bool_inst    = PureThy.get_thm thy24 (Name "cp_bool_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   623
       val cp_prod_inst    = PureThy.get_thm thy24 (Name "cp_prod_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   624
       val cp_list_inst    = PureThy.get_thm thy24 (Name "cp_list_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   625
       val cp_fun_inst     = PureThy.get_thm thy24 (Name "cp_fun_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   626
       val cp_option_inst  = PureThy.get_thm thy24 (Name "cp_option_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   627
       val cp_noption_inst = PureThy.get_thm thy24 (Name "cp_noption_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   628
       val pt_perm_compose = PureThy.get_thm thy24 (Name "pt_perm_compose");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   629
       val dj_pp_forget    = PureThy.get_thm thy24 (Name "dj_perm_perm_forget");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   630
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   631
       (* shows that <aj> is an instance of cp_<ak>_<ai>  *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   632
       (* that needs a three-nested loop *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   633
       val (thy25,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   634
	 foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   635
          foldl_map (fn (thy'', (ak_name'', T'')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   636
            let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   637
              val qu_name =  Sign.full_name (sign_of thy'') ak_name;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   638
              val qu_class = Sign.full_name (sign_of thy'') ("cp_"^ak_name'^"_"^ak_name'');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   639
              val proof =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   640
                (if (ak_name'=ak_name'') then 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   641
		  (let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   642
                    val pt_inst  = PureThy.get_thm thy'' (Name ("pt_"^ak_name''^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   643
		    val at_inst  = PureThy.get_thm thy'' (Name ("at_"^ak_name''^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   644
                  in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   645
		   EVERY [AxClass.intro_classes_tac [], 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   646
                          rtac (at_inst RS (pt_inst RS pt_perm_compose)) 1]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   647
                  end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   648
		else
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   649
		  (let 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   650
                     val dj_inst  = PureThy.get_thm thy'' (Name ("dj_"^ak_name''^"_"^ak_name'));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   651
		     val simp_s = HOL_basic_ss addsimps 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   652
                                        ((dj_inst RS dj_pp_forget)::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   653
                                         (PureThy.get_thmss thy'' 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   654
					   [Name (ak_name' ^"_prm_"^ak_name^"_def"),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   655
                                            Name (ak_name''^"_prm_"^ak_name^"_def")]));  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   656
		  in 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   657
                    EVERY [AxClass.intro_classes_tac [], simp_tac simp_s 1]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   658
                  end))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   659
	      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   660
                (AxClass.add_inst_arity_i (qu_name,[],[qu_class]) proof thy'',())
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   661
	      end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   662
	   (thy', ak_names_types)) (thy, ak_names_types)) (thy24, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   663
      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   664
       (* shows that unit is an instance of cp_<ak>_<ai>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   665
       (* for every <ak>-combination                         *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   666
       val (thy26,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   667
	 foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   668
          let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   669
            val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   670
            val proof = EVERY [AxClass.intro_classes_tac [],rtac (cp_unit_inst RS cp1) 1];     
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   671
	  in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   672
            (AxClass.add_inst_arity_i ("Product_Type.unit",[],[qu_class]) proof thy',())
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   673
	  end) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   674
	   (thy, ak_names_types)) (thy25, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   675
       
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   676
       (* shows that bool is an instance of cp_<ak>_<ai>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   677
       (* for every <ak>-combination                         *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   678
       val (thy27,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   679
	 foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   680
           let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   681
	     val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   682
             val proof = EVERY [AxClass.intro_classes_tac [], rtac (cp_bool_inst RS cp1) 1];     
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   683
	   in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   684
             (AxClass.add_inst_arity_i ("bool",[],[qu_class]) proof thy',())
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   685
	   end) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   686
	   (thy, ak_names_types)) (thy26, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   687
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   688
       (* shows that prod is an instance of cp_<ak>_<ai>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   689
       (* for every <ak>-combination                         *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   690
       val (thy28,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   691
	 foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   692
          let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   693
	    val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   694
            val cp_inst  = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   695
            val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   696
                               rtac ((cp_inst RS (cp_inst RS cp_prod_inst)) RS cp1) 1];     
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   697
	  in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   698
            (AxClass.add_inst_arity_i ("*",[[qu_class],[qu_class]],[qu_class]) proof thy',())
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   699
	  end)  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   700
	  (thy, ak_names_types)) (thy27, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   701
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   702
       (* shows that list is an instance of cp_<ak>_<ai>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   703
       (* for every <ak>-combination                         *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   704
       val (thy29,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   705
	 foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   706
           let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   707
	     val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   708
             val cp_inst  = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   709
             val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   710
                                rtac ((cp_inst RS cp_list_inst) RS cp1) 1];     
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   711
	   in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   712
            (AxClass.add_inst_arity_i ("List.list",[[qu_class]],[qu_class]) proof thy',())
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   713
	   end) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   714
	   (thy, ak_names_types)) (thy28, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   715
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   716
       (* shows that function is an instance of cp_<ak>_<ai>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   717
       (* for every <ak>-combination                             *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   718
       val (thy30,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   719
	 foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   720
          let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   721
	    val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   722
            val cp_inst  = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   723
            val pt_inst  = PureThy.get_thm thy' (Name ("pt_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   724
            val at_inst  = PureThy.get_thm thy' (Name ("at_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   725
            val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   726
                    rtac ((at_inst RS (pt_inst RS (cp_inst RS (cp_inst RS cp_fun_inst)))) RS cp1) 1];  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   727
	  in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   728
            (AxClass.add_inst_arity_i ("fun",[[qu_class],[qu_class]],[qu_class]) proof thy',())
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   729
	  end) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   730
	  (thy, ak_names_types)) (thy29, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   731
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   732
       (* shows that option is an instance of cp_<ak>_<ai>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   733
       (* for every <ak>-combination                           *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   734
       val (thy31,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   735
	 foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   736
          let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   737
	    val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   738
            val cp_inst  = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   739
            val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   740
                               rtac ((cp_inst RS cp_option_inst) RS cp1) 1];     
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   741
	  in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   742
            (AxClass.add_inst_arity_i ("Datatype.option",[[qu_class]],[qu_class]) proof thy',())
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   743
	  end) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   744
	  (thy, ak_names_types)) (thy30, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   745
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   746
       (* shows that nOption is an instance of cp_<ak>_<ai>     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   747
       (* for every <ak>-combination                            *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   748
       val (thy32,_) = foldl_map (fn (thy, (ak_name, T)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   749
	 foldl_map (fn (thy', (ak_name', T')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   750
          let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   751
	    val qu_class = Sign.full_name (sign_of thy') ("cp_"^ak_name^"_"^ak_name');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   752
            val cp_inst  = PureThy.get_thm thy' (Name ("cp_"^ak_name^"_"^ak_name'^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   753
            val proof = EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   754
                               rtac ((cp_inst RS cp_noption_inst) RS cp1) 1];     
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   755
	  in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   756
           (AxClass.add_inst_arity_i ("nominal.nOption",[[qu_class]],[qu_class]) proof thy',())
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   757
	  end) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   758
	  (thy, ak_names_types)) (thy31, ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   759
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   760
       (* abbreviations for some collection of rules *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   761
       (*============================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   762
       val abs_fun_pi        = PureThy.get_thm thy32 (Name ("nominal.abs_fun_pi"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   763
       val abs_fun_pi_ineq   = PureThy.get_thm thy32 (Name ("nominal.abs_fun_pi_ineq"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   764
       val abs_fun_eq        = PureThy.get_thm thy32 (Name ("nominal.abs_fun_eq"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   765
       val dj_perm_forget    = PureThy.get_thm thy32 (Name ("nominal.dj_perm_forget"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   766
       val dj_pp_forget      = PureThy.get_thm thy32 (Name ("nominal.dj_perm_perm_forget"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   767
       val fresh_iff         = PureThy.get_thm thy32 (Name ("nominal.fresh_abs_fun_iff"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   768
       val fresh_iff_ineq    = PureThy.get_thm thy32 (Name ("nominal.fresh_abs_fun_iff_ineq"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   769
       val abs_fun_supp      = PureThy.get_thm thy32 (Name ("nominal.abs_fun_supp"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   770
       val abs_fun_supp_ineq = PureThy.get_thm thy32 (Name ("nominal.abs_fun_supp_ineq"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   771
       val pt_swap_bij       = PureThy.get_thm thy32 (Name ("nominal.pt_swap_bij"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   772
       val pt_fresh_fresh    = PureThy.get_thm thy32 (Name ("nominal.pt_fresh_fresh"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   773
       val pt_bij            = PureThy.get_thm thy32 (Name ("nominal.pt_bij"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   774
       val pt_perm_compose   = PureThy.get_thm thy32 (Name ("nominal.pt_perm_compose"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   775
       val perm_eq_app       = PureThy.get_thm thy32 (Name ("nominal.perm_eq_app"));
18054
2493cb9f5ede added the collection of lemmas "supp_at"
urbanc
parents: 18045
diff changeset
   776
       val at_fresh          = PureThy.get_thm thy32 (Name ("nominal.at_fresh"));
18067
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   777
       val at_calc           = PureThy.get_thms thy32 (Name ("nominal.at_calc"));
18054
2493cb9f5ede added the collection of lemmas "supp_at"
urbanc
parents: 18045
diff changeset
   778
       val at_supp           = PureThy.get_thm thy32 (Name ("nominal.at_supp"));
18067
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   779
       val dj_supp           = PureThy.get_thm thy32 (Name ("nominal.dj_supp"));
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   780
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   781
       (* abs_perm collects all lemmas for simplifying a permutation *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   782
       (* in front of an abs_fun                                     *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   783
       val (thy33,_) = 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   784
	   let 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   785
	     val name = "abs_perm"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   786
             val thm_list = Library.flat (map (fn (ak_name, T) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   787
	        let	
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   788
		  val at_inst = PureThy.get_thm thy32 (Name ("at_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   789
		  val pt_inst = PureThy.get_thm thy32 (Name ("pt_"^ak_name^"_inst"));	      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   790
	          val thm = [pt_inst, at_inst] MRS abs_fun_pi
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   791
                  val thm_list = map (fn (ak_name', T') =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   792
                     let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   793
                      val cp_inst = PureThy.get_thm thy32 (Name ("cp_"^ak_name^"_"^ak_name'^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   794
	             in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   795
                     [pt_inst, pt_inst, at_inst, cp_inst] MRS abs_fun_pi_ineq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   796
	             end) ak_names_types;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   797
                 in thm::thm_list end) (ak_names_types))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   798
           in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   799
             (PureThy.add_thmss [((name, thm_list),[])] thy32)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   800
           end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   801
18054
2493cb9f5ede added the collection of lemmas "supp_at"
urbanc
parents: 18045
diff changeset
   802
       val (thy34,_) = 
18067
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   803
	 let 
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   804
             (* takes a theorem and a list of theorems        *)
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   805
             (* produces a list of theorems of the form       *)
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   806
             (* [t1 RS thm,..,tn RS thm] where t1..tn in thms *) 
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   807
             fun instantiate thm thms = map (fn ti => ti RS thm) thms;
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   808
               
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   809
             (* takes two theorem lists (hopefully of the same length)           *)
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   810
             (* produces a list of theorems of the form                          *)
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   811
             (* [t1 RS m1,..,tn RS mn] where t1..tn in thms1 and m1..mn in thms2 *) 
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   812
             fun instantiate_zip thms1 thms2 = 
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   813
		 map (fn (t1,t2) => t1 RS t2) (thms1 ~~ thms2);
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   814
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   815
             (* list of all at_inst-theorems *)
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   816
             val ats = map (fn ak => PureThy.get_thm thy33 (Name ("at_"^ak^"_inst"))) ak_names
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   817
             (* list of all pt_inst-theorems *)
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   818
             val pts = map (fn ak => PureThy.get_thm thy33 (Name ("pt_"^ak^"_inst"))) ak_names
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   819
             (* list of all cp_inst-theorems *)
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   820
             val cps = 
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   821
	       let fun cps_fun (ak1,ak2) = PureThy.get_thm thy33 (Name ("cp_"^ak1^"_"^ak2^"_inst"))
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   822
	       in map cps_fun (cprod (ak_names,ak_names)) end;	
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   823
             (* list of all dj_inst-theorems *)
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   824
             val djs = 
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   825
	       let fun djs_fun (ak1,ak2) = 
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   826
		    if ak1=ak2 
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   827
		    then NONE
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   828
		    else SOME(PureThy.get_thm thy33 (Name ("dj_"^ak1^"_"^ak2)))
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   829
	       in List.mapPartial I (map djs_fun (cprod (ak_names,ak_names))) end;	
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   830
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   831
             fun inst_pt thms = Library.flat (map (fn ti => instantiate ti pts) thms); 
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   832
             fun inst_at thms = Library.flat (map (fn ti => instantiate ti ats) thms);               
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   833
	     fun inst_pt_at thms = instantiate_zip ats (inst_pt thms);			
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   834
             fun inst_dj thms = Library.flat (map (fn ti => instantiate ti djs) thms);  
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   835
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   836
           in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   837
            thy33 
18067
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   838
	    |>   PureThy.add_thmss [(("alpha", inst_pt_at [abs_fun_eq]),[])]
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   839
            |>>> PureThy.add_thmss [(("perm_swap", inst_pt_at [pt_swap_bij]),[])]
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   840
            |>>> PureThy.add_thmss [(("perm_fresh_fresh", inst_pt_at [pt_fresh_fresh]),[])]
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   841
            |>>> PureThy.add_thmss [(("perm_bij", inst_pt_at [pt_bij]),[])]
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   842
            |>>> PureThy.add_thmss [(("perm_compose", inst_pt_at [pt_perm_compose]),[])]
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   843
            |>>> PureThy.add_thmss [(("perm_app_eq", inst_pt_at [perm_eq_app]),[])]
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   844
            |>>> PureThy.add_thmss [(("supp_atm", (inst_at [at_supp]) @ (inst_dj [dj_supp])),[])]
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   845
            |>>> PureThy.add_thmss [(("fresh_atm", inst_at [at_fresh]),[])]
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   846
            |>>> PureThy.add_thmss [(("calc_atm", inst_at at_calc),[])]
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
   847
            
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   848
	   end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   849
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   850
         (* perm_dj collects all lemmas that forget an permutation *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   851
         (* when it acts on an atom of different type              *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   852
         val (thy35,_) = 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   853
	   let 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   854
	     val name = "perm_dj"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   855
             val thm_list = Library.flat (map (fn (ak_name, T) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   856
	        Library.flat (map (fn (ak_name', T') => 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   857
                 if not (ak_name = ak_name') 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   858
                 then 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   859
		    let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   860
                      val dj_inst = PureThy.get_thm thy34 (Name ("dj_"^ak_name^"_"^ak_name'));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   861
                    in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   862
                      [dj_inst RS dj_perm_forget, dj_inst RS dj_pp_forget]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   863
                    end 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   864
                 else []) ak_names_types)) ak_names_types)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   865
           in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   866
             (PureThy.add_thmss [((name, thm_list),[])] thy34)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   867
           end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   868
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   869
         (* abs_fresh collects all lemmas for simplifying a freshness *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   870
         (* proposition involving an abs_fun                          *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   871
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   872
         val (thy36,_) = 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   873
	   let 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   874
	     val name = "abs_fresh"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   875
             val thm_list = Library.flat (map (fn (ak_name, T) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   876
	        let	
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   877
		  val at_inst = PureThy.get_thm thy35 (Name ("at_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   878
		  val pt_inst = PureThy.get_thm thy35 (Name ("pt_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   879
                  val fs_inst = PureThy.get_thm thy35 (Name ("fs_"^ak_name^"_inst"));	      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   880
	          val thm = [pt_inst, at_inst, (fs_inst RS fs1)] MRS fresh_iff
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   881
                  val thm_list = Library.flat (map (fn (ak_name', T') =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   882
                     (if (not (ak_name = ak_name')) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   883
                     then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   884
                       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   885
                        val cp_inst = PureThy.get_thm thy35 (Name ("cp_"^ak_name^"_"^ak_name'^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   886
	                val dj_inst = PureThy.get_thm thy35 (Name ("dj_"^ak_name'^"_"^ak_name));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   887
                       in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   888
                        [[pt_inst, pt_inst, at_inst, cp_inst, dj_inst] MRS fresh_iff_ineq]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   889
	               end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   890
                     else [])) ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   891
                 in thm::thm_list end) (ak_names_types))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   892
           in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   893
             (PureThy.add_thmss [((name, thm_list),[])] thy35)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   894
           end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   895
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   896
         (* abs_supp collects all lemmas for simplifying  *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   897
         (* support proposition involving an abs_fun      *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   898
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   899
         val (thy37,_) = 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   900
	   let 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   901
	     val name = "abs_supp"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   902
             val thm_list = Library.flat (map (fn (ak_name, T) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   903
	        let	
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   904
		  val at_inst = PureThy.get_thm thy36 (Name ("at_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   905
		  val pt_inst = PureThy.get_thm thy36 (Name ("pt_"^ak_name^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   906
                  val fs_inst = PureThy.get_thm thy36 (Name ("fs_"^ak_name^"_inst"));	      
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   907
	          val thm1 = [pt_inst, at_inst, (fs_inst RS fs1)] MRS abs_fun_supp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   908
                  val thm2 = [pt_inst, at_inst] MRS abs_fun_supp
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   909
                  val thm_list = Library.flat (map (fn (ak_name', T') =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   910
                     (if (not (ak_name = ak_name')) 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   911
                     then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   912
                       let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   913
                        val cp_inst = PureThy.get_thm thy36 (Name ("cp_"^ak_name^"_"^ak_name'^"_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   914
	                val dj_inst = PureThy.get_thm thy36 (Name ("dj_"^ak_name'^"_"^ak_name));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   915
                       in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   916
                        [[pt_inst, pt_inst, at_inst, cp_inst, dj_inst] MRS abs_fun_supp_ineq]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   917
	               end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   918
                     else [])) ak_names_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   919
                 in thm1::thm2::thm_list end) (ak_names_types))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   920
           in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   921
             (PureThy.add_thmss [((name, thm_list),[])] thy36)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   922
           end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   923
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   924
    in NominalData.put (fold Symtab.update (map (rpair ()) full_ak_names)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   925
      (NominalData.get thy11)) thy37
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   926
    end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   927
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   928
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   929
(* syntax und parsing *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   930
structure P = OuterParse and K = OuterKeyword;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   931
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   932
val atom_declP =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   933
  OuterSyntax.command "atom_decl" "Declare new kinds of atoms" K.thy_decl
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   934
    (Scan.repeat1 P.name >> (Toplevel.theory o create_nom_typedecls));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   935
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   936
val _ = OuterSyntax.add_parsers [atom_declP];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   937
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   938
val setup = [NominalData.init];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   939
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   940
(*=======================================================================*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   941
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   942
(** FIXME: DatatypePackage should export this function **)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   943
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   944
local
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   945
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   946
fun dt_recs (DtTFree _) = []
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   947
  | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   948
  | dt_recs (DtRec i) = [i];
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   949
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   950
fun dt_cases (descr: descr) (_, args, constrs) =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   951
  let
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   952
    fun the_bname i = Sign.base_name (#1 (valOf (AList.lookup (op =) descr i)));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   953
    val bnames = map the_bname (distinct (List.concat (map dt_recs args)));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   954
  in map (fn (c, _) => space_implode "_" (Sign.base_name c :: bnames)) constrs end;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   955
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   956
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   957
fun induct_cases descr =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   958
  DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr)));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   959
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   960
fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   961
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   962
in
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   963
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   964
fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   965
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   966
fun mk_case_names_exhausts descr new =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   967
  map (RuleCases.case_names o exhaust_cases descr o #1)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   968
    (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr);
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   969
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   970
end;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   971
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   972
(*******************************)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
   973
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   974
val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   975
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   976
fun read_typ sign ((Ts, sorts), str) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   977
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   978
    val T = Type.no_tvars (Sign.read_typ (sign, (AList.lookup op =)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   979
      (map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   980
  in (Ts @ [T], add_typ_tfrees (T, sorts)) end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   981
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   982
(** taken from HOL/Tools/datatype_aux.ML **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   983
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   984
fun indtac indrule indnames i st =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   985
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   986
    val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   987
    val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   988
      (Logic.strip_imp_concl (List.nth (prems_of st, i - 1))));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   989
    val getP = if can HOLogic.dest_imp (hd ts) then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   990
      (apfst SOME) o HOLogic.dest_imp else pair NONE;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   991
    fun abstr (t1, t2) = (case t1 of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   992
        NONE => (case filter (fn Free (s, _) => s mem indnames | _ => false)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   993
              (term_frees t2) of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   994
            [Free (s, T)] => absfree (s, T, t2)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   995
          | _ => sys_error "indtac")
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   996
      | SOME (_ $ t' $ _) => Abs ("x", fastype_of t', abstract_over (t', t2)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   997
    val cert = cterm_of (Thm.sign_of_thm st);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   998
    val Ps = map (cert o head_of o snd o getP) ts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
   999
    val indrule' = cterm_instantiate (Ps ~~
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1000
      (map (cert o abstr o getP) ts')) indrule
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1001
  in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1002
    rtac indrule' i st
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1003
  end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1004
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1005
fun gen_add_nominal_datatype prep_typ err flat_names new_type_names dts thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1006
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1007
    (* this theory is used just for parsing *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1008
  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1009
    val tmp_thy = thy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1010
      Theory.copy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1011
      Theory.add_types (map (fn (tvs, tname, mx, _) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1012
        (tname, length tvs, mx)) dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1013
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1014
    val sign = Theory.sign_of tmp_thy;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1015
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1016
    val atoms = atoms_of thy;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1017
    val classes = map (NameSpace.map_base (fn s => "pt_" ^ s)) atoms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1018
    val cp_classes = List.concat (map (fn atom1 => map (fn atom2 =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1019
      Sign.intern_class thy ("cp_" ^ Sign.base_name atom1 ^ "_" ^
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1020
        Sign.base_name atom2)) atoms) atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1021
    fun augment_sort S = S union classes;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1022
    val augment_sort_typ = map_type_tfree (fn (s, S) => TFree (s, augment_sort S));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1023
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1024
    fun prep_constr ((constrs, sorts), (cname, cargs, mx)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1025
      let val (cargs', sorts') = Library.foldl (prep_typ sign) (([], sorts), cargs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1026
      in (constrs @ [(cname, cargs', mx)], sorts') end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1027
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1028
    fun prep_dt_spec ((dts, sorts), (tvs, tname, mx, constrs)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1029
      let val (constrs', sorts') = Library.foldl prep_constr (([], sorts), constrs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1030
      in (dts @ [(tvs, tname, mx, constrs')], sorts') end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1031
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1032
    val (dts', sorts) = Library.foldl prep_dt_spec (([], []), dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1033
    val sorts' = map (apsnd augment_sort) sorts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1034
    val tyvars = map #1 dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1035
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1036
    val types_syntax = map (fn (tvs, tname, mx, constrs) => (tname, mx)) dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1037
    val constr_syntax = map (fn (tvs, tname, mx, constrs) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1038
      map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1039
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1040
    val ps = map (fn (_, n, _, _) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1041
      (Sign.full_name sign n, Sign.full_name sign (n ^ "_Rep"))) dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1042
    val rps = map Library.swap ps;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1043
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1044
    fun replace_types (Type ("nominal.ABS", [T, U])) = 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1045
          Type ("fun", [T, Type ("nominal.nOption", [replace_types U])])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1046
      | replace_types (Type (s, Ts)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1047
          Type (getOpt (AList.lookup op = ps s, s), map replace_types Ts)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1048
      | replace_types T = T;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1049
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1050
    fun replace_types' (Type (s, Ts)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1051
          Type (getOpt (AList.lookup op = rps s, s), map replace_types' Ts)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1052
      | replace_types' T = T;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1053
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1054
    val dts'' = map (fn (tvs, tname, mx, constrs) => (tvs, tname ^ "_Rep", NoSyn,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1055
      map (fn (cname, cargs, mx) => (cname,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1056
        map (augment_sort_typ o replace_types) cargs, NoSyn)) constrs)) dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1057
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1058
    val new_type_names' = map (fn n => n ^ "_Rep") new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1059
    val full_new_type_names' = map (Sign.full_name (sign_of thy)) new_type_names';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1060
18045
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
  1061
    val ({induction, ...},thy1) =
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1062
      DatatypePackage.add_datatype_i err flat_names new_type_names' dts'' thy;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1063
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1064
    val SOME {descr, ...} = Symtab.lookup
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1065
      (DatatypePackage.get_datatypes thy1) (hd full_new_type_names');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1066
    val typ_of_dtyp = typ_of_dtyp descr sorts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1067
    fun nth_dtyp i = typ_of_dtyp (DtRec i);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1068
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1069
    (**** define permutation functions ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1070
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1071
    val permT = mk_permT (TFree ("'x", HOLogic.typeS));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1072
    val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1073
    val perm_types = map (fn (i, _) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1074
      let val T = nth_dtyp i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1075
      in permT --> T --> T end) descr;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1076
    val perm_names = replicate (length new_type_names) "nominal.perm" @
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1077
      DatatypeProp.indexify_names (map (fn i => Sign.full_name (sign_of thy1)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1078
        ("perm_" ^ name_of_typ (nth_dtyp i)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1079
          (length new_type_names upto length descr - 1));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1080
    val perm_names_types = perm_names ~~ perm_types;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1081
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1082
    val perm_eqs = List.concat (map (fn (i, (_, _, constrs)) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1083
      let val T = nth_dtyp i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1084
      in map (fn (cname, dts) => 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1085
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1086
          val Ts = map typ_of_dtyp dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1087
          val names = DatatypeProp.make_tnames Ts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1088
          val args = map Free (names ~~ Ts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1089
          val c = Const (cname, Ts ---> T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1090
          fun perm_arg (dt, x) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1091
            let val T = type_of x
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1092
            in if is_rec_type dt then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1093
                let val (Us, _) = strip_type T
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1094
                in list_abs (map (pair "x") Us,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1095
                  Const (List.nth (perm_names_types, body_index dt)) $ pi $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1096
                    list_comb (x, map (fn (i, U) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1097
                      Const ("nominal.perm", permT --> U --> U) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1098
                        (Const ("List.rev", permT --> permT) $ pi) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1099
                        Bound i) ((length Us - 1 downto 0) ~~ Us)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1100
                end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1101
              else Const ("nominal.perm", permT --> T --> T) $ pi $ x
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1102
            end;  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1103
        in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1104
          (("", HOLogic.mk_Trueprop (HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1105
            (Const (List.nth (perm_names_types, i)) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1106
               Free ("pi", mk_permT (TFree ("'x", HOLogic.typeS))) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1107
               list_comb (c, args),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1108
             list_comb (c, map perm_arg (dts ~~ args))))), [])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1109
        end) constrs
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1110
      end) descr);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1111
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1112
    val (thy2, perm_simps) = thy1 |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1113
      Theory.add_consts_i (map (fn (s, T) => (Sign.base_name s, T, NoSyn))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1114
        (List.drop (perm_names_types, length new_type_names))) |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1115
      PrimrecPackage.add_primrec_i "" perm_eqs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1116
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1117
    (**** prove that permutation functions introduced by unfolding are ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1118
    (**** equivalent to already existing permutation functions         ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1119
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1120
    val _ = warning ("length descr: " ^ string_of_int (length descr));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1121
    val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1122
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1123
    val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1124
    val perm_fun_def = PureThy.get_thm thy2 (Name "perm_fun_def");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1125
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1126
    val unfolded_perm_eq_thms =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1127
      if length descr = length new_type_names then []
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1128
      else map standard (List.drop (split_conj_thm
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1129
        (Goal.prove thy2 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1130
          (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1131
            (map (fn (c as (s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1132
               let val [T1, T2] = binder_types T
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1133
               in HOLogic.mk_eq (Const c $ pi $ Free (x, T2),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1134
                 Const ("nominal.perm", T) $ pi $ Free (x, T2))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1135
               end)
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1136
             (perm_names_types ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1137
          (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1138
            ALLGOALS (asm_full_simp_tac
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1139
              (simpset_of thy2 addsimps [perm_fun_def]))])),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1140
        length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1141
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1142
    (**** prove [] \<bullet> t = t ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1143
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1144
    val _ = warning "perm_empty_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1145
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1146
    val perm_empty_thms = List.concat (map (fn a =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1147
      let val permT = mk_permT (Type (a, []))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1148
      in map standard (List.take (split_conj_thm
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1149
        (Goal.prove thy2 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1150
          (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1151
            (map (fn ((s, T), x) => HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1152
                (Const (s, permT --> T --> T) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1153
                   Const ("List.list.Nil", permT) $ Free (x, T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1154
                 Free (x, T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1155
             (perm_names ~~
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1156
              map body_type perm_types ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1157
          (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1158
            ALLGOALS (asm_full_simp_tac (simpset_of thy2))])),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1159
        length new_type_names))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1160
      end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1161
      atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1162
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1163
    (**** prove (pi1 @ pi2) \<bullet> t = pi1 \<bullet> (pi2 \<bullet> t) ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1164
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1165
    val _ = warning "perm_append_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1166
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1167
    (*FIXME: these should be looked up statically*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1168
    val at_pt_inst = PureThy.get_thm thy2 (Name "at_pt_inst");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1169
    val pt2 = PureThy.get_thm thy2 (Name "pt2");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1170
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1171
    val perm_append_thms = List.concat (map (fn a =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1172
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1173
        val permT = mk_permT (Type (a, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1174
        val pi1 = Free ("pi1", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1175
        val pi2 = Free ("pi2", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1176
        val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1177
        val pt2' = pt_inst RS pt2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1178
        val pt2_ax = PureThy.get_thm thy2
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1179
          (Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "2") a));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1180
      in List.take (map standard (split_conj_thm
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1181
        (Goal.prove thy2 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1182
             (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1183
                (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1184
                    let val perm = Const (s, permT --> T --> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1185
                    in HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1186
                      (perm $ (Const ("List.op @", permT --> permT --> permT) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1187
                         pi1 $ pi2) $ Free (x, T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1188
                       perm $ pi1 $ (perm $ pi2 $ Free (x, T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1189
                    end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1190
                  (perm_names ~~
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1191
                   map body_type perm_types ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1192
           (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1193
              ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1194
         length new_type_names)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1195
      end) atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1196
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1197
    (**** prove pi1 ~ pi2 ==> pi1 \<bullet> t = pi2 \<bullet> t ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1198
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1199
    val _ = warning "perm_eq_thms";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1200
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1201
    val pt3 = PureThy.get_thm thy2 (Name "pt3");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1202
    val pt3_rev = PureThy.get_thm thy2 (Name "pt3_rev");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1203
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1204
    val perm_eq_thms = List.concat (map (fn a =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1205
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1206
        val permT = mk_permT (Type (a, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1207
        val pi1 = Free ("pi1", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1208
        val pi2 = Free ("pi2", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1209
        (*FIXME: not robust - better access these theorems using NominalData?*)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1210
        val at_inst = PureThy.get_thm thy2 (Name ("at_" ^ Sign.base_name a ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1211
        val pt_inst = PureThy.get_thm thy2 (Name ("pt_" ^ Sign.base_name a ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1212
        val pt3' = pt_inst RS pt3;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1213
        val pt3_rev' = at_inst RS (pt_inst RS pt3_rev);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1214
        val pt3_ax = PureThy.get_thm thy2
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1215
          (Name (NameSpace.map_base (fn s => "pt_" ^ s ^ "3") a));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1216
      in List.take (map standard (split_conj_thm
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1217
        (Goal.prove thy2 [] [] (Logic.mk_implies
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1218
             (HOLogic.mk_Trueprop (Const ("nominal.prm_eq",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1219
                permT --> permT --> HOLogic.boolT) $ pi1 $ pi2),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1220
              HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1221
                (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1222
                    let val perm = Const (s, permT --> T --> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1223
                    in HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1224
                      (perm $ pi1 $ Free (x, T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1225
                       perm $ pi2 $ Free (x, T))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1226
                    end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1227
                  (perm_names ~~
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1228
                   map body_type perm_types ~~ perm_indnames)))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1229
           (fn _ => EVERY [indtac induction perm_indnames 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1230
              ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1231
         length new_type_names)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1232
      end) atoms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1233
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1234
    (**** prove pi1 \<bullet> (pi2 \<bullet> t) = (pi1 \<bullet> pi2) \<bullet> (pi1 \<bullet> t) ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1235
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1236
    val cp1 = PureThy.get_thm thy2 (Name "cp1");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1237
    val dj_cp = PureThy.get_thm thy2 (Name "dj_cp");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1238
    val pt_perm_compose = PureThy.get_thm thy2 (Name "pt_perm_compose");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1239
    val pt_perm_compose_rev = PureThy.get_thm thy2 (Name "pt_perm_compose_rev");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1240
    val dj_perm_perm_forget = PureThy.get_thm thy2 (Name "dj_perm_perm_forget");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1241
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1242
    fun composition_instance name1 name2 thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1243
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1244
        val name1' = Sign.base_name name1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1245
        val name2' = Sign.base_name name2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1246
        val cp_class = Sign.intern_class thy ("cp_" ^ name1' ^ "_" ^ name2');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1247
        val permT1 = mk_permT (Type (name1, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1248
        val permT2 = mk_permT (Type (name2, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1249
        val augment = map_type_tfree
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1250
          (fn (x, S) => TFree (x, cp_class :: S));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1251
        val Ts = map (augment o body_type) perm_types;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1252
        val cp_inst = PureThy.get_thm thy
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1253
          (Name ("cp_" ^ name1' ^ "_" ^ name2' ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1254
        val simps = simpset_of thy addsimps (perm_fun_def ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1255
          (if name1 <> name2 then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1256
             let val dj = PureThy.get_thm thy (Name ("dj_" ^ name2' ^ "_" ^ name1'))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1257
             in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1258
           else
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1259
             let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1260
               val at_inst = PureThy.get_thm thy (Name ("at_" ^ name1' ^ "_inst"));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1261
               val pt_inst = PureThy.get_thm thy (Name ("pt_" ^ name1' ^ "_inst"))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1262
             in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1263
               [cp_inst RS cp1 RS sym,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1264
                at_inst RS (pt_inst RS pt_perm_compose) RS sym,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1265
                at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1266
            end))
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1267
        val thms = split_conj_thm (standard (Goal.prove thy [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1268
            (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1269
              (map (fn ((s, T), x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1270
                  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1271
                    val pi1 = Free ("pi1", permT1);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1272
                    val pi2 = Free ("pi2", permT2);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1273
                    val perm1 = Const (s, permT1 --> T --> T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1274
                    val perm2 = Const (s, permT2 --> T --> T);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1275
                    val perm3 = Const ("nominal.perm", permT1 --> permT2 --> permT2)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1276
                  in HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1277
                    (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1278
                     perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1279
                  end)
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1280
                (perm_names ~~ Ts ~~ perm_indnames))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1281
          (fn _ => EVERY [indtac induction perm_indnames 1,
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1282
             ALLGOALS (asm_full_simp_tac simps)])))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1283
      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1284
        foldl (fn ((s, tvs), thy) => AxClass.add_inst_arity_i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1285
            (s, replicate (length tvs) (cp_class :: classes), [cp_class])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1286
            (AxClass.intro_classes_tac [] THEN ALLGOALS (resolve_tac thms)) thy)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1287
          thy (full_new_type_names' ~~ tyvars)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1288
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1289
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1290
    val (thy3, perm_thmss) = thy2 |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1291
      fold (fn name1 => fold (composition_instance name1) atoms) atoms |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1292
      curry (Library.foldr (fn ((i, (tyname, args, _)), thy) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1293
        AxClass.add_inst_arity_i (tyname, replicate (length args) classes, classes)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1294
        (AxClass.intro_classes_tac [] THEN REPEAT (EVERY
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1295
           [resolve_tac perm_empty_thms 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1296
            resolve_tac perm_append_thms 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1297
            resolve_tac perm_eq_thms 1, assume_tac 1])) thy))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1298
        (List.take (descr, length new_type_names)) |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1299
      PureThy.add_thmss
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1300
        [((space_implode "_" new_type_names ^ "_unfolded_perm_eq",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1301
          unfolded_perm_eq_thms), [Simplifier.simp_add_global]),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1302
         ((space_implode "_" new_type_names ^ "_perm_empty",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1303
          perm_empty_thms), [Simplifier.simp_add_global]),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1304
         ((space_implode "_" new_type_names ^ "_perm_append",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1305
          perm_append_thms), [Simplifier.simp_add_global]),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1306
         ((space_implode "_" new_type_names ^ "_perm_eq",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1307
          perm_eq_thms), [Simplifier.simp_add_global])];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1308
  
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1309
    (**** Define representing sets ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1310
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1311
    val _ = warning "representing sets";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1312
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1313
    val rep_set_names = map (Sign.full_name thy3) (DatatypeProp.indexify_names
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1314
      (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1315
    val big_rep_name =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1316
      space_implode "_" (DatatypeProp.indexify_names (List.mapPartial
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1317
        (fn (i, ("nominal.nOption", _, _)) => NONE
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1318
          | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1319
    val _ = warning ("big_rep_name: " ^ big_rep_name);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1320
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1321
    fun strip_option (dtf as DtType ("fun", [dt, DtRec i])) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1322
          (case AList.lookup op = descr i of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1323
             SOME ("nominal.nOption", _, [(_, [dt']), _]) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1324
               apfst (cons dt) (strip_option dt')
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1325
           | _ => ([], dtf))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1326
      | strip_option dt = ([], dt);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1327
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1328
    fun make_intr s T (cname, cargs) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1329
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1330
        fun mk_prem (dt, (j, j', prems, ts)) = 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1331
          let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1332
            val (dts, dt') = strip_option dt;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1333
            val (dts', dt'') = strip_dtyp dt';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1334
            val Ts = map typ_of_dtyp dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1335
            val Us = map typ_of_dtyp dts';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1336
            val T = typ_of_dtyp dt'';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1337
            val free = mk_Free "x" (Us ---> T) j;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1338
            val free' = app_bnds free (length Us);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1339
            fun mk_abs_fun (T, (i, t)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1340
              let val U = fastype_of t
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1341
              in (i + 1, Const ("nominal.abs_fun", [T, U, T] --->
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1342
                Type ("nominal.nOption", [U])) $ mk_Free "y" T i $ t)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1343
              end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1344
          in (j + 1, j' + length Ts,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1345
            case dt'' of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1346
                DtRec k => list_all (map (pair "x") Us,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1347
                  HOLogic.mk_Trueprop (HOLogic.mk_mem (free',
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1348
                    Const (List.nth (rep_set_names, k),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1349
                      HOLogic.mk_setT T)))) :: prems
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1350
              | _ => prems,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1351
            snd (foldr mk_abs_fun (j', free) Ts) :: ts)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1352
          end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1353
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1354
        val (_, _, prems, ts) = foldr mk_prem (1, 1, [], []) cargs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1355
        val concl = HOLogic.mk_Trueprop (HOLogic.mk_mem
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1356
          (list_comb (Const (cname, map fastype_of ts ---> T), ts),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1357
           Const (s, HOLogic.mk_setT T)))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1358
      in Logic.list_implies (prems, concl)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1359
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1360
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1361
    val (intr_ts, ind_consts) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1362
      apfst List.concat (ListPair.unzip (List.mapPartial
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1363
        (fn ((_, ("nominal.nOption", _, _)), _) => NONE
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1364
          | ((i, (_, _, constrs)), rep_set_name) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1365
              let val T = nth_dtyp i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1366
              in SOME (map (make_intr rep_set_name T) constrs,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1367
                Const (rep_set_name, HOLogic.mk_setT T))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1368
              end)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1369
                (descr ~~ rep_set_names)));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1370
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1371
    val (thy4, {raw_induct = rep_induct, intrs = rep_intrs, ...}) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1372
      setmp InductivePackage.quiet_mode false
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1373
        (InductivePackage.add_inductive_i false true big_rep_name false true false
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1374
           ind_consts (map (fn x => (("", x), [])) intr_ts) []) thy3;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1375
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1376
    (**** Prove that representing set is closed under permutation ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1377
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1378
    val _ = warning "proving closure under permutation...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1379
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1380
    val perm_indnames' = List.mapPartial
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1381
      (fn (x, (_, ("nominal.nOption", _, _))) => NONE | (x, _) => SOME x)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1382
      (perm_indnames ~~ descr);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1383
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1384
    fun mk_perm_closed name = map (fn th => standard (th RS mp))
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1385
      (List.take (split_conj_thm (Goal.prove thy4 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1386
        (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1387
           (fn (S, x) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1388
              let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1389
                val S = map_term_types (map_type_tfree
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1390
                  (fn (s, cs) => TFree (s, cs union cp_classes))) S;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1391
                val T = HOLogic.dest_setT (fastype_of S);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1392
                val permT = mk_permT (Type (name, []))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1393
              in HOLogic.mk_imp (HOLogic.mk_mem (Free (x, T), S),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1394
                HOLogic.mk_mem (Const ("nominal.perm", permT --> T --> T) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1395
                  Free ("pi", permT) $ Free (x, T), S))
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1396
              end) (ind_consts ~~ perm_indnames'))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1397
        (fn _ => EVERY (* CU: added perm_fun_def in the final tactic in order to deal with funs *)
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1398
           [indtac rep_induct [] 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1399
            ALLGOALS (simp_tac (simpset_of thy4 addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1400
              (symmetric perm_fun_def :: PureThy.get_thms thy4 (Name ("abs_perm"))))),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1401
            ALLGOALS (resolve_tac rep_intrs 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1402
               THEN_ALL_NEW (asm_full_simp_tac (simpset_of thy4 addsimps [perm_fun_def])))])),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1403
        length new_type_names));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1404
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1405
    (* FIXME: theorems are stored in database for testing only *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1406
    val perm_closed_thmss = map mk_perm_closed atoms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1407
    val (thy5, _) = PureThy.add_thmss [(("perm_closed",
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1408
      List.concat perm_closed_thmss), [])] thy4;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1409
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1410
    (**** typedef ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1411
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1412
    val _ = warning "defining type...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1413
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1414
    val (thy6, typedefs) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1415
      foldl_map (fn (thy, ((((name, mx), tvs), c), name')) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1416
        setmp TypedefPackage.quiet_mode true
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1417
          (TypedefPackage.add_typedef_i false (SOME name') (name, tvs, mx) c NONE
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1418
            (rtac exI 1 THEN
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1419
              QUIET_BREADTH_FIRST (has_fewer_prems 1)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1420
              (resolve_tac rep_intrs 1))) thy |> (fn (thy, r) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1421
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1422
          val permT = mk_permT (TFree (variant tvs "'a", HOLogic.typeS));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1423
          val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1424
          val tvs' = map (fn s => TFree (s, the (AList.lookup op = sorts' s))) tvs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1425
          val T = Type (Sign.intern_type thy name, tvs');
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1426
          val Const (_, Type (_, [U])) = c
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1427
        in apsnd (pair r o hd)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1428
          (PureThy.add_defs_i true [(("prm_" ^ name ^ "_def", Logic.mk_equals
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1429
            (Const ("nominal.perm", permT --> T --> T) $ pi $ Free ("x", T),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1430
             Const (Sign.intern_const thy ("Abs_" ^ name), U --> T) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1431
               (Const ("nominal.perm", permT --> U --> U) $ pi $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1432
                 (Const (Sign.intern_const thy ("Rep_" ^ name), T --> U) $
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1433
                   Free ("x", T))))), [])] thy)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1434
        end))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1435
          (thy5, types_syntax ~~ tyvars ~~
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1436
            (List.take (ind_consts, length new_type_names)) ~~ new_type_names);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1437
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1438
    val perm_defs = map snd typedefs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1439
    val Abs_inverse_thms = map (#Abs_inverse o fst) typedefs;
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1440
    val Rep_inverse_thms = map (#Rep_inverse o fst) typedefs;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1441
    val Rep_thms = map (#Rep o fst) typedefs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1442
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1443
    val big_name = space_implode "_" new_type_names;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1444
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1445
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1446
    (** prove that new types are in class pt_<name> **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1447
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1448
    val _ = warning "prove that new types are in class pt_<name> ...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1449
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1450
    fun pt_instance ((class, atom), perm_closed_thms) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1451
      fold (fn (((({Abs_inverse, Rep_inverse, Rep, ...},
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1452
        perm_def), name), tvs), perm_closed) => fn thy =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1453
          AxClass.add_inst_arity_i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1454
            (Sign.intern_type thy name,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1455
              replicate (length tvs) (classes @ cp_classes), [class])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1456
            (EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1457
              rewrite_goals_tac [perm_def],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1458
              asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1459
              asm_full_simp_tac (simpset_of thy addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1460
                [Rep RS perm_closed RS Abs_inverse]) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1461
              asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1462
                (Name ("pt_" ^ Sign.base_name atom ^ "3"))]) 1]) thy)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1463
        (typedefs ~~ new_type_names ~~ tyvars ~~ perm_closed_thms);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1464
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1465
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1466
    (** prove that new types are in class cp_<name1>_<name2> **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1467
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1468
    val _ = warning "prove that new types are in class cp_<name1>_<name2> ...";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1469
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1470
    fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1471
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1472
        val name = "cp_" ^ Sign.base_name atom1 ^ "_" ^ Sign.base_name atom2;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1473
        val class = Sign.intern_class thy name;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1474
        val cp1' = PureThy.get_thm thy (Name (name ^ "_inst")) RS cp1
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1475
      in fold (fn ((((({Abs_inverse, Rep_inverse, Rep, ...},
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1476
        perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1477
          AxClass.add_inst_arity_i
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1478
            (Sign.intern_type thy name,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1479
              replicate (length tvs) (classes @ cp_classes), [class])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1480
            (EVERY [AxClass.intro_classes_tac [],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1481
              rewrite_goals_tac [perm_def],
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1482
              asm_full_simp_tac (simpset_of thy addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1483
                ((Rep RS perm_closed1 RS Abs_inverse) ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1484
                 (if atom1 = atom2 then []
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1485
                  else [Rep RS perm_closed2 RS Abs_inverse]))) 1,
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1486
              cong_tac 1,
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1487
              rtac refl 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1488
              rtac cp1' 1]) thy)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1489
        (typedefs ~~ new_type_names ~~ tyvars ~~ perm_closed_thms1 ~~
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1490
          perm_closed_thms2) thy
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1491
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1492
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1493
    val thy7 = fold (fn x => fn thy => thy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1494
      pt_instance x |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1495
      fold (cp_instance (apfst snd x)) (atoms ~~ perm_closed_thmss))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1496
        (classes ~~ atoms ~~ perm_closed_thmss) thy6;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1497
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1498
    (**** constructors ****)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1499
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1500
    fun mk_abs_fun (x, t) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1501
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1502
        val T = fastype_of x;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1503
        val U = fastype_of t
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1504
      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1505
        Const ("nominal.abs_fun", T --> U --> T -->
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1506
          Type ("nominal.nOption", [U])) $ x $ t
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1507
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1508
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1509
    val (ty_idxs, _) = foldl
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1510
      (fn ((i, ("nominal.nOption", _, _)), p) => p
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1511
        | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1512
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1513
    fun reindex (DtType (s, dts)) = DtType (s, map reindex dts)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1514
      | reindex (DtRec i) = DtRec (the (AList.lookup op = ty_idxs i))
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1515
      | reindex dt = dt;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1516
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1517
    fun strip_suffix i s = implode (List.take (explode s, size s - i));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1518
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1519
    (** strips the "_Rep" in type names *)
18045
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
  1520
    fun strip_nth_name i s = 
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
  1521
      let val xs = NameSpace.unpack s; 
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
  1522
      in NameSpace.pack (Library.nth_map (length xs - i) (strip_suffix 4) xs) end;
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1523
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1524
    val descr'' = List.mapPartial
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1525
      (fn (i, ("nominal.nOption", _, _)) => NONE
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1526
        | (i, (s, dts, constrs)) => SOME (the (AList.lookup op = ty_idxs i),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1527
            (strip_nth_name 1 s,  map reindex dts,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1528
             map (fn (cname, cargs) =>
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1529
               (strip_nth_name 2 cname,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1530
                foldl (fn (dt, dts) =>
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1531
                  let val (dts', dt') = strip_option dt
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1532
                  in (dts @ dts' @ [reindex dt']) end) [] cargs)) constrs))) descr;
18045
6d69a4190eb2 1) have adjusted the swapping of the result type
urbanc
parents: 18017
diff changeset
  1533
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1534
    val (descr1, descr2) = splitAt (length new_type_names, descr'');
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1535
    val descr' = [descr1, descr2];
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1536
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1537
    val typ_of_dtyp' = replace_types' o typ_of_dtyp;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1538
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1539
    val rep_names = map (fn s =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1540
      Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1541
    val abs_names = map (fn s =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1542
      Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1543
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1544
    val recTs' = List.mapPartial
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1545
      (fn ((_, ("nominal.nOption", _, _)), T) => NONE
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1546
        | (_, T) => SOME T) (descr ~~ get_rec_types descr sorts');
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1547
    val recTs = get_rec_types (List.concat descr') sorts';
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1548
    val newTs' = Library.take (length new_type_names, recTs');
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1549
    val newTs = Library.take (length new_type_names, recTs);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1550
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1551
    val full_new_type_names = map (Sign.full_name (sign_of thy)) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1552
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1553
    fun make_constr_def tname T T' ((thy, defs, eqns), ((cname, cargs), (cname', mx))) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1554
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1555
        fun constr_arg (dt, (j, l_args, r_args)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1556
          let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1557
            val x' = mk_Free "x" (typ_of_dtyp' dt) j;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1558
            val (dts, dt') = strip_option dt;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1559
            val xs = map (fn (dt, i) => mk_Free "x" (typ_of_dtyp' dt) i)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1560
              (dts ~~ (j upto j + length dts - 1))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1561
            val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1562
            val (dts', dt'') = strip_dtyp dt'
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1563
          in case dt'' of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1564
              DtRec k => if k < length new_type_names then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1565
                  (j + length dts + 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1566
                   xs @ x :: l_args,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1567
                   foldr mk_abs_fun
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1568
                     (list_abs (map (pair "z" o typ_of_dtyp') dts',
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1569
                       Const (List.nth (rep_names, k), typ_of_dtyp' dt'' -->
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1570
                         typ_of_dtyp dt'') $ app_bnds x (length dts')))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1571
                     xs :: r_args)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1572
                else error "nested recursion not (yet) supported"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1573
            | _ => (j + 1, x' :: l_args, x' :: r_args)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1574
          end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1575
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1576
        val (_, l_args, r_args) = foldr constr_arg (1, [], []) cargs;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1577
        val abs_name = Sign.intern_const (Theory.sign_of thy) ("Abs_" ^ tname);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1578
        val rep_name = Sign.intern_const (Theory.sign_of thy) ("Rep_" ^ tname);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1579
        val constrT = map fastype_of l_args ---> T;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1580
        val lhs = list_comb (Const (Sign.full_name thy (Sign.base_name cname),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1581
          constrT), l_args);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1582
        val rhs = list_comb (Const (cname, map fastype_of r_args ---> T'), r_args);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1583
        val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1584
        val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1585
          (Const (rep_name, T --> T') $ lhs, rhs));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1586
        val def_name = (Sign.base_name cname) ^ "_def";
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1587
        val (thy', [def_thm]) = thy |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1588
          Theory.add_consts_i [(cname', constrT, mx)] |>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1589
          (PureThy.add_defs_i false o map Thm.no_attributes) [(def_name, def)]
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1590
      in (thy', defs @ [def_thm], eqns @ [eqn]) end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1591
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1592
    fun dt_constr_defs ((thy, defs, eqns, dist_lemmas),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1593
        (((((_, (_, _, constrs)), tname), T), T'), constr_syntax)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1594
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1595
        val rep_const = cterm_of thy
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1596
          (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T'));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1597
        val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1598
        val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T')
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1599
          ((Theory.add_path tname thy, defs, []), constrs ~~ constr_syntax)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1600
      in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1601
        (parent_path flat_names thy', defs', eqns @ [eqns'], dist_lemmas @ [dist])
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1602
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1603
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1604
    val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1605
      ((thy7, [], [], []), List.take (descr, length new_type_names) ~~
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1606
        new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1607
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1608
    val abs_inject_thms = map (fn tname =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1609
      PureThy.get_thm thy8 (Name ("Abs_" ^ tname ^ "_inject"))) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1610
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1611
    val rep_inject_thms = map (fn tname =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1612
      PureThy.get_thm thy8 (Name ("Rep_" ^ tname ^ "_inject"))) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1613
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1614
    val rep_thms = map (fn tname =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1615
      PureThy.get_thm thy8 (Name ("Rep_" ^ tname))) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1616
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1617
    val rep_inverse_thms = map (fn tname =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1618
      PureThy.get_thm thy8 (Name ("Rep_" ^ tname ^ "_inverse"))) new_type_names;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1619
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1620
    (* prove theorem  Rep_i (Constr_j ...) = Constr'_j ...  *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1621
    
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1622
    fun prove_constr_rep_thm eqn =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1623
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1624
        val inj_thms = map (fn r => r RS iffD1) abs_inject_thms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1625
        val rewrites = constr_defs @ map mk_meta_eq rep_inverse_thms
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1626
      in standard (Goal.prove thy8 [] [] eqn (fn _ => EVERY
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1627
        [resolve_tac inj_thms 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1628
         rewrite_goals_tac rewrites,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1629
         rtac refl 3,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1630
         resolve_tac rep_intrs 2,
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1631
         REPEAT (resolve_tac rep_thms 1)]))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1632
      end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1633
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1634
    val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1635
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1636
    (* prove theorem  pi \<bullet> Rep_i x = Rep_i (pi \<bullet> x) *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1637
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1638
    fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1639
      let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1640
        val _ $ (_ $ (Rep $ x) $ _) = Logic.unvarify (prop_of th);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1641
        val Type ("fun", [T, U]) = fastype_of Rep;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1642
        val permT = mk_permT (Type (atom, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1643
        val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1644
      in
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1645
        standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1646
            (Const ("nominal.perm", permT --> U --> U) $ pi $ (Rep $ x),
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1647
             Rep $ (Const ("nominal.perm", permT --> T --> T) $ pi $ x))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1648
          (fn _ => simp_tac (HOL_basic_ss addsimps (perm_defs @ Abs_inverse_thms @
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1649
            perm_closed_thms @ Rep_thms)) 1))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1650
      end) Rep_thms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1651
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1652
    val perm_rep_perm_thms = List.concat (map prove_perm_rep_perm
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1653
      (atoms ~~ perm_closed_thmss));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1654
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1655
    (* prove distinctness theorems *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1656
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1657
    val distinct_props = setmp DatatypeProp.dtK 1000
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1658
      (DatatypeProp.make_distincts new_type_names descr' sorts') thy8;
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1659
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1660
    val dist_rewrites = map (fn (rep_thms, dist_lemma) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1661
      dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1662
        (constr_rep_thmss ~~ dist_lemmas);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1663
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1664
    fun prove_distinct_thms (_, []) = []
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1665
      | prove_distinct_thms (p as (rep_thms, dist_lemma), t::ts) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1666
          let
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1667
            val dist_thm = standard (Goal.prove thy8 [] [] t (fn _ =>
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1668
              simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1669
          in dist_thm::(standard (dist_thm RS not_sym))::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1670
            (prove_distinct_thms (p, ts))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1671
          end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1672
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1673
    val distinct_thms = map prove_distinct_thms
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1674
      (constr_rep_thmss ~~ dist_lemmas ~~ distinct_props);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1675
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1676
    (** prove equations for permutation functions **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1677
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1678
    val abs_perm = PureThy.get_thms thy8 (Name "abs_perm"); (* FIXME *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1679
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1680
    val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1681
      let val T = replace_types' (nth_dtyp i)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1682
      in List.concat (map (fn (atom, perm_closed_thms) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1683
          map (fn ((cname, dts), constr_rep_thm) => 
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1684
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1685
          val cname = Sign.intern_const thy8
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1686
            (NameSpace.append tname (Sign.base_name cname));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1687
          val permT = mk_permT (Type (atom, []));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1688
          val pi = Free ("pi", permT);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1689
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1690
          fun perm t =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1691
            let val T = fastype_of t
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1692
            in Const ("nominal.perm", permT --> T --> T) $ pi $ t end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1693
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1694
          fun constr_arg (dt, (j, l_args, r_args)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1695
            let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1696
              val x' = mk_Free "x" (typ_of_dtyp' dt) j;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1697
              val (dts, dt') = strip_option dt;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1698
              val Ts = map typ_of_dtyp' dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1699
              val xs = map (fn (T, i) => mk_Free "x" T i)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1700
                (Ts ~~ (j upto j + length dts - 1))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1701
              val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1702
              val (dts', dt'') = strip_dtyp dt';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1703
            in case dt'' of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1704
                DtRec k => if k < length new_type_names then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1705
                    (j + length dts + 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1706
                     xs @ x :: l_args,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1707
                     map perm (xs @ [x]) @ r_args)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1708
                  else error "nested recursion not (yet) supported"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1709
              | _ => (j + 1, x' :: l_args, perm x' :: r_args)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1710
            end
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1711
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1712
          val (_, l_args, r_args) = foldr constr_arg (1, [], []) dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1713
          val c = Const (cname, map fastype_of l_args ---> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1714
        in
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1715
          standard (Goal.prove thy8 [] []
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1716
            (HOLogic.mk_Trueprop (HOLogic.mk_eq
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1717
              (perm (list_comb (c, l_args)), list_comb (c, r_args))))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1718
            (fn _ => EVERY
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1719
              [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1720
               simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1721
                 constr_defs @ perm_closed_thms)) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1722
               TRY (simp_tac (HOL_basic_ss addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1723
                 (symmetric perm_fun_def :: abs_perm)) 1),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1724
               TRY (simp_tac (HOL_basic_ss addsimps
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1725
                 (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1726
                    perm_closed_thms)) 1)]))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1727
        end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss))
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1728
      end) (List.take (descr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1729
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1730
    (** prove injectivity of constructors **)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1731
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1732
    val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1733
    val alpha = PureThy.get_thms thy8 (Name "alpha");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1734
    val abs_fresh = PureThy.get_thms thy8 (Name "abs_fresh");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1735
    val fresh_def = PureThy.get_thm thy8 (Name "fresh_def");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1736
    val supp_def = PureThy.get_thm thy8 (Name "supp_def");
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1737
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1738
    val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1739
      let val T = replace_types' (nth_dtyp i)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1740
      in List.mapPartial (fn ((cname, dts), constr_rep_thm) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1741
        if null dts then NONE else SOME
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1742
        let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1743
          val cname = Sign.intern_const thy8
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1744
            (NameSpace.append tname (Sign.base_name cname));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1745
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1746
          fun make_inj (dt, (j, args1, args2, eqs)) =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1747
            let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1748
              val x' = mk_Free "x" (typ_of_dtyp' dt) j;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1749
              val y' = mk_Free "y" (typ_of_dtyp' dt) j;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1750
              val (dts, dt') = strip_option dt;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1751
              val Ts_idx = map typ_of_dtyp' dts ~~ (j upto j + length dts - 1);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1752
              val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1753
              val ys = map (fn (T, i) => mk_Free "y" T i) Ts_idx;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1754
              val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1755
              val y = mk_Free "y" (typ_of_dtyp' dt') (j + length dts);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1756
              val (dts', dt'') = strip_dtyp dt';
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1757
            in case dt'' of
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1758
                DtRec k => if k < length new_type_names then
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1759
                    (j + length dts + 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1760
                     xs @ (x :: args1), ys @ (y :: args2),
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1761
                     HOLogic.mk_eq
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1762
                       (foldr mk_abs_fun x xs, foldr mk_abs_fun y ys) :: eqs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1763
                  else error "nested recursion not (yet) supported"
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1764
              | _ => (j + 1, x' :: args1, y' :: args2, HOLogic.mk_eq (x', y') :: eqs)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1765
            end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1766
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1767
          val (_, args1, args2, eqs) = foldr make_inj (1, [], [], []) dts;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1768
          val Ts = map fastype_of args1;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1769
          val c = Const (cname, Ts ---> T)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1770
        in
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1771
          standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1772
              (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)),
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1773
               foldr1 HOLogic.mk_conj eqs)))
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1774
            (fn _ => EVERY
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1775
               [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1776
                  rep_inject_thms')) 1,
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1777
                TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def ::
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1778
                  alpha @ abs_perm @ abs_fresh @ rep_inject_thms @
17874
8be65cf94d2e Improved proof of injectivity theorems to make it work on
berghofe
parents: 17873
diff changeset
  1779
                  perm_rep_perm_thms)) 1),
8be65cf94d2e Improved proof of injectivity theorems to make it work on
berghofe
parents: 17873
diff changeset
  1780
                TRY (asm_full_simp_tac (HOL_basic_ss addsimps (perm_fun_def ::
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1781
                  expand_fun_eq :: rep_inject_thms @ perm_rep_perm_thms)) 1)]))
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1782
        end) (constrs ~~ constr_rep_thms)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1783
      end) (List.take (descr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss);
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1784
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1785
    (** equations for support and freshness **)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1786
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1787
    val Un_assoc = PureThy.get_thm thy8 (Name "Un_assoc");
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1788
    val de_Morgan_conj = PureThy.get_thm thy8 (Name "de_Morgan_conj");
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1789
    val Collect_disj_eq = PureThy.get_thm thy8 (Name "Collect_disj_eq");
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1790
    val finite_Un = PureThy.get_thm thy8 (Name "finite_Un");
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1791
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1792
    val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1793
      (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') =>
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1794
      let val T = replace_types' (nth_dtyp i)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1795
      in List.concat (map (fn (cname, dts) => map (fn atom =>
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1796
        let
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1797
          val cname = Sign.intern_const thy8
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1798
            (NameSpace.append tname (Sign.base_name cname));
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1799
          val atomT = Type (atom, []);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1800
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1801
          fun process_constr (dt, (j, args1, args2)) =
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1802
            let
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1803
              val x' = mk_Free "x" (typ_of_dtyp' dt) j;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1804
              val (dts, dt') = strip_option dt;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1805
              val Ts_idx = map typ_of_dtyp' dts ~~ (j upto j + length dts - 1);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1806
              val xs = map (fn (T, i) => mk_Free "x" T i) Ts_idx;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1807
              val x = mk_Free "x" (typ_of_dtyp' dt') (j + length dts);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1808
              val (dts', dt'') = strip_dtyp dt';
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1809
            in case dt'' of
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1810
                DtRec k => if k < length new_type_names then
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1811
                    (j + length dts + 1,
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1812
                     xs @ (x :: args1), foldr mk_abs_fun x xs :: args2)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1813
                  else error "nested recursion not (yet) supported"
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1814
              | _ => (j + 1, x' :: args1, x' :: args2)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1815
            end;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1816
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1817
          val (_, args1, args2) = foldr process_constr (1, [], []) dts;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1818
          val Ts = map fastype_of args1;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1819
          val c = list_comb (Const (cname, Ts ---> T), args1);
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1820
          fun supp t =
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1821
            Const ("nominal.supp", fastype_of t --> HOLogic.mk_setT atomT) $ t;
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1822
          fun fresh t =
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1823
            Const ("nominal.fresh", atomT --> fastype_of t --> HOLogic.boolT) $
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1824
              Free ("a", atomT) $ t;
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1825
          val supp_thm = standard (Goal.prove thy8 [] []
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1826
              (HOLogic.mk_Trueprop (HOLogic.mk_eq
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1827
                (supp c,
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1828
                 if null dts then Const ("{}", HOLogic.mk_setT atomT)
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1829
                 else foldr1 (HOLogic.mk_binop "op Un") (map supp args2))))
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1830
            (fn _ =>
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1831
              simp_tac (HOL_basic_ss addsimps (supp_def ::
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1832
                 Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un ::
17874
8be65cf94d2e Improved proof of injectivity theorems to make it work on
berghofe
parents: 17873
diff changeset
  1833
                 symmetric empty_def :: Finites.emptyI :: simp_thms @
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1834
                 abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1))
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1835
        in
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1836
          (supp_thm,
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1837
           standard (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1838
              (fresh c,
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1839
               if null dts then HOLogic.true_const
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1840
               else foldr1 HOLogic.mk_conj (map fresh args2))))
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1841
             (fn _ =>
18010
c885c93a9324 Removed legacy prove_goalw_cterm command.
berghofe
parents: 17874
diff changeset
  1842
               simp_tac (simpset_of thy8 addsimps [fresh_def, supp_thm]) 1)))
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1843
        end) atoms) constrs)
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1844
      end) (List.take (descr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps')));
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1845
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1846
    (**** Induction theorem ****)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1847
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1848
    val arities = get_arities (List.concat descr');
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1849
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1850
    fun mk_funs_inv thm =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1851
      let
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1852
        val {sign, prop, ...} = rep_thm thm;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1853
        val _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ S)) $
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1854
          (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1855
        val used = add_term_tfree_names (a, []);
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1856
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1857
        fun mk_thm i =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1858
          let
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1859
            val Ts = map (TFree o rpair HOLogic.typeS)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1860
              (variantlist (replicate i "'t", used));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1861
            val f = Free ("f", Ts ---> U)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1862
          in standard (Goal.prove sign [] [] (Logic.mk_implies
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1863
            (HOLogic.mk_Trueprop (HOLogic.list_all
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1864
               (map (pair "x") Ts, HOLogic.mk_mem (app_bnds f i, S))),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1865
             HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1866
               r $ (a $ app_bnds f i)), f))))
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1867
            (fn _ => EVERY [REPEAT (rtac ext 1), REPEAT (etac allE 1), rtac thm 1, atac 1]))
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1868
          end
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1869
      in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1870
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1871
    fun mk_indrule_lemma ((prems, concls), (((i, _), T), U)) =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1872
      let
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1873
        val Rep_t = Const (List.nth (rep_names, i), T --> U) $
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1874
          mk_Free "x" T i;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1875
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1876
        val Abs_t =  Const (List.nth (abs_names, i), U --> T)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1877
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1878
      in (prems @ [HOLogic.imp $ HOLogic.mk_mem (Rep_t,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1879
            Const (List.nth (rep_set_names, i), HOLogic.mk_setT U)) $
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1880
              (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1881
          concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1882
      end;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1883
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1884
    val (indrule_lemma_prems, indrule_lemma_concls) =
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1885
      Library.foldl mk_indrule_lemma (([], []), (List.concat descr' ~~ recTs ~~ recTs'));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1886
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1887
    val indrule_lemma = standard (Goal.prove thy8 [] []
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1888
      (Logic.mk_implies
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1889
        (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1890
         HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1891
           [REPEAT (etac conjE 1),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1892
            REPEAT (EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1893
              [TRY (rtac conjI 1), full_simp_tac (HOL_basic_ss addsimps Rep_inverse_thms) 1,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1894
               etac mp 1, resolve_tac Rep_thms 1])]));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1895
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1896
    val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1897
    val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1898
      map (Free o apfst fst o dest_Var) Ps;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1899
    val indrule_lemma' = cterm_instantiate
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1900
      (map (cterm_of thy8) Ps ~~ map (cterm_of thy8) frees) indrule_lemma;
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1901
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1902
    val Abs_inverse_thms' = List.concat (map mk_funs_inv Abs_inverse_thms);
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1903
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1904
    val dt_induct_prop = DatatypeProp.make_ind descr' sorts';
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1905
    val dt_induct = standard (Goal.prove thy8 []
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1906
      (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1907
      (fn prems => EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1908
        [rtac indrule_lemma' 1,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1909
         (DatatypeAux.indtac rep_induct THEN_ALL_NEW ObjectLogic.atomize_tac) 1,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1910
         EVERY (map (fn (prem, r) => (EVERY
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1911
           [REPEAT (eresolve_tac Abs_inverse_thms' 1),
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1912
            simp_tac (HOL_basic_ss addsimps [symmetric r]) 1,
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1913
            DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1914
                (prems ~~ constr_defs))]));
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1915
18017
f6abeac6dcb5 fixed case names in the weak induction principle and
urbanc
parents: 18016
diff changeset
  1916
    val case_names_induct = mk_case_names_induct (List.concat descr');
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1917
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1918
    (**** prove that new datatypes have finite support ****)
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1919
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1920
    val indnames = DatatypeProp.make_tnames recTs;
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1921
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1922
    val abs_supp = PureThy.get_thms thy8 (Name "abs_supp");
18067
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
  1923
    val supp_atm = PureThy.get_thms thy8 (Name "supp_atm");
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1924
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1925
    val finite_supp_thms = map (fn atom =>
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1926
      let val atomT = Type (atom, [])
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1927
      in map standard (List.take
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1928
        (split_conj_thm (Goal.prove thy8 [] [] (HOLogic.mk_Trueprop
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1929
           (foldr1 HOLogic.mk_conj (map (fn (s, T) => HOLogic.mk_mem
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1930
             (Const ("nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T),
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1931
              Const ("Finite_Set.Finites", HOLogic.mk_setT (HOLogic.mk_setT atomT))))
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1932
               (indnames ~~ recTs))))
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1933
           (fn _ => indtac dt_induct indnames 1 THEN
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1934
            ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps
18067
8b9848d150ba - completed the list of thms for supp_atm
urbanc
parents: 18066
diff changeset
  1935
              (abs_supp @ supp_atm @
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1936
               PureThy.get_thms thy8 (Name ("fs_" ^ Sign.base_name atom ^ "1")) @
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1937
               List.concat supp_thms))))),
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1938
         length new_type_names))
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1939
      end) atoms;
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1940
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1941
    val (thy9, _) = thy8 |>
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1942
      Theory.add_path big_name |>
18017
f6abeac6dcb5 fixed case names in the weak induction principle and
urbanc
parents: 18016
diff changeset
  1943
      PureThy.add_thms [(("induct_weak", dt_induct), [case_names_induct])] |>>
18016
8f3a80033ba4 Implemented proof of weak induction theorem.
berghofe
parents: 18010
diff changeset
  1944
      Theory.parent_path |>>>
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1945
      DatatypeAux.store_thmss "distinct" new_type_names distinct_thms |>>>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1946
      DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss |>>>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1947
      DatatypeAux.store_thmss "perm" new_type_names perm_simps' |>>>
17872
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1948
      DatatypeAux.store_thmss "inject" new_type_names inject_thms |>>>
f08fc98a164a Implemented proofs for support and freshness theorems.
berghofe
parents: 17870
diff changeset
  1949
      DatatypeAux.store_thmss "supp" new_type_names supp_thms |>>>
18066
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1950
      DatatypeAux.store_thmss "fresh" new_type_names fresh_thms |>>
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1951
      fold (fn (atom, ths) => fn thy =>
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1952
        let val class = Sign.intern_class thy ("fs_" ^ Sign.base_name atom)
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1953
        in fold (fn T => AxClass.add_inst_arity_i
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1954
            (fst (dest_Type T),
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1955
              replicate (length sorts) [class], [class])
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1956
            (AxClass.intro_classes_tac [] THEN resolve_tac ths 1)) newTs thy
d1e47ee13070 Added code for proving that new datatype has finite support.
berghofe
parents: 18054
diff changeset
  1957
        end) (atoms ~~ finite_supp_thms);
17870
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1958
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1959
  in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1960
    (thy9, perm_eq_thms)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1961
  end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1962
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1963
val add_nominal_datatype = gen_add_nominal_datatype read_typ true;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1964
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1965
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1966
(* FIXME: The following stuff should be exported by DatatypePackage *)
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1967
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1968
local structure P = OuterParse and K = OuterKeyword in
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1969
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1970
val datatype_decl =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1971
  Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1972
    (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1973
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1974
fun mk_datatype args =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1975
  let
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1976
    val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1977
    val specs = map (fn ((((_, vs), t), mx), cons) =>
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1978
      (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1979
  in #1 o add_nominal_datatype false names specs end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1980
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1981
val nominal_datatypeP =
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1982
  OuterSyntax.command "nominal_datatype" "define inductive datatypes" K.thy_decl
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1983
    (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1984
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1985
val _ = OuterSyntax.add_parsers [nominal_datatypeP];
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1986
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1987
end;
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1988
c35381811d5c Initial revision.
berghofe
parents:
diff changeset
  1989
end