src/HOL/Tools/datatype_rep_proofs.ML
author wenzelm
Fri, 31 Aug 2001 18:46:48 +0200
changeset 11539 0f17da240450
parent 11471 ba2c252b55ad
child 11628 e57a6e51715e
permissions -rw-r--r--
tuned headers;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Tools/datatype_rep_proofs.ML
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     2
    ID:         $Id$
11539
0f17da240450 tuned headers;
wenzelm
parents: 11471
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
0f17da240450 tuned headers;
wenzelm
parents: 11471
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     5
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     6
Definitional introduction of datatypes
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     7
Proof of characteristic theorems:
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     8
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     9
 - injectivity of constructors
7228
ddb67dcf026c Tuned some comments.
berghofe
parents: 7205
diff changeset
    10
 - distinctness of constructors
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    11
 - induction theorem
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    12
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    13
*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    14
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    15
signature DATATYPE_REP_PROOFS =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    16
sig
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    17
  val representation_proofs : bool -> DatatypeAux.datatype_info Symtab.table ->
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    18
    string list -> (int * (string * DatatypeAux.dtyp list *
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    19
      (string * DatatypeAux.dtyp list) list)) list list -> (string * sort) list ->
8436
8a87fa482baf adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8305
diff changeset
    20
        (string * mixfix) list -> (string * mixfix) list list -> theory attribute
8a87fa482baf adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8305
diff changeset
    21
          -> theory -> theory * thm list list * thm list list * thm list list *
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    22
            DatatypeAux.simproc_dist list * thm
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    23
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    24
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    25
structure DatatypeRepProofs : DATATYPE_REP_PROOFS =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    26
struct
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    27
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    28
open DatatypeAux;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    29
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    30
val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    31
11435
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
    32
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
    33
(** theory context references **)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    34
11435
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
    35
val f_myinv_f = thm "f_myinv_f";
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
    36
val myinv_f_f = thm "myinv_f_f";
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
    37
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    38
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    39
fun exh_thm_of (dt_info : datatype_info Symtab.table) tname =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    40
  #exhaustion (the (Symtab.lookup (dt_info, tname)));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    41
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    42
(******************************************************************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    43
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    44
fun representation_proofs flat_names (dt_info : datatype_info Symtab.table)
8436
8a87fa482baf adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8305
diff changeset
    45
      new_type_names descr sorts types_syntax constr_syntax case_names_induct thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    46
  let
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    47
    val Datatype_thy = theory "Datatype";
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    48
    val node_name = Sign.intern_tycon (Theory.sign_of Datatype_thy) "node";
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    49
    val [In0_name, In1_name, Scons_name, Leaf_name, Numb_name, Lim_name,
7293
959e060f4a2f Moved sum_case stuff from Sum to Datatype.
berghofe
parents: 7257
diff changeset
    50
      Funs_name, o_name, sum_case_name] =
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    51
      map (Sign.intern_const (Theory.sign_of Datatype_thy))
7293
959e060f4a2f Moved sum_case stuff from Sum to Datatype.
berghofe
parents: 7257
diff changeset
    52
        ["In0", "In1", "Scons", "Leaf", "Numb", "Lim", "Funs", "op o", "sum_case"];
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    53
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    54
    val [In0_inject, In1_inject, Scons_inject, Leaf_inject, In0_eq, In1_eq,
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    55
         In0_not_In1, In1_not_In0, Funs_mono, FunsI, Lim_inject,
7293
959e060f4a2f Moved sum_case stuff from Sum to Datatype.
berghofe
parents: 7257
diff changeset
    56
         Funs_inv, FunsD, Funs_rangeE, Funs_nonempty, sum_case_inject] = map (get_thm Datatype_thy)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    57
        ["In0_inject", "In1_inject", "Scons_inject", "Leaf_inject", "In0_eq", "In1_eq",
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    58
         "In0_not_In1", "In1_not_In0", "Funs_mono", "FunsI", "Lim_inject",
7293
959e060f4a2f Moved sum_case stuff from Sum to Datatype.
berghofe
parents: 7257
diff changeset
    59
         "Funs_inv", "FunsD", "Funs_rangeE", "Funs_nonempty", "sum_case_inject"];
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    60
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    61
    val Funs_IntE = (Int_lower2 RS Funs_mono RS
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    62
      (Int_lower1 RS Funs_mono RS Int_greatest) RS subsetD) RS IntE;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    63
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    64
    val descr' = flat descr;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    65
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    66
    val big_name = space_implode "_" new_type_names;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    67
    val thy1 = add_path flat_names big_name thy;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    68
    val big_rec_name = big_name ^ "_rep_set";
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
    69
    val rep_set_names = map (Sign.full_name (Theory.sign_of thy1))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    70
      (if length descr' = 1 then [big_rec_name] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    71
        (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    72
          (1 upto (length descr'))));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    73
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    74
    val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    75
    val leafTs' = get_nonrec_types descr' sorts;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    76
    val branchTs = get_branching_types descr' sorts;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    77
    val branchT = if null branchTs then HOLogic.unitT
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    78
      else fold_bal (fn (T, U) => Type ("+", [T, U])) branchTs;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    79
    val unneeded_vars = hd tyvars \\ foldr add_typ_tfree_names (leafTs' @ branchTs, []);
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    80
    val leafTs = leafTs' @ (map (fn n => TFree (n, the (assoc (sorts, n)))) unneeded_vars);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    81
    val recTs = get_rec_types descr' sorts;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    82
    val newTs = take (length (hd descr), recTs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    83
    val oldTs = drop (length (hd descr), recTs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    84
    val sumT = if null leafTs then HOLogic.unitT
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    85
      else fold_bal (fn (T, U) => Type ("+", [T, U])) leafTs;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    86
    val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT]));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    87
    val UnivT = HOLogic.mk_setT Univ_elT;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    88
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    89
    val In0 = Const (In0_name, Univ_elT --> Univ_elT);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    90
    val In1 = Const (In1_name, Univ_elT --> Univ_elT);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    91
    val Leaf = Const (Leaf_name, sumT --> Univ_elT);
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    92
    val Lim = Const (Lim_name, (branchT --> Univ_elT) --> Univ_elT);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    93
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    94
    (* make injections needed for embedding types in leaves *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    95
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    96
    fun mk_inj T' x =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    97
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    98
        fun mk_inj' T n i =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    99
          if n = 1 then x else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   100
          let val n2 = n div 2;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   101
              val Type (_, [T1, T2]) = T
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   102
          in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   103
            if i <= n2 then
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   104
              Const ("Inl", T1 --> T) $ (mk_inj' T1 n2 i)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   105
            else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   106
              Const ("Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   107
          end
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   108
      in mk_inj' sumT (length leafTs) (1 + find_index_eq T' leafTs)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   109
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   110
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   111
    (* make injections for constructors *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   112
7704
9a6783fdb9a5 eliminated ap/app;
wenzelm
parents: 7499
diff changeset
   113
    fun mk_univ_inj ts = access_bal (fn t => In0 $ t, fn t => In1 $ t, if ts = [] then
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   114
        Const ("arbitrary", Univ_elT)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   115
      else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   116
        foldr1 (HOLogic.mk_binop Scons_name) ts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   117
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   118
    (* function spaces *)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   119
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   120
    fun mk_fun_inj T' x =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   121
      let
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   122
        fun mk_inj T n i =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   123
          if n = 1 then x else
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   124
          let
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   125
            val n2 = n div 2;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   126
            val Type (_, [T1, T2]) = T;
7293
959e060f4a2f Moved sum_case stuff from Sum to Datatype.
berghofe
parents: 7257
diff changeset
   127
            val sum_case = Const (sum_case_name, [T1 --> Univ_elT, T2 --> Univ_elT, T] ---> Univ_elT)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   128
          in
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   129
            if i <= n2 then
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   130
              sum_case $ (mk_inj T1 n2 i) $ Const ("arbitrary", T2 --> Univ_elT)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   131
            else
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   132
              sum_case $ Const ("arbitrary", T1 --> Univ_elT) $ mk_inj T2 (n - n2) (i - n2)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   133
          end
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   134
      in mk_inj branchT (length branchTs) (1 + find_index_eq T' branchTs)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   135
      end;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   136
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   137
    (************** generate introduction rules for representing set **********)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   138
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
   139
    val _ = message "Constructing representing sets ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   140
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   141
    (* make introduction rule for a single constructor *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   142
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   143
    fun make_intr s n (i, (_, cargs)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   144
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   145
        fun mk_prem (DtRec k, (j, prems, ts)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   146
              let val free_t = mk_Free "x" Univ_elT j
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   147
              in (j + 1, (HOLogic.mk_mem (free_t,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   148
                Const (nth_elem (k, rep_set_names), UnivT)))::prems, free_t::ts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   149
              end
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   150
          | mk_prem (DtType ("fun", [T, DtRec k]), (j, prems, ts)) =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   151
              let val T' = typ_of_dtyp descr' sorts T;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   152
                  val free_t = mk_Free "x" (T' --> Univ_elT) j
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   153
              in (j + 1, (HOLogic.mk_mem (free_t,
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   154
                Const (Funs_name, UnivT --> HOLogic.mk_setT (T' --> Univ_elT)) $
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   155
                  Const (nth_elem (k, rep_set_names), UnivT)))::prems,
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   156
                    Lim $ mk_fun_inj T' free_t::ts)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   157
              end
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   158
          | mk_prem (dt, (j, prems, ts)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   159
              let val T = typ_of_dtyp descr' sorts dt
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   160
              in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   161
              end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   162
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   163
        val (_, prems, ts) = foldr mk_prem (cargs, (1, [], []));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   164
        val concl = HOLogic.mk_Trueprop (HOLogic.mk_mem
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   165
          (mk_univ_inj ts n i, Const (s, UnivT)))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   166
      in Logic.list_implies (map HOLogic.mk_Trueprop prems, concl)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   167
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   168
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   169
    val consts = map (fn s => Const (s, UnivT)) rep_set_names;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   170
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   171
    val intr_ts = flat (map (fn ((_, (_, _, constrs)), rep_set_name) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   172
      map (make_intr rep_set_name (length constrs))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   173
        ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   174
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   175
    val (thy2, {raw_induct = rep_induct, intrs = rep_intrs, ...}) =
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
   176
      setmp InductivePackage.quiet_mode (!quiet_mode)
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
   177
        (InductivePackage.add_inductive_i false true big_rec_name false true false
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   178
           consts [] (map (fn x => (("", x), [])) intr_ts) [Funs_mono] []) thy1;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   179
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   180
    (********************************* typedef ********************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   181
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
   182
    val thy3 = add_path flat_names big_name (foldl (fn (thy, ((((name, mx), tvs), c), name')) =>
5696
c2c2214f8037 quiet proofs;
wenzelm
parents: 5661
diff changeset
   183
      setmp TypedefPackage.quiet_mode true
c2c2214f8037 quiet proofs;
wenzelm
parents: 5661
diff changeset
   184
        (TypedefPackage.add_typedef_i_no_def name' (name, tvs, mx) c [] []
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   185
          (Some (QUIET_BREADTH_FIRST (has_fewer_prems 1)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   186
            (resolve_tac (Funs_nonempty::rep_intrs) 1)))) thy)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   187
              (parent_path flat_names thy2, types_syntax ~~ tyvars ~~
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   188
                (take (length newTs, consts)) ~~ new_type_names));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   189
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   190
    (*********************** definition of constructors ***********************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   191
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   192
    val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   193
    val rep_names = map (curry op ^ "Rep_") new_type_names;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   194
    val rep_names' = map (fn i => big_rep_name ^ (string_of_int i))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   195
      (1 upto (length (flat (tl descr))));
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   196
    val all_rep_names = map (Sign.intern_const (Theory.sign_of thy3)) rep_names @
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   197
      map (Sign.full_name (Theory.sign_of thy3)) rep_names';
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   198
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   199
    (* isomorphism declarations *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   200
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   201
    val iso_decls = map (fn (T, s) => (s, T --> Univ_elT, NoSyn))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   202
      (oldTs ~~ rep_names');
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   203
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   204
    (* constructor definitions *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   205
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   206
    fun make_constr_def tname T n ((thy, defs, eqns, i), ((cname, cargs), (cname', mx))) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   207
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   208
        fun constr_arg (dt, (j, l_args, r_args)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   209
          let val T = typ_of_dtyp descr' sorts dt;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   210
              val free_t = mk_Free "x" T j
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   211
          in (case dt of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   212
              DtRec m => (j + 1, free_t::l_args, (Const (nth_elem (m, all_rep_names),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   213
                T --> Univ_elT) $ free_t)::r_args)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   214
            | DtType ("fun", [T', DtRec m]) =>
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   215
                let val ([T''], T''') = strip_type T
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   216
                in (j + 1, free_t::l_args, (Lim $ mk_fun_inj T''
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   217
                  (Const (o_name, [T''' --> Univ_elT, T, T''] ---> Univ_elT) $
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   218
                    Const (nth_elem (m, all_rep_names), T''' --> Univ_elT) $ free_t))::r_args)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   219
                end
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   220
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   221
            | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   222
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   223
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   224
        val (_, l_args, r_args) = foldr constr_arg (cargs, (1, [], []));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   225
        val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T;
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   226
        val abs_name = Sign.intern_const (Theory.sign_of thy) ("Abs_" ^ tname);
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   227
        val rep_name = Sign.intern_const (Theory.sign_of thy) ("Rep_" ^ tname);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   228
        val lhs = list_comb (Const (cname, constrT), l_args);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   229
        val rhs = mk_univ_inj r_args n i;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   230
        val def = equals T $ lhs $ (Const (abs_name, Univ_elT --> T) $ rhs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   231
        val def_name = (Sign.base_name cname) ^ "_def";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   232
        val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   233
          (Const (rep_name, T --> Univ_elT) $ lhs, rhs));
8436
8a87fa482baf adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8305
diff changeset
   234
        val (thy', [def_thm]) = thy |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   235
          Theory.add_consts_i [(cname', constrT, mx)] |>
9315
f793f05024f6 adapted PureThy.add_defs_i;
wenzelm
parents: 8479
diff changeset
   236
          (PureThy.add_defs_i false o map Thm.no_attributes) [(def_name, def)];
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   237
8436
8a87fa482baf adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8305
diff changeset
   238
      in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   239
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   240
    (* constructor definitions for datatype *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   241
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   242
    fun dt_constr_defs ((thy, defs, eqns, rep_congs, dist_lemmas),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   243
        ((((_, (_, _, constrs)), tname), T), constr_syntax)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   244
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   245
        val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   246
        val sg = Theory.sign_of thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   247
        val rep_const = cterm_of sg
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   248
          (Const (Sign.intern_const sg ("Rep_" ^ tname), T --> Univ_elT));
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   249
        val cong' = standard (cterm_instantiate [(cterm_of sg cong_f, rep_const)] arg_cong);
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   250
        val dist = standard (cterm_instantiate [(cterm_of sg distinct_f, rep_const)] distinct_lemma);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   251
        val (thy', defs', eqns', _) = foldl ((make_constr_def tname T) (length constrs))
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
   252
          ((add_path flat_names tname thy, defs, [], 1), constrs ~~ constr_syntax)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   253
      in
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
   254
        (parent_path flat_names thy', defs', eqns @ [eqns'],
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   255
          rep_congs @ [cong'], dist_lemmas @ [dist])
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   256
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   257
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   258
    val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = foldl dt_constr_defs
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
   259
      ((thy3 |> Theory.add_consts_i iso_decls |> parent_path flat_names, [], [], [], []),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   260
        hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   261
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   262
    (*********** isomorphisms for new types (introduced by typedef) ***********)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   263
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
   264
    val _ = message "Proving isomorphism properties ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   265
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   266
    (* get axioms from theory *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   267
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   268
    val newT_iso_axms = map (fn s =>
7904
2b551893583e proper handling of axioms / defs;
wenzelm
parents: 7704
diff changeset
   269
      (get_thm thy4 ("Abs_" ^ s ^ "_inverse"),
2b551893583e proper handling of axioms / defs;
wenzelm
parents: 7704
diff changeset
   270
       get_thm thy4 ("Rep_" ^ s ^ "_inverse"),
2b551893583e proper handling of axioms / defs;
wenzelm
parents: 7704
diff changeset
   271
       get_thm thy4 ("Rep_" ^ s))) new_type_names;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   272
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   273
    (*------------------------------------------------*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   274
    (* prove additional theorems:                     *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   275
    (*  inj_on dt_Abs_i rep_set_i  and  inj dt_Rep_i  *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   276
    (*------------------------------------------------*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   277
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   278
    fun prove_newT_iso_inj_thm (((s, (thm1, thm2, _)), T), rep_set_name) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   279
      let
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   280
        val sg = Theory.sign_of thy4;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   281
        val RepT = T --> Univ_elT;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   282
        val Rep_name = Sign.intern_const sg ("Rep_" ^ s);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   283
        val AbsT = Univ_elT --> T;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   284
        val Abs_name = Sign.intern_const sg ("Abs_" ^ s);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   285
6171
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   286
        val inj_Abs_thm = 
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   287
	    prove_goalw_cterm [] 
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   288
	      (cterm_of sg
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   289
	       (HOLogic.mk_Trueprop 
11435
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   290
		(Const ("Fun.inj_on", [AbsT, UnivT] ---> HOLogic.boolT) $
6171
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   291
		 Const (Abs_name, AbsT) $ Const (rep_set_name, UnivT))))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   292
              (fn _ => [rtac inj_on_inverseI 1, etac thm1 1]);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   293
6171
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   294
        val setT = HOLogic.mk_setT T
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   295
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   296
        val inj_Rep_thm =
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   297
	    prove_goalw_cterm []
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   298
	      (cterm_of sg
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   299
	       (HOLogic.mk_Trueprop
11435
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   300
		(Const ("Fun.inj_on", [RepT, setT] ---> HOLogic.boolT) $
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   301
		 Const (Rep_name, RepT) $ Const ("UNIV", setT))))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   302
              (fn _ => [rtac inj_inverseI 1, rtac thm2 1])
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   303
6171
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   304
      in (inj_Abs_thm, inj_Rep_thm) end;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   305
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   306
    val newT_iso_inj_thms = map prove_newT_iso_inj_thm
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   307
      (new_type_names ~~ newT_iso_axms ~~ newTs ~~
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   308
        take (length newTs, rep_set_names));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   309
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   310
    (********* isomorphisms between existing types and "unfolded" types *******)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   311
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   312
    (*---------------------------------------------------------------------*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   313
    (* isomorphisms are defined using primrec-combinators:                 *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   314
    (* generate appropriate functions for instantiating primrec-combinator *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   315
    (*                                                                     *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   316
    (*   e.g.  dt_Rep_i = list_rec ... (%h t y. In1 ((Leaf h) $ y))        *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   317
    (*                                                                     *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   318
    (* also generate characteristic equations for isomorphisms             *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   319
    (*                                                                     *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   320
    (*   e.g.  dt_Rep_i (cons h t) = In1 ((dt_Rep_j h) $ (dt_Rep_i t))     *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   321
    (*---------------------------------------------------------------------*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   322
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   323
    fun make_iso_def k ks n ((fs, eqns, i), (cname, cargs)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   324
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   325
        val argTs = map (typ_of_dtyp descr' sorts) cargs;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   326
        val T = nth_elem (k, recTs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   327
        val rep_name = nth_elem (k, all_rep_names);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   328
        val rep_const = Const (rep_name, T --> Univ_elT);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   329
        val constr = Const (cname, argTs ---> T);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   330
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   331
        fun process_arg ks' ((i2, i2', ts, Ts), dt) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   332
          let val T' = typ_of_dtyp descr' sorts dt
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   333
          in (case dt of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   334
              DtRec j => if j mem ks' then
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   335
                  (i2 + 1, i2' + 1, ts @ [mk_Free "y" Univ_elT i2'], Ts @ [Univ_elT])
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   336
                else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   337
                  (i2 + 1, i2', ts @ [Const (nth_elem (j, all_rep_names),
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   338
                    T' --> Univ_elT) $ mk_Free "x" T' i2], Ts)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   339
            | (DtType ("fun", [_, DtRec j])) =>
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   340
                let val ([T''], T''') = strip_type T'
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   341
                in if j mem ks' then
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   342
                    (i2 + 1, i2' + 1, ts @ [Lim $ mk_fun_inj T''
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   343
                      (mk_Free "y" (T'' --> Univ_elT) i2')], Ts @ [T'' --> Univ_elT])
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   344
                  else
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   345
                    (i2 + 1, i2', ts @ [Lim $ mk_fun_inj T''
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   346
                      (Const (o_name, [T''' --> Univ_elT, T', T''] ---> Univ_elT) $
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   347
                        Const (nth_elem (j, all_rep_names), T''' --> Univ_elT) $
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   348
                          mk_Free "x" T' i2)], Ts)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   349
                end
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   350
            | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   351
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   352
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   353
        val (i2, i2', ts, Ts) = foldl (process_arg ks) ((1, 1, [], []), cargs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   354
        val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   355
        val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   356
        val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   357
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   358
        val (_, _, ts', _) = foldl (process_arg []) ((1, 1, [], []), cargs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   359
        val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   360
          (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   361
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   362
      in (fs @ [f], eqns @ [eqn], i + 1) end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   363
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   364
    (* define isomorphisms for all mutually recursive datatypes in list ds *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   365
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   366
    fun make_iso_defs (ds, (thy, char_thms)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   367
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   368
        val ks = map fst ds;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   369
        val (_, (tname, _, _)) = hd ds;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   370
        val {rec_rewrites, rec_names, ...} = the (Symtab.lookup (dt_info, tname));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   371
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   372
        fun process_dt ((fs, eqns, isos), (k, (tname, _, constrs))) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   373
          let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   374
            val (fs', eqns', _) = foldl (make_iso_def k ks (length constrs))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   375
              ((fs, eqns, 1), constrs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   376
            val iso = (nth_elem (k, recTs), nth_elem (k, all_rep_names))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   377
          in (fs', eqns', isos @ [iso]) end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   378
        
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   379
        val (fs, eqns, isos) = foldl process_dt (([], [], []), ds);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   380
        val fTs = map fastype_of fs;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   381
        val defs = map (fn (rec_name, (T, iso_name)) => ((Sign.base_name iso_name) ^ "_def",
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   382
          equals (T --> Univ_elT) $ Const (iso_name, T --> Univ_elT) $
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   383
            list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs))) (rec_names ~~ isos);
9315
f793f05024f6 adapted PureThy.add_defs_i;
wenzelm
parents: 8479
diff changeset
   384
        val (thy', def_thms) = (PureThy.add_defs_i false o map Thm.no_attributes) defs thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   385
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   386
        (* prove characteristic equations *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   387
5553
ae42b36a50c2 renamed mk_meta_eq to mk_eq
oheimb
parents: 5303
diff changeset
   388
        val rewrites = def_thms @ (map mk_meta_eq rec_rewrites);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   389
        val char_thms' = map (fn eqn => prove_goalw_cterm rewrites
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   390
          (cterm_of (Theory.sign_of thy') eqn) (fn _ => [rtac refl 1])) eqns;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   391
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   392
      in (thy', char_thms' @ char_thms) end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   393
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
   394
    val (thy5, iso_char_thms) = foldr make_iso_defs
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
   395
      (tl descr, (add_path flat_names big_name thy4, []));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   396
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   397
    (* prove isomorphism properties *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   398
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   399
    fun mk_funs_inv thm =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   400
      let
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   401
        val [_, t] = prems_of Funs_inv;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   402
        val [_ $ (_ $ _ $ R)] = Logic.strip_assums_hyp t;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   403
        val _ $ (_ $ (r $ (a $ _)) $ _) = Logic.strip_assums_concl t;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   404
        val [_ $ (_ $ _ $ R')] = prems_of thm;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   405
        val _ $ (_ $ (r' $ (a' $ _)) $ _) = concl_of thm;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   406
        val inv' = cterm_instantiate (map 
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   407
          ((pairself (cterm_of (sign_of_thm thm))) o
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   408
           (apsnd (map_term_types (incr_tvar 1))))
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   409
             [(R, R'), (r, r'), (a, a')]) Funs_inv
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   410
      in
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   411
        rule_by_tactic (atac 2) (thm RSN (2, inv'))
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   412
      end;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   413
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   414
    (* prove  inj dt_Rep_i  and  dt_Rep_i x : dt_rep_set_i *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   415
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   416
    fun prove_iso_thms (ds, (inj_thms, elem_thms)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   417
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   418
        val (_, (tname, _, _)) = hd ds;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   419
        val {induction, ...} = the (Symtab.lookup (dt_info, tname));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   420
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   421
        fun mk_ind_concl (i, _) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   422
          let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   423
            val T = nth_elem (i, recTs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   424
            val Rep_t = Const (nth_elem (i, all_rep_names), T --> Univ_elT);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   425
            val rep_set_name = nth_elem (i, rep_set_names)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   426
          in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   427
                HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   428
                  HOLogic.mk_eq (mk_Free "x" T i, Bound 0)),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   429
              HOLogic.mk_mem (Rep_t $ mk_Free "x" T i, Const (rep_set_name, UnivT)))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   430
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   431
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   432
        val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   433
5553
ae42b36a50c2 renamed mk_meta_eq to mk_eq
oheimb
parents: 5303
diff changeset
   434
        val rewrites = map mk_meta_eq iso_char_thms;
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   435
        val inj_thms' = flat (map (fn r => [r RS injD, r RS inj_o])
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   436
          (map snd newT_iso_inj_thms @ inj_thms));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   437
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   438
        val inj_thm = prove_goalw_cterm [] (cterm_of (Theory.sign_of thy5)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   439
          (HOLogic.mk_Trueprop (mk_conj ind_concl1))) (fn _ =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   440
            [indtac induction 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   441
             REPEAT (EVERY
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   442
               [rtac allI 1, rtac impI 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   443
                exh_tac (exh_thm_of dt_info) 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   444
                REPEAT (EVERY
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   445
                  [hyp_subst_tac 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   446
                   rewrite_goals_tac rewrites,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   447
                   REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   448
                   (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   449
                   ORELSE (EVERY
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   450
                     [REPEAT (etac Scons_inject 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   451
                      REPEAT (dresolve_tac
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   452
                        (inj_thms' @ [Leaf_inject, Lim_inject, Inl_inject, Inr_inject]) 1),
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   453
                      REPEAT ((EVERY [etac allE 1, dtac mp 1, atac 1]) ORELSE
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   454
                              (dtac inj_fun_lemma 1 THEN atac 1)),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   455
                      TRY (hyp_subst_tac 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   456
                      rtac refl 1])])])]);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   457
6171
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   458
        val inj_thms'' = map (fn r => r RS datatype_injI)
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   459
                             (split_conj_thm inj_thm);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   460
6171
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   461
        val elem_thm = 
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   462
	    prove_goalw_cterm []
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   463
	      (cterm_of (Theory.sign_of thy5)
6171
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   464
	       (HOLogic.mk_Trueprop (mk_conj ind_concl2)))
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   465
	      (fn _ =>
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   466
	       [indtac induction 1,
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   467
		rewrite_goals_tac (o_def :: rewrites),
6171
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   468
		REPEAT (EVERY
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   469
			[resolve_tac rep_intrs 1,
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   470
			 REPEAT (FIRST [atac 1, etac spec 1,
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   471
				 resolve_tac (FunsI :: elem_thms) 1])])]);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   472
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   473
      in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm))
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   474
      end;
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   475
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   476
    val (iso_inj_thms_unfolded, iso_elem_thms) = foldr prove_iso_thms
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   477
      (tl descr, ([], map #3 newT_iso_axms));
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   478
    val iso_inj_thms = map snd newT_iso_inj_thms @ iso_inj_thms_unfolded;
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   479
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   480
    (* prove  x : dt_rep_set_i --> x : range dt_Rep_i *)
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   481
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   482
    fun mk_iso_t (((set_name, iso_name), i), T) =
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   483
      let val isoT = T --> Univ_elT
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   484
      in HOLogic.imp $ 
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   485
        HOLogic.mk_mem (mk_Free "x" Univ_elT i, Const (set_name, UnivT)) $
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   486
          (if i < length newTs then Const ("True", HOLogic.boolT)
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   487
           else HOLogic.mk_mem (mk_Free "x" Univ_elT i,
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   488
             Const ("image", [isoT, HOLogic.mk_setT T] ---> UnivT) $
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   489
               Const (iso_name, isoT) $ Const ("UNIV", HOLogic.mk_setT T)))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   490
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   491
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   492
    val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   493
      (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs)));
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   494
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   495
    (* all the theorems are proved by one single simultaneous induction *)
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   496
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   497
    val iso_thms = if length descr = 1 then [] else
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   498
      drop (length newTs, split_conj_thm
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   499
        (prove_goalw_cterm [] (cterm_of (Theory.sign_of thy5) iso_t) (fn _ =>
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   500
           [indtac rep_induct 1,
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   501
            REPEAT (rtac TrueI 1),
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   502
            REPEAT (EVERY
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   503
              [rewrite_goals_tac [mk_meta_eq Collect_mem_eq],
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   504
               REPEAT (etac Funs_IntE 1),
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   505
               REPEAT (eresolve_tac (rangeE ::
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   506
                 map (fn r => r RS Funs_rangeE) iso_inj_thms_unfolded) 1),
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   507
               REPEAT (eresolve_tac (map (fn (iso, _, _) => iso RS subst) newT_iso_axms @
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   508
                 map (fn (iso, _, _) => mk_funs_inv iso RS subst) newT_iso_axms) 1),
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   509
               TRY (hyp_subst_tac 1),
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   510
               rtac (sym RS range_eqI) 1,
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   511
               resolve_tac iso_char_thms 1])])));
11435
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   512
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   513
    val Abs_inverse_thms' =
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   514
      map #1 newT_iso_axms @
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   515
      map2 (fn (r_inj, r) => f_myinv_f OF [r_inj, r RS mp])
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   516
        (iso_inj_thms_unfolded, iso_thms);
11435
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   517
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   518
    val Abs_inverse_thms = map (fn r => r RS subst) (Abs_inverse_thms' @
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   519
      map mk_funs_inv Abs_inverse_thms');
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   520
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   521
    (******************* freeness theorems for constructors *******************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   522
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
   523
    val _ = message "Proving freeness of constructors ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   524
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   525
    (* prove theorem  Rep_i (Constr_j ...) = Inj_j ...  *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   526
    
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   527
    fun prove_constr_rep_thm eqn =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   528
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   529
        val inj_thms = map (fn (r, _) => r RS inj_onD) newT_iso_inj_thms;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   530
        val rewrites = o_def :: constr_defs @ (map (mk_meta_eq o #2) newT_iso_axms)
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   531
      in prove_goalw_cterm [] (cterm_of (Theory.sign_of thy5) eqn) (fn _ =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   532
        [resolve_tac inj_thms 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   533
         rewrite_goals_tac rewrites,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   534
         rtac refl 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   535
         resolve_tac rep_intrs 2,
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   536
         REPEAT (resolve_tac (FunsI :: iso_elem_thms) 1)])
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   537
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   538
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   539
    (*--------------------------------------------------------------*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   540
    (* constr_rep_thms and rep_congs are used to prove distinctness *)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   541
    (* of constructors.                                             *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   542
    (*--------------------------------------------------------------*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   543
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   544
    val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   545
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   546
    val dist_rewrites = map (fn (rep_thms, dist_lemma) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   547
      dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   548
        (constr_rep_thms ~~ dist_lemmas);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   549
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   550
    fun prove_distinct_thms (_, []) = []
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   551
      | prove_distinct_thms (dist_rewrites', t::_::ts) =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   552
          let
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   553
            val dist_thm = prove_goalw_cterm [] (cterm_of (Theory.sign_of thy5) t) (fn _ =>
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   554
              [simp_tac (HOL_ss addsimps dist_rewrites') 1])
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   555
          in dist_thm::(standard (dist_thm RS not_sym))::
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   556
            (prove_distinct_thms (dist_rewrites', ts))
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   557
          end;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   558
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   559
    val distinct_thms = map prove_distinct_thms (dist_rewrites ~~
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   560
      DatatypeProp.make_distincts new_type_names descr sorts thy5);
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   561
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   562
    val simproc_dists = map (fn ((((_, (_, _, constrs)), rep_thms), congr), dists) =>
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   563
      if length constrs < !DatatypeProp.dtK then FewConstrs dists
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   564
      else ManyConstrs (congr, HOL_basic_ss addsimps rep_thms)) (hd descr ~~
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   565
        constr_rep_thms ~~ rep_congs ~~ distinct_thms);
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   566
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   567
    (* prove injectivity of constructors *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   568
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   569
    fun prove_constr_inj_thm rep_thms t =
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   570
      let val inj_thms = Scons_inject::sum_case_inject::(map make_elim
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   571
        ((map (fn r => r RS injD) iso_inj_thms) @
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   572
          [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject, Lim_inject]))
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   573
      in prove_goalw_cterm [] (cterm_of (Theory.sign_of thy5) t) (fn _ =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   574
        [rtac iffI 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   575
         REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   576
         dresolve_tac rep_congs 1, dtac box_equals 1,
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7293
diff changeset
   577
         REPEAT (resolve_tac rep_thms 1), rewtac o_def,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   578
         REPEAT (eresolve_tac inj_thms 1),
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   579
         REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [rtac ext 1, dtac fun_cong 1,
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   580
                  eresolve_tac inj_thms 1, atac 1]))])
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   581
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   582
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   583
    val constr_inject = map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   584
      ((DatatypeProp.make_injs descr sorts) ~~ constr_rep_thms);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   585
8479
5d327a46dc61 Now returns theorems with correct names in derivations.
berghofe
parents: 8436
diff changeset
   586
    val (thy6, (constr_inject', distinct_thms'))= thy5 |> parent_path flat_names |>
5d327a46dc61 Now returns theorems with correct names in derivations.
berghofe
parents: 8436
diff changeset
   587
      store_thmss "inject" new_type_names constr_inject |>>>
5d327a46dc61 Now returns theorems with correct names in derivations.
berghofe
parents: 8436
diff changeset
   588
      store_thmss "distinct" new_type_names distinct_thms;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   589
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   590
    (*************************** induction theorem ****************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   591
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
   592
    val _ = message "Proving induction rule for datatypes ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   593
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   594
    val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   595
      (map (fn r => r RS myinv_f_f RS subst) iso_inj_thms_unfolded);
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   596
    val Rep_inverse_thms' = map (fn r => r RS myinv_f_f) iso_inj_thms_unfolded;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   597
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   598
    fun mk_indrule_lemma ((prems, concls), ((i, _), T)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   599
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   600
        val Rep_t = Const (nth_elem (i, all_rep_names), T --> Univ_elT) $
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   601
          mk_Free "x" T i;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   602
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   603
        val Abs_t = if i < length newTs then
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   604
            Const (Sign.intern_const (Theory.sign_of thy6)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   605
              ("Abs_" ^ (nth_elem (i, new_type_names))), Univ_elT --> T)
11435
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   606
          else Const ("Inductive.myinv", [T --> Univ_elT, Univ_elT] ---> T) $
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   607
            Const (nth_elem (i, all_rep_names), T --> Univ_elT)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   608
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   609
      in (prems @ [HOLogic.imp $ HOLogic.mk_mem (Rep_t,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   610
            Const (nth_elem (i, rep_set_names), UnivT)) $
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   611
              (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   612
          concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   613
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   614
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   615
    val (indrule_lemma_prems, indrule_lemma_concls) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   616
      foldl mk_indrule_lemma (([], []), (descr' ~~ recTs));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   617
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6171
diff changeset
   618
    val cert = cterm_of (Theory.sign_of thy6);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   619
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   620
    val indrule_lemma = prove_goalw_cterm [] (cert
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   621
      (Logic.mk_implies
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   622
        (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   623
         HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls)))) (fn prems =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   624
           [cut_facts_tac prems 1, REPEAT (etac conjE 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   625
            REPEAT (EVERY
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   626
              [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   627
               etac mp 1, resolve_tac iso_elem_thms 1])]);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   628
8305
93aa21ec5494 HOLogic.dest_conj;
wenzelm
parents: 8005
diff changeset
   629
    val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   630
    val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   631
      map (Free o apfst fst o dest_Var) Ps;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   632
    val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   633
10911
eb5721204b38 proper induction rule for arbitrarily branching datatype;
wenzelm
parents: 9315
diff changeset
   634
    val dt_induct = prove_goalw_cterm [InductivePackage.inductive_forall_def] (cert
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   635
      (DatatypeProp.make_ind descr sorts)) (fn prems =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   636
        [rtac indrule_lemma' 1, indtac rep_induct 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   637
         EVERY (map (fn (prem, r) => (EVERY
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   638
           [REPEAT (eresolve_tac (Funs_IntE::Abs_inverse_thms) 1),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   639
            simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1,
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7293
diff changeset
   640
            DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE (EVERY [rewtac o_def,
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   641
              rtac allI 1, dtac FunsD 1, etac CollectD 1]))]))
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   642
                (prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   643
8436
8a87fa482baf adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8305
diff changeset
   644
    val (thy7, [dt_induct']) = thy6 |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
   645
      Theory.add_path big_name |>
10911
eb5721204b38 proper induction rule for arbitrarily branching datatype;
wenzelm
parents: 9315
diff changeset
   646
      PureThy.add_thms [(("induct", dt_induct),
eb5721204b38 proper induction rule for arbitrarily branching datatype;
wenzelm
parents: 9315
diff changeset
   647
        [Drule.rule_attribute (K InductivePackage.rulify), case_names_induct])] |>>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
   648
      Theory.parent_path;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   649
8479
5d327a46dc61 Now returns theorems with correct names in derivations.
berghofe
parents: 8436
diff changeset
   650
  in (thy7, constr_inject', distinct_thms', dist_rewrites, simproc_dists, dt_induct')
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   651
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   652
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   653
end;