src/HOL/Tools/Datatype/datatype_rep_proofs.ML
author wenzelm
Thu, 12 Nov 2009 22:02:11 +0100
changeset 33643 b275f26a638b
parent 33338 de76079f973a
child 33669 ae9a2ea9a989
permissions -rw-r--r--
eliminated obsolete "internal" kind -- collapsed to unspecific "";
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_rep_proofs.ML
11539
0f17da240450 tuned headers;
wenzelm
parents: 11471
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     3
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     4
Definitional introduction of datatypes
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     5
Proof of characteristic theorems:
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     6
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     7
 - injectivity of constructors
7228
ddb67dcf026c Tuned some comments.
berghofe
parents: 7205
diff changeset
     8
 - distinctness of constructors
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     9
 - induction theorem
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    10
*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    11
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    12
signature DATATYPE_REP_PROOFS =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    13
sig
31737
b3f63611784e simplified names of common datatype types
haftmann
parents: 31723
diff changeset
    14
  include DATATYPE_COMMON
b3f63611784e simplified names of common datatype types
haftmann
parents: 31723
diff changeset
    15
  val representation_proofs : config -> info Symtab.table ->
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31643
diff changeset
    16
    string list -> descr list -> (string * sort) list ->
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30307
diff changeset
    17
      (binding * mixfix) list -> (binding * mixfix) list list -> attribute
32907
0300f8dd63ea dropped rule duplicates
haftmann
parents: 32904
diff changeset
    18
        -> theory -> (thm list list * thm list list * thm) * theory
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    19
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    20
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    21
structure DatatypeRepProofs : DATATYPE_REP_PROOFS =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    22
struct
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    23
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    24
open DatatypeAux;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    25
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    26
val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    27
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
    28
val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
    29
11435
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
    30
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
    31
(** theory context references **)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    32
31737
b3f63611784e simplified names of common datatype types
haftmann
parents: 31723
diff changeset
    33
fun exh_thm_of (dt_info : info Symtab.table) tname =
32712
ec5976f4d3d8 registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
haftmann
parents: 32124
diff changeset
    34
  #exhaust (the (Symtab.lookup dt_info tname));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    35
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    36
(******************************************************************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    37
31737
b3f63611784e simplified names of common datatype types
haftmann
parents: 31723
diff changeset
    38
fun representation_proofs (config : config) (dt_info : info Symtab.table)
8436
8a87fa482baf adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8305
diff changeset
    39
      new_type_names descr sorts types_syntax constr_syntax case_names_induct thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    40
  let
19540
d036bff01c23 ThyInfo.the_theory;
wenzelm
parents: 19346
diff changeset
    41
    val Datatype_thy = ThyInfo.the_theory "Datatype" thy;
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20071
diff changeset
    42
    val node_name = "Datatype.node";
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20071
diff changeset
    43
    val In0_name = "Datatype.In0";
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20071
diff changeset
    44
    val In1_name = "Datatype.In1";
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20071
diff changeset
    45
    val Scons_name = "Datatype.Scons";
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20071
diff changeset
    46
    val Leaf_name = "Datatype.Leaf";
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20071
diff changeset
    47
    val Numb_name = "Datatype.Numb";
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20071
diff changeset
    48
    val Lim_name = "Datatype.Lim";
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
    49
    val Suml_name = "Datatype.Suml";
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
    50
    val Sumr_name = "Datatype.Sumr";
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    51
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
    52
    val [In0_inject, In1_inject, Scons_inject, Leaf_inject,
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
    53
         In0_eq, In1_eq, In0_not_In1, In1_not_In0,
26343
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26336
diff changeset
    54
         Lim_inject, Suml_inject, Sumr_inject] = map (PureThy.get_thm Datatype_thy)
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26336
diff changeset
    55
          ["In0_inject", "In1_inject", "Scons_inject", "Leaf_inject",
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26336
diff changeset
    56
           "In0_eq", "In1_eq", "In0_not_In1", "In1_not_In0",
0dd2eab7b296 simplified get_thm(s): back to plain name argument;
wenzelm
parents: 26336
diff changeset
    57
           "Lim_inject", "Suml_inject", "Sumr_inject"];
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    58
28362
b0ebac91c8d5 explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
wenzelm
parents: 28084
diff changeset
    59
    val descr' = flat descr;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    60
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    61
    val big_name = space_implode "_" new_type_names;
32124
954321008785 dropped ancient flat_names option
haftmann
parents: 31986
diff changeset
    62
    val thy1 = Sign.add_path big_name thy;
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    63
    val big_rec_name = big_name ^ "_rep_set";
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
    64
    val rep_set_names' =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    65
      (if length descr' = 1 then [big_rec_name] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    66
        (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    67
          (1 upto (length descr'))));
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28662
diff changeset
    68
    val rep_set_names = map (Sign.full_bname thy1) rep_set_names';
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    69
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    70
    val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5553
diff changeset
    71
    val leafTs' = get_nonrec_types descr' sorts;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    72
    val branchTs = get_branching_types descr' sorts;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    73
    val branchT = if null branchTs then HOLogic.unitT
32765
3032c0308019 modernized Balanced_Tree;
wenzelm
parents: 32727
diff changeset
    74
      else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) branchTs;
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 32999
diff changeset
    75
    val arities = remove (op =) 0 (get_arities descr');
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
    76
    val unneeded_vars =
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
    77
      subtract (op =) (List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs)) (hd tyvars);
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
    78
    val leafTs = leafTs' @ map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    79
    val recTs = get_rec_types descr' sorts;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    80
    val newTs = Library.take (length (hd descr), recTs);
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    81
    val oldTs = Library.drop (length (hd descr), recTs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    82
    val sumT = if null leafTs then HOLogic.unitT
32765
3032c0308019 modernized Balanced_Tree;
wenzelm
parents: 32727
diff changeset
    83
      else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) leafTs;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    84
    val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT]));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    85
    val UnivT = HOLogic.mk_setT Univ_elT;
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
    86
    val UnivT' = Univ_elT --> HOLogic.boolT;
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
    87
    val Collect = Const ("Collect", UnivT' --> UnivT);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    88
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    89
    val In0 = Const (In0_name, Univ_elT --> Univ_elT);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    90
    val In1 = Const (In1_name, Univ_elT --> Univ_elT);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    91
    val Leaf = Const (Leaf_name, sumT --> Univ_elT);
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
    92
    val Lim = Const (Lim_name, (branchT --> Univ_elT) --> Univ_elT);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    93
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    94
    (* make injections needed for embedding types in leaves *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    95
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    96
    fun mk_inj T' x =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    97
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    98
        fun mk_inj' T n i =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    99
          if n = 1 then x else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   100
          let val n2 = n div 2;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   101
              val Type (_, [T1, T2]) = T
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   102
          in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   103
            if i <= n2 then
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 14981
diff changeset
   104
              Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   105
            else
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 14981
diff changeset
   106
              Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   107
          end
31986
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31949
diff changeset
   108
      in mk_inj' sumT (length leafTs) (1 + find_index (fn T'' => T'' = T') leafTs)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   109
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   110
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   111
    (* make injections for constructors *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   112
32765
3032c0308019 modernized Balanced_Tree;
wenzelm
parents: 32727
diff changeset
   113
    fun mk_univ_inj ts = Balanced_Tree.access
23419
8c30dd4b3b22 BalancedTree;
wenzelm
parents: 22596
diff changeset
   114
      {left = fn t => In0 $ t,
8c30dd4b3b22 BalancedTree;
wenzelm
parents: 22596
diff changeset
   115
        right = fn t => In1 $ t,
8c30dd4b3b22 BalancedTree;
wenzelm
parents: 22596
diff changeset
   116
        init =
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 28362
diff changeset
   117
          if ts = [] then Const (@{const_name undefined}, Univ_elT)
23419
8c30dd4b3b22 BalancedTree;
wenzelm
parents: 22596
diff changeset
   118
          else foldr1 (HOLogic.mk_binop Scons_name) ts};
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   119
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   120
    (* function spaces *)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   121
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   122
    fun mk_fun_inj T' x =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   123
      let
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   124
        fun mk_inj T n i =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   125
          if n = 1 then x else
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   126
          let
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   127
            val n2 = n div 2;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   128
            val Type (_, [T1, T2]) = T;
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   129
            fun mkT U = (U --> Univ_elT) --> T --> Univ_elT
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   130
          in
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   131
            if i <= n2 then Const (Suml_name, mkT T1) $ mk_inj T1 n2 i
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   132
            else Const (Sumr_name, mkT T2) $ mk_inj T2 (n - n2) (i - n2)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   133
          end
31986
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31949
diff changeset
   134
      in mk_inj branchT (length branchTs) (1 + find_index (fn T'' => T'' = T') branchTs)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   135
      end;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   136
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
   137
    fun mk_lim t Ts = fold_rev (fn T => fn t => Lim $ mk_fun_inj T (Abs ("x", T, t))) Ts t;
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   138
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   139
    (************** generate introduction rules for representing set **********)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   140
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31643
diff changeset
   141
    val _ = message config "Constructing representing sets ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   142
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   143
    (* make introduction rule for a single constructor *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   144
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   145
    fun make_intr s n (i, (_, cargs)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   146
      let
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
   147
        fun mk_prem dt (j, prems, ts) =
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
   148
          (case strip_dtyp dt of
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   149
            (dts, DtRec k) =>
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   150
              let
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   151
                val Ts = map (typ_of_dtyp descr' sorts) dts;
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   152
                val free_t =
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   153
                  app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts)
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   154
              in (j + 1, list_all (map (pair "x") Ts,
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   155
                  HOLogic.mk_Trueprop
31949
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   156
                    (Free (nth rep_set_names' k, UnivT') $ free_t)) :: prems,
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   157
                mk_lim free_t Ts :: ts)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   158
              end
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   159
          | _ =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   160
              let val T = typ_of_dtyp descr' sorts dt
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   161
              in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts)
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   162
              end);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   163
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
   164
        val (_, prems, ts) = fold_rev mk_prem cargs (1, [], []);
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   165
        val concl = HOLogic.mk_Trueprop
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   166
          (Free (s, UnivT') $ mk_univ_inj ts n i)
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   167
      in Logic.list_implies (prems, concl)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   168
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   169
28362
b0ebac91c8d5 explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
wenzelm
parents: 28084
diff changeset
   170
    val intr_ts = maps (fn ((_, (_, _, constrs)), rep_set_name) =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   171
      map (make_intr rep_set_name (length constrs))
28362
b0ebac91c8d5 explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
wenzelm
parents: 28084
diff changeset
   172
        ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names');
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   173
21365
4ee8e2702241 InductivePackage.add_inductive_i: canonical argument order;
wenzelm
parents: 21291
diff changeset
   174
    val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33244
diff changeset
   175
      thy1
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33244
diff changeset
   176
      |> Sign.map_naming Name_Space.conceal
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33244
diff changeset
   177
      |> Inductive.add_inductive_global (serial ())
33643
b275f26a638b eliminated obsolete "internal" kind -- collapsed to unspecific "";
wenzelm
parents: 33338
diff changeset
   178
          {quiet_mode = #quiet config, verbose = false, kind = "",
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28662
diff changeset
   179
           alt_name = Binding.name big_rec_name, coind = false, no_elim = true, no_ind = false,
29389
0a49f940d729 inductive: added fork_mono flag;
wenzelm
parents: 29270
diff changeset
   180
           skip_mono = true, fork_mono = false}
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28662
diff changeset
   181
          (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33244
diff changeset
   182
          (map (fn x => (Attrib.empty_binding, x)) intr_ts) []
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33244
diff changeset
   183
      ||> Sign.restore_naming thy1
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33244
diff changeset
   184
      ||> Theory.checkpoint;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   185
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   186
    (********************************* typedef ********************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   187
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   188
    val (typedefs, thy3) = thy2 |>
32124
954321008785 dropped ancient flat_names option
haftmann
parents: 31986
diff changeset
   189
      Sign.parent_path |>
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   190
      fold_map (fn ((((name, mx), tvs), c), name') =>
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31668
diff changeset
   191
          Typedef.add_typedef false (SOME (Binding.name name')) (name, tvs, mx)
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   192
            (Collect $ Const (c, UnivT')) NONE
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   193
            (rtac exI 1 THEN rtac CollectI 1 THEN
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   194
              QUIET_BREADTH_FIRST (has_fewer_prems 1)
26475
3cc1e48d0ce1 eliminated quiet_mode ref of some packages (avoid CRITICAL setmp!);
wenzelm
parents: 26359
diff changeset
   195
              (resolve_tac rep_intrs 1)))
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   196
                (types_syntax ~~ tyvars ~~
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   197
                  (Library.take (length newTs, rep_set_names)) ~~ new_type_names) ||>
32124
954321008785 dropped ancient flat_names option
haftmann
parents: 31986
diff changeset
   198
      Sign.add_path big_name;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   199
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   200
    (*********************** definition of constructors ***********************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   201
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   202
    val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   203
    val rep_names = map (curry op ^ "Rep_") new_type_names;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   204
    val rep_names' = map (fn i => big_rep_name ^ (string_of_int i))
28362
b0ebac91c8d5 explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
wenzelm
parents: 28084
diff changeset
   205
      (1 upto (length (flat (tl descr))));
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21525
diff changeset
   206
    val all_rep_names = map (Sign.intern_const thy3) rep_names @
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28662
diff changeset
   207
      map (Sign.full_bname thy3) rep_names';
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   208
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   209
    (* isomorphism declarations *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   210
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30307
diff changeset
   211
    val iso_decls = map (fn (T, s) => (Binding.name s, T --> Univ_elT, NoSyn))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   212
      (oldTs ~~ rep_names');
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   213
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   214
    (* constructor definitions *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   215
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   216
    fun make_constr_def tname T n ((cname, cargs), (cname', mx)) (thy, defs, eqns, i) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   217
      let
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
   218
        fun constr_arg dt (j, l_args, r_args) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   219
          let val T = typ_of_dtyp descr' sorts dt;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   220
              val free_t = mk_Free "x" T j
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   221
          in (case (strip_dtyp dt, strip_type T) of
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   222
              ((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim
31949
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   223
                (Const (nth all_rep_names m, U --> Univ_elT) $
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   224
                   app_bnds free_t (length Us)) Us :: r_args)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   225
            | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   226
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   227
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
   228
        val (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   229
        val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T;
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21525
diff changeset
   230
        val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21525
diff changeset
   231
        val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   232
        val lhs = list_comb (Const (cname, constrT), l_args);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   233
        val rhs = mk_univ_inj r_args n i;
27330
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27300
diff changeset
   234
        val def = Logic.mk_equals (lhs, Const (abs_name, Univ_elT --> T) $ rhs);
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   235
        val def_name = Long_Name.base_name cname ^ "_def";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   236
        val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   237
          (Const (rep_name, T --> Univ_elT) $ lhs, rhs));
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
   238
        val ([def_thm], thy') =
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
   239
          thy
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24112
diff changeset
   240
          |> Sign.add_consts_i [(cname', constrT, mx)]
29579
cb520b766e00 binding replaces bstring
haftmann
parents: 29389
diff changeset
   241
          |> (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)];
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   242
8436
8a87fa482baf adapted to new PureThy.add_thms etc.;
wenzelm
parents: 8305
diff changeset
   243
      in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   244
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   245
    (* constructor definitions for datatype *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   246
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   247
    fun dt_constr_defs ((((_, (_, _, constrs)), tname), T), constr_syntax)
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   248
        (thy, defs, eqns, rep_congs, dist_lemmas) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   249
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   250
        val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21525
diff changeset
   251
        val rep_const = cterm_of thy
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21525
diff changeset
   252
          (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT));
32957
675c0c7e6a37 explicitly qualify Drule.standard;
wenzelm
parents: 32907
diff changeset
   253
        val cong' =
675c0c7e6a37 explicitly qualify Drule.standard;
wenzelm
parents: 32907
diff changeset
   254
          Drule.standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong);
675c0c7e6a37 explicitly qualify Drule.standard;
wenzelm
parents: 32907
diff changeset
   255
        val dist =
675c0c7e6a37 explicitly qualify Drule.standard;
wenzelm
parents: 32907
diff changeset
   256
          Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   257
        val (thy', defs', eqns', _) = fold ((make_constr_def tname T) (length constrs))
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   258
          (constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   259
      in
32124
954321008785 dropped ancient flat_names option
haftmann
parents: 31986
diff changeset
   260
        (Sign.parent_path thy', defs', eqns @ [eqns'],
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   261
          rep_congs @ [cong'], dist_lemmas @ [dist])
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   262
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   263
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   264
    val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   265
      fold dt_constr_defs
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   266
        (hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax)
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   267
        (thy3 |> Sign.add_consts_i iso_decls |> Sign.parent_path, [], [], [], []);
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   268
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   269
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   270
    (*********** isomorphisms for new types (introduced by typedef) ***********)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   271
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31643
diff changeset
   272
    val _ = message config "Proving isomorphism properties ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   273
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   274
    val newT_iso_axms = map (fn (_, td) =>
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   275
      (collect_simp (#Abs_inverse td), #Rep_inverse td,
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   276
       collect_simp (#Rep td))) typedefs;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   277
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   278
    val newT_iso_inj_thms = map (fn (_, td) =>
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   279
      (collect_simp (#Abs_inject td) RS iffD1, #Rep_inject td RS iffD1)) typedefs;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   280
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   281
    (********* isomorphisms between existing types and "unfolded" types *******)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   282
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   283
    (*---------------------------------------------------------------------*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   284
    (* isomorphisms are defined using primrec-combinators:                 *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   285
    (* generate appropriate functions for instantiating primrec-combinator *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   286
    (*                                                                     *)
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   287
    (*   e.g.  dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y))    *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   288
    (*                                                                     *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   289
    (* also generate characteristic equations for isomorphisms             *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   290
    (*                                                                     *)
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   291
    (*   e.g.  dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   292
    (*---------------------------------------------------------------------*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   293
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   294
    fun make_iso_def k ks n (cname, cargs) (fs, eqns, i) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   295
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   296
        val argTs = map (typ_of_dtyp descr' sorts) cargs;
31949
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   297
        val T = nth recTs k;
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   298
        val rep_name = nth all_rep_names k;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   299
        val rep_const = Const (rep_name, T --> Univ_elT);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   300
        val constr = Const (cname, argTs ---> T);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   301
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   302
        fun process_arg ks' dt (i2, i2', ts, Ts) =
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   303
          let
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   304
            val T' = typ_of_dtyp descr' sorts dt;
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   305
            val (Us, U) = strip_type T'
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   306
          in (case strip_dtyp dt of
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   307
              (_, DtRec j) => if j mem ks' then
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   308
                  (i2 + 1, i2' + 1, ts @ [mk_lim (app_bnds
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   309
                     (mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us],
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   310
                   Ts @ [Us ---> Univ_elT])
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   311
                else
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   312
                  (i2 + 1, i2', ts @ [mk_lim
31949
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   313
                     (Const (nth all_rep_names j, U --> Univ_elT) $
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   314
                        app_bnds (mk_Free "x" T' i2) (length Us)) Us], Ts)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   315
            | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   316
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   317
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   318
        val (i2, i2', ts, Ts) = fold (process_arg ks) cargs (1, 1, [], []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   319
        val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   320
        val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   321
        val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   322
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   323
        val (_, _, ts', _) = fold (process_arg []) cargs (1, 1, [], []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   324
        val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   325
          (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   326
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   327
      in (fs @ [f], eqns @ [eqn], i + 1) end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   328
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   329
    (* define isomorphisms for all mutually recursive datatypes in list ds *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   330
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   331
    fun make_iso_defs ds (thy, char_thms) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   332
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   333
        val ks = map fst ds;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   334
        val (_, (tname, _, _)) = hd ds;
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17261
diff changeset
   335
        val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   336
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   337
        fun process_dt (k, (tname, _, constrs)) (fs, eqns, isos) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   338
          let
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   339
            val (fs', eqns', _) =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   340
              fold (make_iso_def k ks (length constrs)) constrs (fs, eqns, 1);
31949
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   341
            val iso = (nth recTs k, nth all_rep_names k)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   342
          in (fs', eqns', isos @ [iso]) end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   343
        
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   344
        val (fs, eqns, isos) = fold process_dt ds ([], [], []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   345
        val fTs = map fastype_of fs;
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   346
        val defs = map (fn (rec_name, (T, iso_name)) => (Binding.name (Long_Name.base_name iso_name ^ "_def"),
27330
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27300
diff changeset
   347
          Logic.mk_equals (Const (iso_name, T --> Univ_elT),
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 27300
diff changeset
   348
            list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos);
28362
b0ebac91c8d5 explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
wenzelm
parents: 28084
diff changeset
   349
        val (def_thms, thy') =
b0ebac91c8d5 explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
wenzelm
parents: 28084
diff changeset
   350
          apsnd Theory.checkpoint ((PureThy.add_defs false o map Thm.no_attributes) defs thy);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   351
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   352
        (* prove characteristic equations *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   353
5553
ae42b36a50c2 renamed mk_meta_eq to mk_eq
oheimb
parents: 5303
diff changeset
   354
        val rewrites = def_thms @ (map mk_meta_eq rec_rewrites);
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32957
diff changeset
   355
        val char_thms' = map (fn eqn => Skip_Proof.prove_global thy' [] [] eqn
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19540
diff changeset
   356
          (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   357
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   358
      in (thy', char_thms' @ char_thms) end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   359
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   360
    val (thy5, iso_char_thms) = apfst Theory.checkpoint (fold_rev make_iso_defs
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   361
        (tl descr) (Sign.add_path big_name thy4, []));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   362
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   363
    (* prove isomorphism properties *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   364
28362
b0ebac91c8d5 explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
wenzelm
parents: 28084
diff changeset
   365
    fun mk_funs_inv thy thm =
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   366
      let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26531
diff changeset
   367
        val prop = Thm.prop_of thm;
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   368
        val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $
16287
7a03b4b4df67 Type.freeze;
wenzelm
parents: 15574
diff changeset
   369
          (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop;
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 28965
diff changeset
   370
        val used = OldTerm.add_term_tfree_names (a, []);
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   371
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   372
        fun mk_thm i =
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   373
          let
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   374
            val Ts = map (TFree o rpair HOLogic.typeS)
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 20046
diff changeset
   375
              (Name.variant_list used (replicate i "'t"));
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   376
            val f = Free ("f", Ts ---> U)
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32957
diff changeset
   377
          in Skip_Proof.prove_global thy [] [] (Logic.mk_implies
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   378
            (HOLogic.mk_Trueprop (HOLogic.list_all
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   379
               (map (pair "x") Ts, S $ app_bnds f i)),
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   380
             HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts,
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   381
               r $ (a $ app_bnds f i)), f))))
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26711
diff changeset
   382
            (fn _ => EVERY [REPEAT_DETERM_N i (rtac ext 1),
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26711
diff changeset
   383
               REPEAT (etac allE 1), rtac thm 1, atac 1])
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   384
          end
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   385
      in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   386
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   387
    (* prove  inj dt_Rep_i  and  dt_Rep_i x : dt_rep_set_i *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   388
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26711
diff changeset
   389
    val fun_congs = map (fn T => make_elim (Drule.instantiate'
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26711
diff changeset
   390
      [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs;
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26711
diff changeset
   391
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
   392
    fun prove_iso_thms ds (inj_thms, elem_thms) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   393
      let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   394
        val (_, (tname, _, _)) = hd ds;
32727
9072201cd69d avoid compound fields in datatype info record
haftmann
parents: 32712
diff changeset
   395
        val induct = (#induct o the o Symtab.lookup dt_info) tname;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   396
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   397
        fun mk_ind_concl (i, _) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   398
          let
31949
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   399
            val T = nth recTs i;
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   400
            val Rep_t = Const (nth all_rep_names i, T --> Univ_elT);
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   401
            val rep_set_name = nth rep_set_names i
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   402
          in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   403
                HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   404
                  HOLogic.mk_eq (mk_Free "x" T i, Bound 0)),
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   405
              Const (rep_set_name, UnivT') $ (Rep_t $ mk_Free "x" T i))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   406
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   407
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   408
        val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   409
5553
ae42b36a50c2 renamed mk_meta_eq to mk_eq
oheimb
parents: 5303
diff changeset
   410
        val rewrites = map mk_meta_eq iso_char_thms;
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   411
        val inj_thms' = map snd newT_iso_inj_thms @
26359
6d437bde2f1d more antiquotations
haftmann
parents: 26343
diff changeset
   412
          map (fn r => r RS @{thm injD}) inj_thms;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   413
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32957
diff changeset
   414
        val inj_thm = Skip_Proof.prove_global thy5 [] []
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   415
          (HOLogic.mk_Trueprop (mk_conj ind_concl1)) (fn _ => EVERY
32712
ec5976f4d3d8 registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
haftmann
parents: 32124
diff changeset
   416
            [(indtac induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   417
             REPEAT (EVERY
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   418
               [rtac allI 1, rtac impI 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   419
                exh_tac (exh_thm_of dt_info) 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   420
                REPEAT (EVERY
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   421
                  [hyp_subst_tac 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   422
                   rewrite_goals_tac rewrites,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   423
                   REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   424
                   (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   425
                   ORELSE (EVERY
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   426
                     [REPEAT (eresolve_tac (Scons_inject ::
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   427
                        map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1),
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   428
                      REPEAT (cong_tac 1), rtac refl 1,
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   429
                      REPEAT (atac 1 ORELSE (EVERY
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   430
                        [REPEAT (rtac ext 1),
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   431
                         REPEAT (eresolve_tac (mp :: allE ::
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   432
                           map make_elim (Suml_inject :: Sumr_inject ::
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26711
diff changeset
   433
                             Lim_inject :: inj_thms') @ fun_congs) 1),
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19540
diff changeset
   434
                         atac 1]))])])])]);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   435
26359
6d437bde2f1d more antiquotations
haftmann
parents: 26343
diff changeset
   436
        val inj_thms'' = map (fn r => r RS @{thm datatype_injI})
6171
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   437
                             (split_conj_thm inj_thm);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   438
6171
cd237a10cbf8 inj is now a translation of inj_on
paulson
parents: 6092
diff changeset
   439
        val elem_thm = 
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32957
diff changeset
   440
            Skip_Proof.prove_global thy5 [] [] (HOLogic.mk_Trueprop (mk_conj ind_concl2))
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19540
diff changeset
   441
              (fn _ =>
32712
ec5976f4d3d8 registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
haftmann
parents: 32124
diff changeset
   442
               EVERY [(indtac induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19540
diff changeset
   443
                rewrite_goals_tac rewrites,
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19540
diff changeset
   444
                REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19540
diff changeset
   445
                  ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   446
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   447
      in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm))
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   448
      end;
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   449
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
   450
    val (iso_inj_thms_unfolded, iso_elem_thms) =
de76079f973a eliminated some old folds;
wenzelm
parents: 33278
diff changeset
   451
      fold_rev prove_iso_thms (tl descr) ([], map #3 newT_iso_axms);
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   452
    val iso_inj_thms = map snd newT_iso_inj_thms @
26359
6d437bde2f1d more antiquotations
haftmann
parents: 26343
diff changeset
   453
      map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded;
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   454
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   455
    (* prove  dt_rep_set_i x --> x : range dt_Rep_i *)
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   456
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   457
    fun mk_iso_t (((set_name, iso_name), i), T) =
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   458
      let val isoT = T --> Univ_elT
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   459
      in HOLogic.imp $ 
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   460
        (Const (set_name, UnivT') $ mk_Free "x" Univ_elT i) $
31643
b040f1679f77 authentic syntax for Pow and image
haftmann
parents: 31604
diff changeset
   461
          (if i < length newTs then HOLogic.true_const
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   462
           else HOLogic.mk_mem (mk_Free "x" Univ_elT i,
31643
b040f1679f77 authentic syntax for Pow and image
haftmann
parents: 31604
diff changeset
   463
             Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $
30304
d8e4cd2ac2a1 set operations Int, Un, INTER, UNION, Inter, Union, empty, UNIV are now proper qualified constants with authentic syntax
haftmann
parents: 29579
diff changeset
   464
               Const (iso_name, isoT) $ Const (@{const_name UNIV}, HOLogic.mk_setT T)))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   465
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   466
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   467
    val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   468
      (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs)));
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   469
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   470
    (* all the theorems are proved by one single simultaneous induction *)
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   471
26359
6d437bde2f1d more antiquotations
haftmann
parents: 26343
diff changeset
   472
    val range_eqs = map (fn r => mk_meta_eq (r RS @{thm range_ex1_eq}))
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   473
      iso_inj_thms_unfolded;
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   474
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   475
    val iso_thms = if length descr = 1 then [] else
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   476
      Library.drop (length newTs, split_conj_thm
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32957
diff changeset
   477
        (Skip_Proof.prove_global thy5 [] [] iso_t (fn _ => EVERY
25678
2495389bc1f5 Adapted to changes in interface of indtac.
berghofe
parents: 24814
diff changeset
   478
           [(indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   479
            REPEAT (rtac TrueI 1),
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   480
            rewrite_goals_tac (mk_meta_eq choice_eq ::
26359
6d437bde2f1d more antiquotations
haftmann
parents: 26343
diff changeset
   481
              symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs),
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   482
            rewrite_goals_tac (map symmetric range_eqs),
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   483
            REPEAT (EVERY
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   484
              [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @
28362
b0ebac91c8d5 explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
wenzelm
parents: 28084
diff changeset
   485
                 maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
11471
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   486
               TRY (hyp_subst_tac 1),
ba2c252b55ad - Fixed bug in isomorphism proofs (caused by migration from SOME to THE)
berghofe
parents: 11435
diff changeset
   487
               rtac (sym RS range_eqI) 1,
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19540
diff changeset
   488
               resolve_tac iso_char_thms 1])])));
11435
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   489
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   490
    val Abs_inverse_thms' =
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   491
      map #1 newT_iso_axms @
33057
764547b68538 inv_onto -> inv_into
nipkow
parents: 33040
diff changeset
   492
      map2 (fn r_inj => fn r => @{thm f_the_inv_into_f} OF [r_inj, r RS mp])
18330
444f16d232a2 introduced new map2, fold
haftmann
parents: 18314
diff changeset
   493
        iso_inj_thms_unfolded iso_thms;
11435
bd1a7f53c11b replaced "Eps" by "The";
wenzelm
parents: 10911
diff changeset
   494
28362
b0ebac91c8d5 explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
wenzelm
parents: 28084
diff changeset
   495
    val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms';
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   496
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   497
    (******************* freeness theorems for constructors *******************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   498
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31643
diff changeset
   499
    val _ = message config "Proving freeness of constructors ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   500
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   501
    (* prove theorem  Rep_i (Constr_j ...) = Inj_j ...  *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   502
    
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   503
    fun prove_constr_rep_thm eqn =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   504
      let
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   505
        val inj_thms = map fst newT_iso_inj_thms;
26359
6d437bde2f1d more antiquotations
haftmann
parents: 26343
diff changeset
   506
        val rewrites = @{thm o_def} :: constr_defs @ (map (mk_meta_eq o #2) newT_iso_axms)
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32957
diff changeset
   507
      in Skip_Proof.prove_global thy5 [] [] eqn (fn _ => EVERY
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   508
        [resolve_tac inj_thms 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   509
         rewrite_goals_tac rewrites,
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   510
         rtac refl 3,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   511
         resolve_tac rep_intrs 2,
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19540
diff changeset
   512
         REPEAT (resolve_tac iso_elem_thms 1)])
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   513
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   514
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   515
    (*--------------------------------------------------------------*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   516
    (* constr_rep_thms and rep_congs are used to prove distinctness *)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   517
    (* of constructors.                                             *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   518
    (*--------------------------------------------------------------*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   519
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   520
    val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   521
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   522
    val dist_rewrites = map (fn (rep_thms, dist_lemma) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   523
      dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   524
        (constr_rep_thms ~~ dist_lemmas);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   525
32900
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32874
diff changeset
   526
    fun prove_distinct_thms dist_rewrites' (k, ts) =
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32874
diff changeset
   527
      let
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32874
diff changeset
   528
        fun prove [] = []
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32874
diff changeset
   529
          | prove (t :: ts) =
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32874
diff changeset
   530
              let
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32957
diff changeset
   531
                val dist_thm = Skip_Proof.prove_global thy5 [] [] t (fn _ =>
32900
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32874
diff changeset
   532
                  EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1])
32957
675c0c7e6a37 explicitly qualify Drule.standard;
wenzelm
parents: 32907
diff changeset
   533
              in dist_thm :: Drule.standard (dist_thm RS not_sym) :: prove ts end;
32900
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32874
diff changeset
   534
      in prove ts end;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   535
32900
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32874
diff changeset
   536
    val distinct_thms = map2 (prove_distinct_thms)
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32874
diff changeset
   537
      dist_rewrites (DatatypeProp.make_distincts descr sorts);
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6522
diff changeset
   538
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   539
    (* prove injectivity of constructors *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   540
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   541
    fun prove_constr_inj_thm rep_thms t =
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   542
      let val inj_thms = Scons_inject :: (map make_elim
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   543
        (iso_inj_thms @
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   544
          [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject,
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   545
           Lim_inject, Suml_inject, Sumr_inject]))
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32957
diff changeset
   546
      in Skip_Proof.prove_global thy5 [] [] t (fn _ => EVERY
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   547
        [rtac iffI 1,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   548
         REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   549
         dresolve_tac rep_congs 1, dtac box_equals 1,
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   550
         REPEAT (resolve_tac rep_thms 1),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   551
         REPEAT (eresolve_tac inj_thms 1),
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   552
         REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1),
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   553
           REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1),
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19540
diff changeset
   554
           atac 1]))])
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   555
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   556
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   557
    val constr_inject = map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   558
      ((DatatypeProp.make_injs descr sorts) ~~ constr_rep_thms);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   559
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
   560
    val ((constr_inject', distinct_thms'), thy6) =
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
   561
      thy5
32124
954321008785 dropped ancient flat_names option
haftmann
parents: 31986
diff changeset
   562
      |> Sign.parent_path
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
   563
      |> store_thmss "inject" new_type_names constr_inject
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
   564
      ||>> store_thmss "distinct" new_type_names distinct_thms;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   565
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   566
    (*************************** induction theorem ****************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   567
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31643
diff changeset
   568
    val _ = message config "Proving induction rule for datatypes ...";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   569
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   570
    val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @
32999
d603c567170b Replaced inv by the_inv_onto.
berghofe
parents: 32970
diff changeset
   571
      (map (fn r => r RS @{thm the_inv_f_f} RS subst) iso_inj_thms_unfolded);
d603c567170b Replaced inv by the_inv_onto.
berghofe
parents: 32970
diff changeset
   572
    val Rep_inverse_thms' = map (fn r => r RS @{thm the_inv_f_f}) iso_inj_thms_unfolded;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   573
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   574
    fun mk_indrule_lemma ((i, _), T) (prems, concls) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   575
      let
31949
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   576
        val Rep_t = Const (nth all_rep_names i, T --> Univ_elT) $
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   577
          mk_Free "x" T i;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   578
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   579
        val Abs_t = if i < length newTs then
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21525
diff changeset
   580
            Const (Sign.intern_const thy6
31949
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   581
              ("Abs_" ^ (nth new_type_names i)), Univ_elT --> T)
33057
764547b68538 inv_onto -> inv_into
nipkow
parents: 33040
diff changeset
   582
          else Const (@{const_name the_inv_into},
32999
d603c567170b Replaced inv by the_inv_onto.
berghofe
parents: 32970
diff changeset
   583
              [HOLogic.mk_setT T, T --> Univ_elT, Univ_elT] ---> T) $
d603c567170b Replaced inv by the_inv_onto.
berghofe
parents: 32970
diff changeset
   584
            HOLogic.mk_UNIV T $ Const (nth all_rep_names i, T --> Univ_elT)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   585
21021
6f19e5eb3a44 Adapted to new inductive definition package.
berghofe
parents: 20820
diff changeset
   586
      in (prems @ [HOLogic.imp $
31949
3f933687fae9 moved Inductive.myinv to Fun.inv; tuned
haftmann
parents: 31775
diff changeset
   587
            (Const (nth rep_set_names i, UnivT') $ Rep_t) $
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   588
              (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   589
          concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   590
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   591
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   592
    val (indrule_lemma_prems, indrule_lemma_concls) =
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   593
      fold mk_indrule_lemma (descr' ~~ recTs) ([], []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   594
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 21525
diff changeset
   595
    val cert = cterm_of thy6;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   596
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32957
diff changeset
   597
    val indrule_lemma = Skip_Proof.prove_global thy6 [] []
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   598
      (Logic.mk_implies
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   599
        (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   600
         HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   601
           [REPEAT (etac conjE 1),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   602
            REPEAT (EVERY
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   603
              [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1,
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19540
diff changeset
   604
               etac mp 1, resolve_tac iso_elem_thms 1])]);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   605
8305
93aa21ec5494 HOLogic.dest_conj;
wenzelm
parents: 8005
diff changeset
   606
    val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   607
    val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   608
      map (Free o apfst fst o dest_Var) Ps;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   609
    val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   610
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   611
    val dt_induct_prop = DatatypeProp.make_ind descr sorts;
32970
fbd2bb2489a8 operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
wenzelm
parents: 32957
diff changeset
   612
    val dt_induct = Skip_Proof.prove_global thy6 []
17985
d5d576b72371 avoid legacy goals;
wenzelm
parents: 17959
diff changeset
   613
      (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
26711
3a478bfa1650 prove_global: pass context;
wenzelm
parents: 26626
diff changeset
   614
      (fn {prems, ...} => EVERY
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   615
        [rtac indrule_lemma' 1,
25678
2495389bc1f5 Adapted to changes in interface of indtac.
berghofe
parents: 24814
diff changeset
   616
         (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   617
         EVERY (map (fn (prem, r) => (EVERY
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   618
           [REPEAT (eresolve_tac Abs_inverse_thms 1),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   619
            simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1,
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13585
diff changeset
   620
            DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
20046
9c8909fc5865 Goal.prove_global;
wenzelm
parents: 19540
diff changeset
   621
                (prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   622
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
   623
    val ([dt_induct'], thy7) =
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18358
diff changeset
   624
      thy6
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24112
diff changeset
   625
      |> Sign.add_path big_name
29579
cb520b766e00 binding replaces bstring
haftmann
parents: 29389
diff changeset
   626
      |> PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])]
28362
b0ebac91c8d5 explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
wenzelm
parents: 28084
diff changeset
   627
      ||> Sign.parent_path
b0ebac91c8d5 explicit checkpoint for low-level (global) theory operations admits concurrent SkipProof.prove;
wenzelm
parents: 28084
diff changeset
   628
      ||> Theory.checkpoint;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   629
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 17985
diff changeset
   630
  in
32907
0300f8dd63ea dropped rule duplicates
haftmann
parents: 32904
diff changeset
   631
    ((constr_inject', distinct_thms', dt_induct'), thy7)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   632
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   633
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   634
end;