src/HOL/Tools/datatype_abs_proofs.ML
author wenzelm
Sun, 27 Feb 2000 15:26:47 +0100
changeset 8305 93aa21ec5494
parent 7904 2b551893583e
child 8436 8a87fa482baf
permissions -rw-r--r--
HOLogic.dest_conj;
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_abs_proofs.ML
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     2
    ID:         $Id$
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     3
    Author:     Stefan Berghofer
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     4
    Copyright   1998  TU Muenchen
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     5
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     6
Proofs and defintions independent of concrete representation
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     7
of datatypes  (i.e. requiring only abstract properties such as
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     8
injectivity / distinctness of constructors and induction)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     9
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    10
 - case distinction (exhaustion) theorems
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    11
 - characteristic equations for primrec combinators
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    12
 - characteristic equations for case combinators
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    13
 - equations for splitting "P (case ...)" expressions
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    14
 - datatype size function
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    15
 - "nchotomy" and "case_cong" theorems for TFL
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    16
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    17
*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    18
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    19
signature DATATYPE_ABS_PROOFS =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    20
sig
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    21
  val prove_casedist_thms : string list -> (int * (string * DatatypeAux.dtyp list *
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    22
    (string * DatatypeAux.dtyp list) list)) list list -> (string * sort) list ->
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    23
      thm -> theory -> theory * thm list
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    24
  val prove_primrec_thms : bool -> string list -> (int * (string * DatatypeAux.dtyp list *
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    25
    (string * DatatypeAux.dtyp list) list)) list list -> (string * sort) list ->
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    26
      DatatypeAux.datatype_info Symtab.table -> thm list list -> thm list list ->
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    27
        simpset -> thm -> theory -> theory * string list * thm list
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    28
  val prove_case_thms : bool -> string list -> (int * (string * DatatypeAux.dtyp list *
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    29
    (string * DatatypeAux.dtyp list) list)) list list -> (string * sort) list ->
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    30
      string list -> thm list -> theory -> theory * string list * thm list list
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    31
  val prove_split_thms : string list -> (int * (string * DatatypeAux.dtyp list *
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    32
    (string * DatatypeAux.dtyp list) list)) list list -> (string * sort) list ->
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    33
      thm list list -> thm list list -> thm list -> thm list list -> theory ->
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    34
        theory * (thm * thm) list
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    35
  val prove_size_thms : bool -> string list -> (int * (string * DatatypeAux.dtyp list *
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    36
    (string * DatatypeAux.dtyp list) list)) list list -> (string * sort) list ->
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    37
      string list -> thm list -> theory -> theory * thm list
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    38
  val prove_nchotomys : string list -> (int * (string * DatatypeAux.dtyp list *
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    39
    (string * DatatypeAux.dtyp list) list)) list list -> (string * sort) list ->
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    40
      thm list -> theory -> theory * thm list
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    41
  val prove_case_congs : string list -> (int * (string * DatatypeAux.dtyp list *
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    42
    (string * DatatypeAux.dtyp list) list)) list list -> (string * sort) list ->
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    43
      thm list -> thm list list -> theory -> theory * thm list
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    44
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    45
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    46
structure DatatypeAbsProofs : DATATYPE_ABS_PROOFS =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    47
struct
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    48
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    49
open DatatypeAux;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    50
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
    51
val thin = read_instantiate_sg (Theory.sign_of Set.thy) [("V", "?X : ?Y")] thin_rl;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    52
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    53
val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    54
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    55
(************************ case distinction theorems ***************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    56
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    57
fun prove_casedist_thms new_type_names descr sorts induct thy =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    58
  let
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
    59
    val _ = message "Proving case distinction theorems ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    60
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    61
    val descr' = flat descr;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    62
    val recTs = get_rec_types descr' sorts;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    63
    val newTs = take (length (hd descr), recTs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    64
8305
93aa21ec5494 HOLogic.dest_conj;
wenzelm
parents: 7904
diff changeset
    65
    val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    66
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    67
    fun prove_casedist_thm ((i, t), T) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    68
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    69
        val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    70
          Abs ("z", T', Const ("True", T''))) induct_Ps;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    71
        val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", 0), T), Bound 0) $
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    72
          Var (("P", 0), HOLogic.boolT))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    73
        val insts = take (i, dummyPs) @ (P::(drop (i + 1, dummyPs)));
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
    74
        val cert = cterm_of (Theory.sign_of thy);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    75
        val insts' = (map cert induct_Ps) ~~ (map cert insts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    76
        val induct' = refl RS ((nth_elem (i,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    77
          split_conj_thm (cterm_instantiate insts' induct))) RSN (2, rev_mp))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    78
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    79
      in prove_goalw_cterm [] (cert t) (fn prems =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    80
        [rtac induct' 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    81
         REPEAT (rtac TrueI 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    82
         REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    83
         REPEAT (rtac TrueI 1)])
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    84
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    85
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    86
    val casedist_thms = map prove_casedist_thm ((0 upto (length newTs - 1)) ~~
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    87
      (DatatypeProp.make_casedists descr sorts) ~~ newTs)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    88
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    89
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    90
    (store_thms "exhaust" new_type_names casedist_thms thy, casedist_thms)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    91
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    92
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    93
(*************************** primrec combinators ******************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    94
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    95
fun prove_primrec_thms flat_names new_type_names descr sorts
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    96
    (dt_info : datatype_info Symtab.table) constr_inject dist_rewrites dist_ss induct thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    97
  let
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
    98
    val _ = message "Constructing primrec combinators ...";
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    99
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   100
    val fun_rel_comp_name = Sign.intern_const (sign_of Relation.thy) "fun_rel_comp";
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   101
    val [fun_rel_comp_def, o_def] =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   102
      map (get_thm Relation.thy) ["fun_rel_comp_def", "o_def"];
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   103
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   104
    val big_name = space_implode "_" new_type_names;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   105
    val thy0 = add_path flat_names big_name thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   106
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   107
    val descr' = flat descr;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   108
    val recTs = get_rec_types descr' sorts;
5578
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5553
diff changeset
   109
    val used = foldr add_typ_tfree_names (recTs, []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   110
    val newTs = take (length (hd descr), recTs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   111
8305
93aa21ec5494 HOLogic.dest_conj;
wenzelm
parents: 7904
diff changeset
   112
    val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   113
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   114
    val big_rec_name' = big_name ^ "_rec_set";
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
   115
    val rec_set_names = map (Sign.full_name (Theory.sign_of thy0))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   116
      (if length descr' = 1 then [big_rec_name'] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   117
        (map ((curry (op ^) (big_rec_name' ^ "_")) o string_of_int)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   118
          (1 upto (length descr'))));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   119
5578
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5553
diff changeset
   120
    val rec_result_Ts = map TFree (variantlist (replicate (length descr') "'t", used) ~~
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5553
diff changeset
   121
      replicate (length descr') HOLogic.termS);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   122
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   123
    val reccomb_fn_Ts = flat (map (fn (i, (_, _, constrs)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   124
      map (fn (_, cargs) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   125
        let
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   126
          val Ts = map (typ_of_dtyp descr' sorts) cargs;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   127
          val recs = filter (is_rec_type o fst) (cargs ~~ Ts);
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   128
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   129
          fun mk_argT (DtRec k, _) = nth_elem (k, rec_result_Ts)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   130
            | mk_argT (DtType ("fun", [_, DtRec k]), Type ("fun", [T, _])) =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   131
               T --> nth_elem (k, rec_result_Ts);
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   132
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   133
          val argTs = Ts @ map mk_argT recs
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   134
        in argTs ---> nth_elem (i, rec_result_Ts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   135
        end) constrs) descr');
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   136
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   137
    val rec_set_Ts = map (fn (T1, T2) => reccomb_fn_Ts ---> HOLogic.mk_setT
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   138
      (HOLogic.mk_prodT (T1, T2))) (recTs ~~ rec_result_Ts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   139
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   140
    val rec_fns = map (uncurry (mk_Free "f"))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   141
      (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   142
    val rec_sets = map (fn c => list_comb (Const c, rec_fns))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   143
      (rec_set_names ~~ rec_set_Ts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   144
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   145
    (* introduction rules for graph of primrec function *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   146
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   147
    fun make_rec_intr T set_name ((rec_intr_ts, l), (cname, cargs)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   148
      let
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   149
        fun mk_prem ((dt, U), (j, k, prems, t1s, t2s)) =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   150
          let val free1 = mk_Free "x" U j
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   151
          in (case (dt, U) of
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   152
             (DtRec m, _) =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   153
               let val free2 = mk_Free "y" (nth_elem (m, rec_result_Ts)) k
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   154
               in (j + 1, k + 1, (HOLogic.mk_Trueprop (HOLogic.mk_mem
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   155
                 (HOLogic.mk_prod (free1, free2), nth_elem (m, rec_sets))))::prems,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   156
                   free1::t1s, free2::t2s)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   157
               end
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   158
           | (DtType ("fun", [_, DtRec m]), U' as Type ("fun", [T', _])) =>
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   159
               let val free2 = mk_Free "y" (T' --> nth_elem (m, rec_result_Ts)) k
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   160
               in (j + 1, k + 1, (HOLogic.mk_Trueprop (HOLogic.mk_mem (free2,
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   161
                 Const (fun_rel_comp_name, [U', snd (strip_type (nth_elem (m, rec_set_Ts)))] --->
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   162
                   HOLogic.mk_setT (T' --> nth_elem (m, rec_result_Ts))) $
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   163
                     free1 $ nth_elem (m, rec_sets))))::prems, free1::t1s, free2::t2s)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   164
               end
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   165
           | _ => (j + 1, k, prems, free1::t1s, t2s))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   166
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   167
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   168
        val Ts = map (typ_of_dtyp descr' sorts) cargs;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   169
        val (_, _, prems, t1s, t2s) = foldr mk_prem (cargs ~~ Ts, (1, 1, [], [], []))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   170
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   171
      in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop (HOLogic.mk_mem
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   172
        (HOLogic.mk_prod (list_comb (Const (cname, Ts ---> T), t1s),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   173
          list_comb (nth_elem (l, rec_fns), t1s @ t2s)), set_name)))], l + 1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   174
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   175
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   176
    val (rec_intr_ts, _) = foldl (fn (x, ((d, T), set_name)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   177
      foldl (make_rec_intr T set_name) (x, #3 (snd d)))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   178
        (([], 0), descr' ~~ recTs ~~ rec_sets);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   179
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   180
    val (thy1, {intrs = rec_intrs, elims = rec_elims, ...}) =
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   181
      setmp InductivePackage.quiet_mode (!quiet_mode)
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   182
        (InductivePackage.add_inductive_i false true big_rec_name' false false true
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   183
           rec_sets [] (map (fn x => (("", x), [])) rec_intr_ts) [fun_rel_comp_mono] []) thy0;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   184
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   185
    (* prove uniqueness and termination of primrec combinators *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   186
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
   187
    val _ = message "Proving termination and uniqueness of primrec functions ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   188
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   189
    fun mk_unique_tac ((tac, intrs), ((((i, (tname, _, constrs)), elim), T), T')) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   190
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   191
        val distinct_tac = (etac Pair_inject 1) THEN
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   192
          (if i < length newTs then
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   193
             full_simp_tac (HOL_ss addsimps (nth_elem (i, dist_rewrites))) 1
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   194
           else full_simp_tac dist_ss 1);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   195
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   196
        val inject = map (fn r => r RS iffD1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   197
          (if i < length newTs then nth_elem (i, constr_inject)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   198
            else #inject (the (Symtab.lookup (dt_info, tname))));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   199
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   200
        fun mk_unique_constr_tac n ((tac, intr::intrs, j), (cname, cargs)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   201
          let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   202
            val k = length (filter is_rec_type cargs)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   203
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   204
          in (EVERY [DETERM tac,
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   205
                REPEAT (dtac fun_rel_comp_unique 1),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   206
                REPEAT (etac ex1E 1), rtac ex1I 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   207
                DEPTH_SOLVE_1 (ares_tac [intr] 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   208
                REPEAT_DETERM_N k (etac thin 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   209
                etac elim 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   210
                REPEAT_DETERM_N j distinct_tac,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   211
                etac Pair_inject 1, TRY (dresolve_tac inject 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   212
                REPEAT (etac conjE 1), hyp_subst_tac 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   213
                REPEAT (etac allE 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   214
                REPEAT (dtac mp 1 THEN atac 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   215
                TRY (hyp_subst_tac 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   216
                rtac refl 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   217
                REPEAT_DETERM_N (n - j - 1) distinct_tac],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   218
              intrs, j + 1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   219
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   220
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   221
        val (tac', intrs', _) = foldl (mk_unique_constr_tac (length constrs))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   222
          ((tac, intrs, 0), constrs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   223
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   224
      in (tac', intrs') end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   225
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   226
    val rec_unique_thms =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   227
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   228
        val rec_unique_ts = map (fn (((set_t, T1), T2), i) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   229
          Const ("Ex1", (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   230
            absfree ("y", T2, HOLogic.mk_mem (HOLogic.mk_prod
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   231
              (mk_Free "x" T1 i, Free ("y", T2)), set_t)))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   232
                (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
   233
        val cert = cterm_of (Theory.sign_of thy1)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   234
        val insts = map (fn ((i, T), t) => absfree ("x" ^ (string_of_int i), T, t))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   235
          ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   236
        val induct' = cterm_instantiate ((map cert induct_Ps) ~~
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   237
          (map cert insts)) induct;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   238
        val (tac, _) = foldl mk_unique_tac
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   239
          ((rtac induct' 1, rec_intrs), descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   240
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   241
      in split_conj_thm (prove_goalw_cterm []
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   242
        (cert (HOLogic.mk_Trueprop (mk_conj rec_unique_ts))) (K [tac]))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   243
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   244
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   245
    val rec_total_thms = map (fn r =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   246
      r RS ex1_implies_ex RS (select_eq_Ex RS iffD2)) rec_unique_thms;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   247
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   248
    (* define primrec combinators *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   249
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   250
    val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
   251
    val reccomb_names = map (Sign.full_name (Theory.sign_of thy1))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   252
      (if length descr' = 1 then [big_reccomb_name] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   253
        (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   254
          (1 upto (length descr'))));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   255
    val reccombs = map (fn ((name, T), T') => list_comb
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   256
      (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   257
        (reccomb_names ~~ recTs ~~ rec_result_Ts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   258
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   259
    val thy2 = thy1 |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   260
      Theory.add_consts_i (map (fn ((name, T), T') =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   261
        (Sign.base_name name, reccomb_fn_Ts @ [T] ---> T', NoSyn))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   262
          (reccomb_names ~~ recTs ~~ rec_result_Ts)) |>
7904
2b551893583e proper handling of axioms / defs;
wenzelm
parents: 7704
diff changeset
   263
      (PureThy.add_defs_i o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   264
        ((Sign.base_name name) ^ "_def", Logic.mk_equals (comb, absfree ("x", T,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   265
           Const ("Eps", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   266
             HOLogic.mk_mem (HOLogic.mk_prod (Free ("x", T), Free ("y", T')), set))))))
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   267
               (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts)) |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   268
      parent_path flat_names;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   269
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   270
    val reccomb_defs = map ((get_def thy2) o Sign.base_name) reccomb_names;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   271
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   272
    (* prove characteristic equations for primrec combinators *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   273
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
   274
    val _ = message "Proving characteristic theorems for primrec combinators ..."
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   275
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   276
    val rec_thms = map (fn t => prove_goalw_cterm reccomb_defs
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
   277
      (cterm_of (Theory.sign_of thy2) t) (fn _ =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   278
        [rtac select1_equality 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   279
         resolve_tac rec_unique_thms 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   280
         resolve_tac rec_intrs 1,
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   281
         rewrite_goals_tac [o_def, fun_rel_comp_def],
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   282
         REPEAT ((rtac CollectI 1 THEN rtac allI 1) ORELSE resolve_tac rec_total_thms 1)]))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   283
           (DatatypeProp.make_primrecs new_type_names descr sorts thy2)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   284
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   285
  in
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   286
    (thy2 |> Theory.add_path (space_implode "_" new_type_names) |>
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5891
diff changeset
   287
             PureThy.add_thmss [(("recs", rec_thms), [])] |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   288
             Theory.parent_path,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   289
     reccomb_names, rec_thms)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   290
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   291
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   292
(***************************** case combinators *******************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   293
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   294
fun prove_case_thms flat_names new_type_names descr sorts reccomb_names primrec_thms thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   295
  let
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
   296
    val _ = message "Proving characteristic theorems for case combinators ...";
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   297
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   298
    val thy1 = add_path flat_names (space_implode "_" new_type_names) thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   299
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   300
    val descr' = flat descr;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   301
    val recTs = get_rec_types descr' sorts;
5578
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5553
diff changeset
   302
    val used = foldr add_typ_tfree_names (recTs, []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   303
    val newTs = take (length (hd descr), recTs);
5578
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5553
diff changeset
   304
    val T' = TFree (variant used "'t", HOLogic.termS);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   305
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   306
    fun mk_dummyT (DtRec _) = T'
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   307
      | mk_dummyT (DtType ("fun", [T, _])) = typ_of_dtyp descr' sorts T --> T'
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   308
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   309
    val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   310
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   311
        val Ts = map (typ_of_dtyp descr' sorts) cargs;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   312
        val Ts' = map mk_dummyT (filter is_rec_type cargs)
5578
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5553
diff changeset
   313
      in Const ("arbitrary", Ts @ Ts' ---> T')
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   314
      end) constrs) descr';
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   315
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   316
    val case_names = map (fn s =>
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
   317
      Sign.full_name (Theory.sign_of thy1) (s ^ "_case")) new_type_names;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   318
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   319
    (* define case combinators via primrec combinators *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   320
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   321
    val (case_defs, thy2) = foldl (fn ((defs, thy),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   322
      ((((i, (_, _, constrs)), T), name), recname)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   323
        let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   324
          val (fns1, fns2) = ListPair.unzip (map (fn ((_, cargs), j) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   325
            let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   326
              val Ts = map (typ_of_dtyp descr' sorts) cargs;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   327
              val Ts' = Ts @ map mk_dummyT (filter is_rec_type cargs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   328
              val frees' = map (uncurry (mk_Free "x")) (Ts' ~~ (1 upto length Ts'));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   329
              val frees = take (length cargs, frees');
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   330
              val free = mk_Free "f" (Ts ---> T') j
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   331
            in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   332
             (free, list_abs_free (map dest_Free frees',
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   333
               list_comb (free, frees)))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   334
            end) (constrs ~~ (1 upto length constrs)));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   335
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   336
          val caseT = (map (snd o dest_Free) fns1) @ [T] ---> T';
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   337
          val fns = (flat (take (i, case_dummy_fns))) @
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   338
            fns2 @ (flat (drop (i + 1, case_dummy_fns)));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   339
          val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   340
          val decl = (Sign.base_name name, caseT, NoSyn);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   341
          val def = ((Sign.base_name name) ^ "_def",
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   342
            Logic.mk_equals (list_comb (Const (name, caseT), fns1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   343
              list_comb (reccomb, (flat (take (i, case_dummy_fns))) @
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   344
                fns2 @ (flat (drop (i + 1, case_dummy_fns))) )));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   345
          val thy' = thy |>
7904
2b551893583e proper handling of axioms / defs;
wenzelm
parents: 7704
diff changeset
   346
            Theory.add_consts_i [decl] |> (PureThy.add_defs_i o map Thm.no_attributes) [def];
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   347
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   348
        in (defs @ [get_def thy' (Sign.base_name name)], thy')
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   349
        end) (([], thy1), (hd descr) ~~ newTs ~~ case_names ~~
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   350
          (take (length newTs, reccomb_names)));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   351
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   352
    val case_thms = map (map (fn t => prove_goalw_cterm (case_defs @
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
   353
      (map mk_meta_eq primrec_thms)) (cterm_of (Theory.sign_of thy2) t)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   354
        (fn _ => [rtac refl 1])))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   355
          (DatatypeProp.make_cases new_type_names descr sorts thy2);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   356
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   357
    val thy3 = thy2 |> Theory.add_trrules_i
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   358
      (DatatypeProp.make_case_trrules new_type_names descr) |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   359
      parent_path flat_names;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   360
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   361
  in
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   362
    (store_thmss "cases" new_type_names case_thms thy3, case_names, case_thms)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   363
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   364
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   365
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   366
(******************************* case splitting *******************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   367
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   368
fun prove_split_thms new_type_names descr sorts constr_inject dist_rewrites
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   369
    casedist_thms case_thms thy =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   370
  let
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
   371
    val _ = message "Proving equations for case splitting ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   372
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   373
    val descr' = flat descr;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   374
    val recTs = get_rec_types descr' sorts;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   375
    val newTs = take (length (hd descr), recTs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   376
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   377
    fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   378
        exhaustion), case_thms'), T) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   379
      let
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
   380
        val cert = cterm_of (Theory.sign_of thy);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   381
        val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   382
        val exhaustion' = cterm_instantiate
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   383
          [(cert lhs, cert (Free ("x", T)))] exhaustion;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   384
        val tacsf = K [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   385
          (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))]
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   386
      in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   387
        (prove_goalw_cterm [] (cert t1) tacsf,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   388
         prove_goalw_cterm [] (cert t2) tacsf)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   389
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   390
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   391
    val split_thm_pairs = map prove_split_thms
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   392
      ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   393
        dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   394
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   395
    val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   396
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   397
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   398
    (thy |> store_thms "split" new_type_names split_thms |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   399
            store_thms "split_asm" new_type_names split_asm_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   400
     split_thm_pairs)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   401
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   402
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   403
(******************************* size functions *******************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   404
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   405
fun prove_size_thms flat_names new_type_names descr sorts reccomb_names primrec_thms thy =
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   406
  if exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   407
    (fn (DtType ("fun", [_, DtRec _])) => true | _ => false) cargs) constrs) (flat descr)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   408
  then
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   409
    (thy, [])
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   410
  else
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   411
  let
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
   412
    val _ = message "Proving equations for size function ...";
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   413
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   414
    val big_name = space_implode "_" new_type_names;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   415
    val thy1 = add_path flat_names big_name thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   416
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   417
    val descr' = flat descr;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   418
    val recTs = get_rec_types descr' sorts;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   419
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   420
    val big_size_name = space_implode "_" new_type_names ^ "_size";
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   421
    val size_name = Sign.intern_const (Theory.sign_of (theory "Arith")) "size";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   422
    val size_names = replicate (length (hd descr)) size_name @
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
   423
      map (Sign.full_name (Theory.sign_of thy1))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   424
        (if length (flat (tl descr)) = 1 then [big_size_name] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   425
          map (fn i => big_size_name ^ "_" ^ string_of_int i)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   426
            (1 upto length (flat (tl descr))));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   427
    val def_names = map (fn i => big_size_name ^ "_def_" ^ string_of_int i)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   428
      (1 upto length recTs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   429
7704
9a6783fdb9a5 eliminated ap/app;
wenzelm
parents: 7228
diff changeset
   430
    fun plus (t1, t2) = Const ("op +", [HOLogic.natT, HOLogic.natT] ---> HOLogic.natT) $ t1 $ t2;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   431
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   432
    fun make_sizefun (_, cargs) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   433
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   434
        val Ts = map (typ_of_dtyp descr' sorts) cargs;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   435
        val k = length (filter is_rec_type cargs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   436
        val t = if k = 0 then HOLogic.zero else
7704
9a6783fdb9a5 eliminated ap/app;
wenzelm
parents: 7228
diff changeset
   437
          foldl1 plus (map Bound (k - 1 downto 0) @ [HOLogic.mk_nat 1])
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   438
      in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   439
        foldr (fn (T, t') => Abs ("x", T, t')) (Ts @ replicate k HOLogic.natT, t)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   440
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   441
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   442
    val fs = flat (map (fn (_, (_, _, constrs)) => map make_sizefun constrs) descr');
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   443
    val fTs = map fastype_of fs;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   444
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   445
    val thy' = thy1 |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   446
      Theory.add_consts_i (map (fn (s, T) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   447
        (Sign.base_name s, T --> HOLogic.natT, NoSyn))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   448
          (drop (length (hd descr), size_names ~~ recTs))) |>
7904
2b551893583e proper handling of axioms / defs;
wenzelm
parents: 7704
diff changeset
   449
      (PureThy.add_defs_i o map Thm.no_attributes) (map (fn (((s, T), def_name), rec_name) =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   450
        (def_name, Logic.mk_equals (Const (s, T --> HOLogic.natT),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   451
          list_comb (Const (rec_name, fTs @ [T] ---> HOLogic.natT), fs))))
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   452
            (size_names ~~ recTs ~~ def_names ~~ reccomb_names)) |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   453
      parent_path flat_names;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   454
7904
2b551893583e proper handling of axioms / defs;
wenzelm
parents: 7704
diff changeset
   455
    val size_def_thms = map (get_thm thy') def_names;
5553
ae42b36a50c2 renamed mk_meta_eq to mk_eq
oheimb
parents: 5303
diff changeset
   456
    val rewrites = size_def_thms @ map mk_meta_eq primrec_thms;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   457
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   458
    val size_thms = map (fn t => prove_goalw_cterm rewrites
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
   459
      (cterm_of (Theory.sign_of thy') t) (fn _ => [rtac refl 1]))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   460
        (DatatypeProp.make_size new_type_names descr sorts thy')
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   461
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   462
  in
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   463
    (thy' |> Theory.add_path big_name |>
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5891
diff changeset
   464
             PureThy.add_thmss [(("size", size_thms), [])] |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   465
             Theory.parent_path,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   466
     size_thms)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   467
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   468
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   469
(************************* additional theorems for TFL ************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   470
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   471
fun prove_nchotomys new_type_names descr sorts casedist_thms thy =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   472
  let
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6422
diff changeset
   473
    val _ = message "Proving additional theorems for TFL ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   474
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   475
    fun prove_nchotomy (t, exhaustion) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   476
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   477
        (* For goal i, select the correct disjunct to attack, then prove it *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   478
        fun tac i 0 = EVERY [TRY (rtac disjI1 i),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   479
              hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   480
          | tac i n = rtac disjI2 i THEN tac i (n - 1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   481
      in 
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
   482
        prove_goalw_cterm [] (cterm_of (Theory.sign_of thy) t) (fn _ =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   483
          [rtac allI 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   484
           exh_tac (K exhaustion) 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   485
           ALLGOALS (fn i => tac i (i-1))])
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   486
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   487
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   488
    val nchotomys =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   489
      map prove_nchotomy (DatatypeProp.make_nchotomys descr sorts ~~ casedist_thms)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   490
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   491
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   492
    (store_thms "nchotomy" new_type_names nchotomys thy, nchotomys)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   493
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   494
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   495
fun prove_case_congs new_type_names descr sorts nchotomys case_thms thy =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   496
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   497
    fun prove_case_cong ((t, nchotomy), case_rewrites) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   498
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   499
        val (Const ("==>", _) $ tm $ _) = t;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   500
        val (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ Ma)) = tm;
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6092
diff changeset
   501
        val cert = cterm_of (Theory.sign_of thy);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   502
        val nchotomy' = nchotomy RS spec;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   503
        val nchotomy'' = cterm_instantiate
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   504
          [(cert (hd (add_term_vars (concl_of nchotomy', []))), cert Ma)] nchotomy'
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   505
      in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   506
        prove_goalw_cterm [] (cert t) (fn prems => 
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   507
          let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   508
          in [simp_tac (HOL_ss addsimps [hd prems]) 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   509
              cut_facts_tac [nchotomy''] 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   510
              REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   511
              REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   512
          end)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   513
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   514
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   515
    val case_congs = map prove_case_cong (DatatypeProp.make_case_congs
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   516
      new_type_names descr sorts thy ~~ nchotomys ~~ case_thms)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   517
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   518
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   519
    (store_thms "case_cong" new_type_names case_congs thy, case_congs)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   520
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   521
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   522
end;