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