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