src/HOL/Tools/Datatype/datatype_abs_proofs.ML
author wenzelm
Thu, 09 Jun 2011 16:34:49 +0200
changeset 44206 2b47822868e4
parent 43246 774df7c59508
child 44937 ae85c5d64913
permissions -rw-r--r--
discontinued Name.variant to emphasize that this is old-style / indirect;
haftmann@33967
     1
(*  Title:      HOL/Tools/Datatype/datatype_abs_proofs.ML
wenzelm@11539
     2
    Author:     Stefan Berghofer, TU Muenchen
berghofe@5177
     3
haftmann@33967
     4
Datatype package: proofs and defintions independent of concrete
haftmann@33967
     5
representation of datatypes  (i.e. requiring only abstract
haftmann@33967
     6
properties: injectivity / distinctness of constructors and induction).
berghofe@5177
     7
*)
berghofe@5177
     8
berghofe@5177
     9
signature DATATYPE_ABS_PROOFS =
berghofe@5177
    10
sig
haftmann@31737
    11
  include DATATYPE_COMMON
haftmann@31737
    12
  val prove_casedist_thms : config -> string list ->
haftmann@31668
    13
    descr list -> (string * sort) list -> thm ->
wenzelm@18728
    14
    attribute list -> theory -> thm list * theory
haftmann@31737
    15
  val prove_primrec_thms : config -> string list ->
haftmann@31668
    16
    descr list -> (string * sort) list ->
haftmann@32915
    17
      (string -> thm list) -> thm list list -> thm list list * thm list list ->
haftmann@32915
    18
        thm -> theory -> (string list * thm list) * theory
haftmann@31737
    19
  val prove_case_thms : config -> string list ->
haftmann@31668
    20
    descr list -> (string * sort) list ->
haftmann@18314
    21
      string list -> thm list -> theory -> (thm list list * string list) * theory
haftmann@31737
    22
  val prove_split_thms : config -> string list ->
haftmann@31668
    23
    descr list -> (string * sort) list ->
berghofe@5177
    24
      thm list list -> thm list list -> thm list -> thm list list -> theory ->
haftmann@18314
    25
        (thm * thm) list * theory
haftmann@31737
    26
  val prove_nchotomys : config -> string list -> descr list ->
haftmann@18314
    27
    (string * sort) list -> thm list -> theory -> thm list * theory
haftmann@31668
    28
  val prove_weak_case_congs : string list -> descr list ->
haftmann@18314
    29
    (string * sort) list -> theory -> thm list * theory
berghofe@13641
    30
  val prove_case_congs : string list ->
haftmann@31668
    31
    descr list -> (string * sort) list ->
haftmann@18314
    32
      thm list -> thm list list -> theory -> thm list * theory
berghofe@5177
    33
end;
berghofe@5177
    34
haftmann@33967
    35
structure Datatype_Abs_Proofs: DATATYPE_ABS_PROOFS =
berghofe@5177
    36
struct
berghofe@5177
    37
berghofe@5177
    38
(************************ case distinction theorems ***************************)
berghofe@5177
    39
wenzelm@41671
    40
fun prove_casedist_thms (config : Datatype_Aux.config)
wenzelm@41671
    41
    new_type_names descr sorts induct case_names_exhausts thy =
berghofe@5177
    42
  let
wenzelm@41671
    43
    val _ = Datatype_Aux.message config "Proving case distinction theorems ...";
berghofe@5177
    44
wenzelm@32952
    45
    val descr' = flat descr;
wenzelm@41671
    46
    val recTs = Datatype_Aux.get_rec_types descr' sorts;
haftmann@33956
    47
    val newTs = take (length (hd descr)) recTs;
berghofe@5177
    48
berghofe@8477
    49
    val {maxidx, ...} = rep_thm induct;
wenzelm@8305
    50
    val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
berghofe@5177
    51
haftmann@33061
    52
    fun prove_casedist_thm (i, (T, t)) =
berghofe@5177
    53
      let
berghofe@5177
    54
        val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
wenzelm@35364
    55
          Abs ("z", T', Const (@{const_name True}, T''))) induct_Ps;
berghofe@8477
    56
        val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx+1), T), Bound 0) $
berghofe@5177
    57
          Var (("P", 0), HOLogic.boolT))
haftmann@33956
    58
        val insts = take i dummyPs @ (P::(drop (i + 1) dummyPs));
wenzelm@17985
    59
        val cert = cterm_of thy;
berghofe@5177
    60
        val insts' = (map cert induct_Ps) ~~ (map cert insts);
haftmann@32905
    61
        val induct' = refl RS ((nth
wenzelm@41671
    62
          (Datatype_Aux.split_conj_thm (cterm_instantiate insts' induct)) i) RSN (2, rev_mp))
berghofe@5177
    63
wenzelm@17985
    64
      in
wenzelm@32970
    65
        Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
wenzelm@26711
    66
          (fn {prems, ...} => EVERY
wenzelm@17985
    67
            [rtac induct' 1,
wenzelm@17985
    68
             REPEAT (rtac TrueI 1),
wenzelm@17985
    69
             REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
wenzelm@20046
    70
             REPEAT (rtac TrueI 1)])
berghofe@5177
    71
      end;
berghofe@5177
    72
haftmann@33061
    73
    val casedist_thms = map_index prove_casedist_thm
haftmann@33967
    74
      (newTs ~~ Datatype_Prop.make_casedists descr sorts)
haftmann@18314
    75
  in
haftmann@18314
    76
    thy
wenzelm@41671
    77
    |> Datatype_Aux.store_thms_atts "exhaust" new_type_names
wenzelm@41671
    78
      (map single case_names_exhausts) casedist_thms
haftmann@18314
    79
  end;
berghofe@5177
    80
berghofe@5177
    81
berghofe@5177
    82
(*************************** primrec combinators ******************************)
berghofe@5177
    83
wenzelm@41671
    84
fun prove_primrec_thms (config : Datatype_Aux.config) new_type_names descr sorts
haftmann@32915
    85
    injects_of constr_inject (dist_rewrites, other_dist_rewrites) induct thy =
berghofe@5177
    86
  let
wenzelm@41671
    87
    val _ = Datatype_Aux.message config "Constructing primrec combinators ...";
berghofe@5661
    88
berghofe@5661
    89
    val big_name = space_implode "_" new_type_names;
haftmann@32102
    90
    val thy0 = Sign.add_path big_name thy;
berghofe@5177
    91
wenzelm@32952
    92
    val descr' = flat descr;
wenzelm@41671
    93
    val recTs = Datatype_Aux.get_rec_types descr' sorts;
wenzelm@30193
    94
    val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
haftmann@33956
    95
    val newTs = take (length (hd descr)) recTs;
berghofe@5177
    96
wenzelm@8305
    97
    val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
berghofe@5177
    98
berghofe@5661
    99
    val big_rec_name' = big_name ^ "_rec_set";
berghofe@21021
   100
    val rec_set_names' =
berghofe@21021
   101
      if length descr' = 1 then [big_rec_name'] else
berghofe@21021
   102
        map ((curry (op ^) (big_rec_name' ^ "_")) o string_of_int)
berghofe@21021
   103
          (1 upto (length descr'));
haftmann@28965
   104
    val rec_set_names = map (Sign.full_bname thy0) rec_set_names';
berghofe@5177
   105
haftmann@33967
   106
    val (rec_result_Ts, reccomb_fn_Ts) = Datatype_Prop.make_primrec_Ts descr sorts used;
berghofe@5177
   107
berghofe@21021
   108
    val rec_set_Ts = map (fn (T1, T2) =>
berghofe@21021
   109
      reccomb_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
berghofe@5177
   110
wenzelm@41671
   111
    val rec_fns = map (uncurry (Datatype_Aux.mk_Free "f"))
berghofe@5177
   112
      (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
berghofe@21021
   113
    val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
berghofe@21021
   114
      (rec_set_names' ~~ rec_set_Ts);
berghofe@5177
   115
    val rec_sets = map (fn c => list_comb (Const c, rec_fns))
berghofe@5177
   116
      (rec_set_names ~~ rec_set_Ts);
berghofe@5177
   117
berghofe@5177
   118
    (* introduction rules for graph of primrec function *)
berghofe@5177
   119
haftmann@32906
   120
    fun make_rec_intr T rec_set (cname, cargs) (rec_intr_ts, l) =
berghofe@5177
   121
      let
wenzelm@33345
   122
        fun mk_prem (dt, U) (j, k, prems, t1s, t2s) =
wenzelm@41671
   123
          let val free1 = Datatype_Aux.mk_Free "x" U j
wenzelm@41671
   124
          in (case (Datatype_Aux.strip_dtyp dt, strip_type U) of
wenzelm@41671
   125
             ((_, Datatype_Aux.DtRec m), (Us, _)) =>
berghofe@13641
   126
               let
wenzelm@41671
   127
                 val free2 = Datatype_Aux.mk_Free "y" (Us ---> nth rec_result_Ts m) k;
berghofe@13641
   128
                 val i = length Us
berghofe@13641
   129
               in (j + 1, k + 1, HOLogic.mk_Trueprop (HOLogic.list_all
haftmann@32905
   130
                     (map (pair "x") Us, nth rec_sets' m $
wenzelm@41671
   131
                       Datatype_Aux.app_bnds free1 i $ Datatype_Aux.app_bnds free2 i)) :: prems,
berghofe@5177
   132
                   free1::t1s, free2::t2s)
berghofe@5177
   133
               end
berghofe@5177
   134
           | _ => (j + 1, k, prems, free1::t1s, t2s))
berghofe@5177
   135
          end;
berghofe@5177
   136
wenzelm@41671
   137
        val Ts = map (Datatype_Aux.typ_of_dtyp descr' sorts) cargs;
wenzelm@33345
   138
        val (_, _, prems, t1s, t2s) = fold_rev mk_prem (cargs ~~ Ts) (1, 1, [], [], [])
berghofe@5177
   139
berghofe@21021
   140
      in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop
berghofe@21021
   141
        (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
haftmann@32905
   142
          list_comb (nth rec_fns l, t1s @ t2s)))], l + 1)
berghofe@5177
   143
      end;
berghofe@5177
   144
haftmann@32906
   145
    val (rec_intr_ts, _) = fold (fn ((d, T), set_name) =>
haftmann@32906
   146
      fold (make_rec_intr T set_name) (#3 (snd d)))
haftmann@32906
   147
        (descr' ~~ recTs ~~ rec_sets') ([], 0);
berghofe@5177
   148
wenzelm@21365
   149
    val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
wenzelm@33278
   150
      thy0
wenzelm@33278
   151
      |> Sign.map_naming Name_Space.conceal
wenzelm@33736
   152
      |> Inductive.add_inductive_global
wenzelm@33671
   153
          {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name',
wenzelm@33671
   154
            coind = false, no_elim = false, no_ind = true, skip_mono = true, fork_mono = false}
haftmann@28965
   155
          (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
wenzelm@26128
   156
          (map dest_Free rec_fns)
wenzelm@33278
   157
          (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) []
wenzelm@33278
   158
      ||> Sign.restore_naming thy0
wenzelm@33278
   159
      ||> Theory.checkpoint;
berghofe@5177
   160
berghofe@5177
   161
    (* prove uniqueness and termination of primrec combinators *)
berghofe@5177
   162
wenzelm@41671
   163
    val _ = Datatype_Aux.message config "Proving termination and uniqueness of primrec functions ...";
berghofe@5177
   164
haftmann@32906
   165
    fun mk_unique_tac ((((i, (tname, _, constrs)), elim), T), T') (tac, intrs) =
berghofe@5177
   166
      let
berghofe@21021
   167
        val distinct_tac =
berghofe@5177
   168
          (if i < length newTs then
haftmann@32905
   169
             full_simp_tac (HOL_ss addsimps (nth dist_rewrites i)) 1
haftmann@32915
   170
           else full_simp_tac (HOL_ss addsimps (flat other_dist_rewrites)) 1);
berghofe@5177
   171
berghofe@5177
   172
        val inject = map (fn r => r RS iffD1)
haftmann@32905
   173
          (if i < length newTs then nth constr_inject i
haftmann@32726
   174
            else injects_of tname);
berghofe@5177
   175
haftmann@32906
   176
        fun mk_unique_constr_tac n (cname, cargs) (tac, intr::intrs, j) =
berghofe@5177
   177
          let
wenzelm@41671
   178
            val k = length (filter Datatype_Aux.is_rec_type cargs)
berghofe@5177
   179
berghofe@5177
   180
          in (EVERY [DETERM tac,
berghofe@5177
   181
                REPEAT (etac ex1E 1), rtac ex1I 1,
berghofe@5177
   182
                DEPTH_SOLVE_1 (ares_tac [intr] 1),
berghofe@13641
   183
                REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1),
berghofe@5177
   184
                etac elim 1,
berghofe@5177
   185
                REPEAT_DETERM_N j distinct_tac,
berghofe@21021
   186
                TRY (dresolve_tac inject 1),
berghofe@5177
   187
                REPEAT (etac conjE 1), hyp_subst_tac 1,
berghofe@13641
   188
                REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
berghofe@5177
   189
                TRY (hyp_subst_tac 1),
berghofe@5177
   190
                rtac refl 1,
berghofe@5177
   191
                REPEAT_DETERM_N (n - j - 1) distinct_tac],
berghofe@5177
   192
              intrs, j + 1)
berghofe@5177
   193
          end;
berghofe@5177
   194
haftmann@32906
   195
        val (tac', intrs', _) = fold (mk_unique_constr_tac (length constrs))
haftmann@32906
   196
          constrs (tac, intrs, 0);
berghofe@5177
   197
berghofe@5177
   198
      in (tac', intrs') end;
berghofe@5177
   199
berghofe@5177
   200
    val rec_unique_thms =
berghofe@5177
   201
      let
berghofe@5177
   202
        val rec_unique_ts = map (fn (((set_t, T1), T2), i) =>
wenzelm@35364
   203
          Const (@{const_name Ex1}, (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
wenzelm@41671
   204
            absfree ("y", T2, set_t $ Datatype_Aux.mk_Free "x" T1 i $ Free ("y", T2)))
berghofe@21021
   205
              (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
wenzelm@17985
   206
        val cert = cterm_of thy1
berghofe@5177
   207
        val insts = map (fn ((i, T), t) => absfree ("x" ^ (string_of_int i), T, t))
berghofe@5177
   208
          ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
berghofe@5177
   209
        val induct' = cterm_instantiate ((map cert induct_Ps) ~~
berghofe@5177
   210
          (map cert insts)) induct;
haftmann@32906
   211
        val (tac, _) = fold mk_unique_tac (descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts)
wenzelm@35625
   212
           (((rtac induct' THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1
wenzelm@35410
   213
              THEN rewrite_goals_tac [mk_meta_eq @{thm choice_eq}], rec_intrs));
berghofe@5177
   214
wenzelm@41671
   215
      in
wenzelm@41671
   216
        Datatype_Aux.split_conj_thm (Skip_Proof.prove_global thy1 [] []
wenzelm@41671
   217
          (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj rec_unique_ts)) (K tac))
berghofe@5177
   218
      end;
berghofe@5177
   219
wenzelm@35410
   220
    val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms;
berghofe@5177
   221
berghofe@5177
   222
    (* define primrec combinators *)
berghofe@5177
   223
berghofe@5177
   224
    val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
haftmann@28965
   225
    val reccomb_names = map (Sign.full_bname thy1)
berghofe@5177
   226
      (if length descr' = 1 then [big_reccomb_name] else
berghofe@5177
   227
        (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
berghofe@5177
   228
          (1 upto (length descr'))));
berghofe@5177
   229
    val reccombs = map (fn ((name, T), T') => list_comb
berghofe@5177
   230
      (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
berghofe@5177
   231
        (reccomb_names ~~ recTs ~~ rec_result_Ts);
berghofe@5177
   232
haftmann@18358
   233
    val (reccomb_defs, thy2) =
haftmann@18358
   234
      thy1
wenzelm@24712
   235
      |> Sign.add_consts_i (map (fn ((name, T), T') =>
wenzelm@30364
   236
          (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn))
haftmann@18358
   237
          (reccomb_names ~~ recTs ~~ rec_result_Ts))
wenzelm@39814
   238
      |> (Global_Theory.add_defs false o map Thm.no_attributes)
wenzelm@39814
   239
          (map (fn ((((name, comb), set), T), T') =>
wenzelm@39814
   240
            (Binding.name (Long_Name.base_name name ^ "_def"),
wenzelm@39814
   241
              Logic.mk_equals (comb, absfree ("x", T,
wenzelm@39814
   242
               Const (@{const_name The}, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
wenzelm@39814
   243
                 set $ Free ("x", T) $ Free ("y", T'))))))
wenzelm@39814
   244
                   (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
haftmann@32102
   245
      ||> Sign.parent_path
wenzelm@28361
   246
      ||> Theory.checkpoint;
berghofe@5177
   247
berghofe@5177
   248
berghofe@5177
   249
    (* prove characteristic equations for primrec combinators *)
berghofe@5177
   250
wenzelm@41671
   251
    val _ = Datatype_Aux.message config "Proving characteristic theorems for primrec combinators ...";
berghofe@5177
   252
wenzelm@32970
   253
    val rec_thms = map (fn t => Skip_Proof.prove_global thy2 [] [] t
wenzelm@17985
   254
      (fn _ => EVERY
wenzelm@17985
   255
        [rewrite_goals_tac reccomb_defs,
wenzelm@35410
   256
         rtac @{thm the1_equality} 1,
berghofe@5177
   257
         resolve_tac rec_unique_thms 1,
berghofe@5177
   258
         resolve_tac rec_intrs 1,
wenzelm@20046
   259
         REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
wenzelm@41671
   260
           (Datatype_Prop.make_primrecs new_type_names descr sorts thy2);
berghofe@5177
   261
  in
haftmann@18314
   262
    thy2
wenzelm@24712
   263
    |> Sign.add_path (space_implode "_" new_type_names)
wenzelm@39814
   264
    |> Global_Theory.add_thmss [((Binding.name "recs", rec_thms), [Nitpick_Simps.add])]
wenzelm@24712
   265
    ||> Sign.parent_path
wenzelm@28361
   266
    ||> Theory.checkpoint
haftmann@32906
   267
    |-> (fn thms => pair (reccomb_names, flat thms))
berghofe@5177
   268
  end;
berghofe@5177
   269
berghofe@8477
   270
berghofe@5177
   271
(***************************** case combinators *******************************)
berghofe@5177
   272
wenzelm@41671
   273
fun prove_case_thms (config : Datatype_Aux.config)
wenzelm@41671
   274
    new_type_names descr sorts reccomb_names primrec_thms thy =
berghofe@5177
   275
  let
wenzelm@41671
   276
    val _ = Datatype_Aux.message config "Proving characteristic theorems for case combinators ...";
berghofe@5661
   277
haftmann@32102
   278
    val thy1 = Sign.add_path (space_implode "_" new_type_names) thy;
berghofe@5177
   279
wenzelm@32952
   280
    val descr' = flat descr;
wenzelm@41671
   281
    val recTs = Datatype_Aux.get_rec_types descr' sorts;
wenzelm@30193
   282
    val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
haftmann@33956
   283
    val newTs = take (length (hd descr)) recTs;
wenzelm@44206
   284
    val T' = TFree (singleton (Name.variant_list used) "'t", HOLogic.typeS);
berghofe@5177
   285
wenzelm@41671
   286
    fun mk_dummyT dt = binder_types (Datatype_Aux.typ_of_dtyp descr' sorts dt) ---> T';
berghofe@7015
   287
berghofe@5177
   288
    val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
berghofe@5177
   289
      let
wenzelm@41671
   290
        val Ts = map (Datatype_Aux.typ_of_dtyp descr' sorts) cargs;
wenzelm@41671
   291
        val Ts' = map mk_dummyT (filter Datatype_Aux.is_rec_type cargs)
haftmann@28524
   292
      in Const (@{const_name undefined}, Ts @ Ts' ---> T')
berghofe@5177
   293
      end) constrs) descr';
berghofe@5177
   294
haftmann@28965
   295
    val case_names = map (fn s => Sign.full_bname thy1 (s ^ "_case")) new_type_names;
berghofe@5177
   296
berghofe@5177
   297
    (* define case combinators via primrec combinators *)
berghofe@5177
   298
haftmann@32906
   299
    val (case_defs, thy2) = fold (fn ((((i, (_, _, constrs)), T), name), recname) => fn (defs, thy) =>
berghofe@5177
   300
        let
haftmann@32906
   301
          val (fns1, fns2) = split_list (map (fn ((_, cargs), j) =>
berghofe@5177
   302
            let
wenzelm@41671
   303
              val Ts = map (Datatype_Aux.typ_of_dtyp descr' sorts) cargs;
wenzelm@41671
   304
              val Ts' = Ts @ map mk_dummyT (filter Datatype_Aux.is_rec_type cargs);
wenzelm@41671
   305
              val frees' = map2 (Datatype_Aux.mk_Free "x") Ts' (1 upto length Ts');
haftmann@33956
   306
              val frees = take (length cargs) frees';
wenzelm@41671
   307
              val free = Datatype_Aux.mk_Free "f" (Ts ---> T') j
berghofe@5177
   308
            in
berghofe@5177
   309
             (free, list_abs_free (map dest_Free frees',
berghofe@5177
   310
               list_comb (free, frees)))
berghofe@5177
   311
            end) (constrs ~~ (1 upto length constrs)));
berghofe@5177
   312
berghofe@5177
   313
          val caseT = (map (snd o dest_Free) fns1) @ [T] ---> T';
haftmann@33956
   314
          val fns = flat (take i case_dummy_fns) @
haftmann@33956
   315
            fns2 @ flat (drop (i + 1) case_dummy_fns);
berghofe@5177
   316
          val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
wenzelm@30364
   317
          val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn);
wenzelm@30364
   318
          val def = (Binding.name (Long_Name.base_name name ^ "_def"),
berghofe@5177
   319
            Logic.mk_equals (list_comb (Const (name, caseT), fns1),
haftmann@33956
   320
              list_comb (reccomb, (flat (take i case_dummy_fns)) @
haftmann@33956
   321
                fns2 @ (flat (drop (i + 1) case_dummy_fns)))));
haftmann@18358
   322
          val ([def_thm], thy') =
haftmann@18358
   323
            thy
wenzelm@43246
   324
            |> Sign.declare_const_global decl |> snd
wenzelm@39814
   325
            |> (Global_Theory.add_defs false o map Thm.no_attributes) [def];
berghofe@5177
   326
wenzelm@8436
   327
        in (defs @ [def_thm], thy')
haftmann@32906
   328
        end) (hd descr ~~ newTs ~~ case_names ~~
haftmann@33956
   329
          take (length newTs) reccomb_names) ([], thy1)
wenzelm@28361
   330
      ||> Theory.checkpoint;
berghofe@5177
   331
wenzelm@32970
   332
    val case_thms = map (map (fn t => Skip_Proof.prove_global thy2 [] [] t
wenzelm@20046
   333
      (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1])))
haftmann@33967
   334
          (Datatype_Prop.make_cases new_type_names descr sorts thy2)
berghofe@8477
   335
  in
haftmann@18314
   336
    thy2
wenzelm@33465
   337
    |> Context.theory_map ((fold o fold) Nitpick_Simps.add_thm case_thms)
haftmann@32102
   338
    |> Sign.parent_path
wenzelm@41671
   339
    |> Datatype_Aux.store_thmss "cases" new_type_names case_thms
haftmann@18314
   340
    |-> (fn thmss => pair (thmss, case_names))
berghofe@8477
   341
  end;
berghofe@5177
   342
berghofe@5177
   343
berghofe@5177
   344
(******************************* case splitting *******************************)
berghofe@5177
   345
wenzelm@41671
   346
fun prove_split_thms (config : Datatype_Aux.config)
wenzelm@41671
   347
    new_type_names descr sorts constr_inject dist_rewrites casedist_thms case_thms thy =
berghofe@5177
   348
  let
wenzelm@41671
   349
    val _ = Datatype_Aux.message config "Proving equations for case splitting ...";
berghofe@5177
   350
haftmann@31668
   351
    val descr' = flat descr;
wenzelm@41671
   352
    val recTs = Datatype_Aux.get_rec_types descr' sorts;
haftmann@33956
   353
    val newTs = take (length (hd descr)) recTs;
berghofe@5177
   354
berghofe@5177
   355
    fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'),
berghofe@5177
   356
        exhaustion), case_thms'), T) =
berghofe@5177
   357
      let
wenzelm@17985
   358
        val cert = cterm_of thy;
berghofe@5177
   359
        val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
berghofe@5177
   360
        val exhaustion' = cterm_instantiate
berghofe@5177
   361
          [(cert lhs, cert (Free ("x", T)))] exhaustion;
wenzelm@17985
   362
        val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
wenzelm@17985
   363
          (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))])
berghofe@5177
   364
      in
wenzelm@32970
   365
        (Skip_Proof.prove_global thy [] [] t1 tacf,
wenzelm@32970
   366
         Skip_Proof.prove_global thy [] [] t2 tacf)
berghofe@5177
   367
      end;
berghofe@5177
   368
berghofe@5177
   369
    val split_thm_pairs = map prove_split_thms
haftmann@33967
   370
      ((Datatype_Prop.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
berghofe@5177
   371
        dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
berghofe@5177
   372
berghofe@5177
   373
    val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs
berghofe@5177
   374
berghofe@5177
   375
  in
haftmann@18314
   376
    thy
wenzelm@41671
   377
    |> Datatype_Aux.store_thms "split" new_type_names split_thms
wenzelm@41671
   378
    ||>> Datatype_Aux.store_thms "split_asm" new_type_names split_asm_thms
haftmann@18314
   379
    |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
berghofe@5177
   380
  end;
berghofe@5177
   381
nipkow@8601
   382
fun prove_weak_case_congs new_type_names descr sorts thy =
nipkow@8601
   383
  let
nipkow@8601
   384
    fun prove_weak_case_cong t =
wenzelm@32970
   385
       Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
wenzelm@26711
   386
         (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1])
nipkow@8601
   387
haftmann@33967
   388
    val weak_case_congs = map prove_weak_case_cong (Datatype_Prop.make_weak_case_congs
nipkow@8601
   389
      new_type_names descr sorts thy)
nipkow@8601
   390
wenzelm@41671
   391
  in thy |> Datatype_Aux.store_thms "weak_case_cong" new_type_names weak_case_congs end;
berghofe@8477
   392
berghofe@5177
   393
(************************* additional theorems for TFL ************************)
berghofe@5177
   394
wenzelm@41671
   395
fun prove_nchotomys (config : Datatype_Aux.config)
wenzelm@41671
   396
    new_type_names descr sorts casedist_thms thy =
berghofe@5177
   397
  let
wenzelm@41671
   398
    val _ = Datatype_Aux.message config "Proving additional theorems for TFL ...";
berghofe@5177
   399
berghofe@5177
   400
    fun prove_nchotomy (t, exhaustion) =
berghofe@5177
   401
      let
berghofe@5177
   402
        (* For goal i, select the correct disjunct to attack, then prove it *)
berghofe@5177
   403
        fun tac i 0 = EVERY [TRY (rtac disjI1 i),
berghofe@5177
   404
              hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
berghofe@5177
   405
          | tac i n = rtac disjI2 i THEN tac i (n - 1)
berghofe@5177
   406
      in 
wenzelm@32970
   407
        Skip_Proof.prove_global thy [] [] t (fn _ =>
wenzelm@17985
   408
          EVERY [rtac allI 1,
wenzelm@41671
   409
           Datatype_Aux.exh_tac (K exhaustion) 1,
wenzelm@20046
   410
           ALLGOALS (fn i => tac i (i-1))])
berghofe@5177
   411
      end;
berghofe@5177
   412
berghofe@5177
   413
    val nchotomys =
haftmann@33967
   414
      map prove_nchotomy (Datatype_Prop.make_nchotomys descr sorts ~~ casedist_thms)
berghofe@5177
   415
wenzelm@41671
   416
  in thy |> Datatype_Aux.store_thms "nchotomy" new_type_names nchotomys end;
berghofe@5177
   417
berghofe@5177
   418
fun prove_case_congs new_type_names descr sorts nchotomys case_thms thy =
berghofe@5177
   419
  let
berghofe@5177
   420
    fun prove_case_cong ((t, nchotomy), case_rewrites) =
berghofe@5177
   421
      let
berghofe@5177
   422
        val (Const ("==>", _) $ tm $ _) = t;
haftmann@39093
   423
        val (Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ Ma)) = tm;
wenzelm@22578
   424
        val cert = cterm_of thy;
berghofe@5177
   425
        val nchotomy' = nchotomy RS spec;
wenzelm@29264
   426
        val [v] = Term.add_vars (concl_of nchotomy') [];
wenzelm@29264
   427
        val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy'
berghofe@5177
   428
      in
wenzelm@32970
   429
        Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
wenzelm@26711
   430
          (fn {prems, ...} => 
wenzelm@17985
   431
            let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
wenzelm@17985
   432
            in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1,
wenzelm@17985
   433
                cut_facts_tac [nchotomy''] 1,
wenzelm@17985
   434
                REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
wenzelm@17985
   435
                REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
wenzelm@20046
   436
            end)
berghofe@5177
   437
      end;
berghofe@5177
   438
haftmann@33967
   439
    val case_congs = map prove_case_cong (Datatype_Prop.make_case_congs
berghofe@5177
   440
      new_type_names descr sorts thy ~~ nchotomys ~~ case_thms)
berghofe@5177
   441
wenzelm@41671
   442
  in thy |> Datatype_Aux.store_thms "case_cong" new_type_names case_congs end;
wenzelm@41671
   443
wenzelm@41671
   444
wenzelm@41671
   445
open Datatype_Aux;
berghofe@5177
   446
berghofe@5177
   447
end;