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