src/HOL/Tools/Datatype/datatype.ML
author wenzelm
Wed, 14 Dec 2011 20:36:17 +0100
changeset 45878 2dad374cf440
parent 45863 afdb92130f5a
child 45879 71b8d0d170b1
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
     1
(*  Title:      HOL/Tools/Datatype/datatype.ML
11539
0f17da240450 tuned headers;
wenzelm
parents: 11345
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     3
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
     4
Datatype package: definitional introduction of datatypes
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
     5
with proof of characteristic theorems: injectivity / distinctness
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
     6
of constructors and induction.  Main interface to datatypes
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
     7
after full bootstrap of datatype package.
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     8
*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     9
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31689
diff changeset
    10
signature DATATYPE =
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    11
sig
33963
977b94b64905 renamed former datatype.ML to datatype_data.ML; datatype.ML provides uniform view on datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33959
diff changeset
    12
  include DATATYPE_DATA
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    13
  type spec =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    14
    (binding * (string * sort) list * mixfix) *
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    15
    (binding * typ list * mixfix) list
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    16
  type spec_cmd =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    17
    (binding * (string * string option) list * mixfix) *
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    18
    (binding * string list * mixfix) list
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    19
  val read_specs: spec_cmd list -> theory -> spec list * Proof.context
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    20
  val check_specs: spec list -> theory -> spec list * Proof.context
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    21
  val add_datatype: config -> spec list -> theory -> string list * theory
45863
afdb92130f5a tuned signature;
wenzelm
parents: 45839
diff changeset
    22
  val add_datatype_cmd: config -> spec_cmd list -> theory -> string list * theory
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    23
  val spec_cmd: spec_cmd parser
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    24
end;
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    25
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    26
structure Datatype : DATATYPE =
33963
977b94b64905 renamed former datatype.ML to datatype_data.ML; datatype.ML provides uniform view on datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33959
diff changeset
    27
struct
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    28
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    29
(** auxiliary **)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    30
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    31
val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    32
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    33
val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    34
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
    35
fun exh_thm_of (dt_info : Datatype_Aux.info Symtab.table) tname =
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    36
  #exhaust (the (Symtab.lookup dt_info tname));
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    37
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    38
val In0_inject = @{thm In0_inject};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    39
val In1_inject = @{thm In1_inject};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    40
val Scons_inject = @{thm Scons_inject};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    41
val Leaf_inject = @{thm Leaf_inject};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    42
val In0_eq = @{thm In0_eq};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    43
val In1_eq = @{thm In1_eq};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    44
val In0_not_In1 = @{thm In0_not_In1};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    45
val In1_not_In0 = @{thm In1_not_In0};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    46
val Lim_inject = @{thm Lim_inject};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    47
val Inl_inject = @{thm Inl_inject};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    48
val Inr_inject = @{thm Inr_inject};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    49
val Suml_inject = @{thm Suml_inject};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    50
val Sumr_inject = @{thm Sumr_inject};
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    51
40719
acb830207103 keep private things private, without comments;
wenzelm
parents: 40237
diff changeset
    52
val datatype_injI =
acb830207103 keep private things private, without comments;
wenzelm
parents: 40237
diff changeset
    53
  @{lemma "(!!x. ALL y. f x = f y --> x = y) ==> inj f" by (simp add: inj_on_def)};
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    54
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    55
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    56
(** proof of characteristic theorems **)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    57
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
    58
fun representation_proofs (config : Datatype_Aux.config) (dt_info : Datatype_Aux.info Symtab.table)
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
    59
    descr types_syntax constr_syntax case_names_induct thy =
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    60
  let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    61
    val descr' = flat descr;
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
    62
    val new_type_names = map (Binding.name_of o fst) types_syntax;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    63
    val big_name = space_implode "_" new_type_names;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    64
    val thy1 = Sign.add_path big_name thy;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    65
    val big_rec_name = big_name ^ "_rep_set";
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    66
    val rep_set_names' =
45743
857b7fcb0365 misc tuning;
wenzelm
parents: 45742
diff changeset
    67
      if length descr' = 1 then [big_rec_name]
45821
wenzelm
parents: 45743
diff changeset
    68
      else map (prefix (big_rec_name ^ "_") o string_of_int) (1 upto length descr');
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    69
    val rep_set_names = map (Sign.full_bname thy1) rep_set_names';
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    70
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
    71
    val tyvars = map (fn (_, (_, Ts, _)) => map Datatype_Aux.dest_DtTFree Ts) (hd descr);
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
    72
    val leafTs' = Datatype_Aux.get_nonrec_types descr';
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
    73
    val branchTs = Datatype_Aux.get_branching_types descr';
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
    74
    val branchT =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
    75
      if null branchTs then HOLogic.unitT
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 36960
diff changeset
    76
      else Balanced_Tree.make (fn (T, U) => Type (@{type_name Sum_Type.sum}, [T, U])) branchTs;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
    77
    val arities = remove (op =) 0 (Datatype_Aux.get_arities descr');
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    78
    val unneeded_vars =
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
    79
      subtract (op =) (fold Term.add_tfreesT (leafTs' @ branchTs) []) (hd tyvars);
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
    80
    val leafTs = leafTs' @ map TFree unneeded_vars;
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
    81
    val recTs = Datatype_Aux.get_rec_types descr';
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    82
    val (newTs, oldTs) = chop (length (hd descr)) recTs;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
    83
    val sumT =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
    84
      if null leafTs then HOLogic.unitT
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 36960
diff changeset
    85
      else Balanced_Tree.make (fn (T, U) => Type (@{type_name Sum_Type.sum}, [T, U])) leafTs;
45878
wenzelm
parents: 45863
diff changeset
    86
    val Univ_elT = HOLogic.mk_setT (Type (@{type_name Datatype.node}, [sumT, branchT]));
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    87
    val UnivT = HOLogic.mk_setT Univ_elT;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    88
    val UnivT' = Univ_elT --> HOLogic.boolT;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    89
    val Collect = Const (@{const_name Collect}, UnivT' --> UnivT);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    90
45878
wenzelm
parents: 45863
diff changeset
    91
    val In0 = Const (@{const_name Datatype.In0}, Univ_elT --> Univ_elT);
wenzelm
parents: 45863
diff changeset
    92
    val In1 = Const (@{const_name Datatype.In1}, Univ_elT --> Univ_elT);
wenzelm
parents: 45863
diff changeset
    93
    val Leaf = Const (@{const_name Datatype.Leaf}, sumT --> Univ_elT);
wenzelm
parents: 45863
diff changeset
    94
    val Lim = Const (@{const_name Datatype.Lim}, (branchT --> Univ_elT) --> Univ_elT);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    95
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    96
    (* make injections needed for embedding types in leaves *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    97
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    98
    fun mk_inj T' x =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    99
      let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   100
        fun mk_inj' T n i =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   101
          if n = 1 then x
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   102
          else
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   103
            let
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   104
              val n2 = n div 2;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   105
              val Type (_, [T1, T2]) = T;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   106
            in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   107
              if i <= n2
45821
wenzelm
parents: 45743
diff changeset
   108
              then Const (@{const_name Inl}, T1 --> T) $ mk_inj' T1 n2 i
wenzelm
parents: 45743
diff changeset
   109
              else Const (@{const_name Inr}, T2 --> T) $ mk_inj' T2 (n - n2) (i - n2)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   110
            end;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   111
      in mk_inj' sumT (length leafTs) (1 + find_index (fn T'' => T'' = T') leafTs) end;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   112
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   113
    (* make injections for constructors *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   114
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   115
    fun mk_univ_inj ts = Balanced_Tree.access
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   116
      {left = fn t => In0 $ t,
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   117
        right = fn t => In1 $ t,
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   118
        init =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   119
          if ts = [] then Const (@{const_name undefined}, Univ_elT)
45878
wenzelm
parents: 45863
diff changeset
   120
          else foldr1 (HOLogic.mk_binop @{const_name Datatype.Scons}) ts};
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   121
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   122
    (* function spaces *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   123
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   124
    fun mk_fun_inj T' x =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   125
      let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   126
        fun mk_inj T n i =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   127
          if n = 1 then x
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   128
          else
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   129
            let
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   130
              val n2 = n div 2;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   131
              val Type (_, [T1, T2]) = T;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   132
              fun mkT U = (U --> Univ_elT) --> T --> Univ_elT;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   133
            in
45878
wenzelm
parents: 45863
diff changeset
   134
              if i <= n2 then Const (@{const_name Sum_Type.Suml}, mkT T1) $ mk_inj T1 n2 i
wenzelm
parents: 45863
diff changeset
   135
              else Const (@{const_name Sum_Type.Sumr}, mkT T2) $ mk_inj T2 (n - n2) (i - n2)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   136
            end;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   137
      in mk_inj branchT (length branchTs) (1 + find_index (fn T'' => T'' = T') branchTs) end;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   138
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   139
    fun mk_lim t Ts = fold_rev (fn T => fn t => Lim $ mk_fun_inj T (Abs ("x", T, t))) Ts t;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   140
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   141
    (************** generate introduction rules for representing set **********)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   142
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   143
    val _ = Datatype_Aux.message config "Constructing representing sets ...";
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   144
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   145
    (* make introduction rule for a single constructor *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   146
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   147
    fun make_intr s n (i, (_, cargs)) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   148
      let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   149
        fun mk_prem dt (j, prems, ts) =
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   150
          (case Datatype_Aux.strip_dtyp dt of
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   151
            (dts, Datatype_Aux.DtRec k) =>
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   152
              let
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   153
                val Ts = map (Datatype_Aux.typ_of_dtyp descr') dts;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   154
                val free_t =
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   155
                  Datatype_Aux.app_bnds (Datatype_Aux.mk_Free "x" (Ts ---> Univ_elT) j) (length Ts)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   156
              in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   157
                (j + 1, list_all (map (pair "x") Ts,
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   158
                  HOLogic.mk_Trueprop
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   159
                    (Free (nth rep_set_names' k, UnivT') $ free_t)) :: prems,
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   160
                mk_lim free_t Ts :: ts)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   161
              end
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   162
          | _ =>
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   163
              let val T = Datatype_Aux.typ_of_dtyp descr' dt
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   164
              in (j + 1, prems, (Leaf $ mk_inj T (Datatype_Aux.mk_Free "x" T j)) :: ts) end);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   165
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   166
        val (_, prems, ts) = fold_rev mk_prem cargs (1, [], []);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   167
        val concl = HOLogic.mk_Trueprop (Free (s, UnivT') $ mk_univ_inj ts n i);
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   168
      in Logic.list_implies (prems, concl) end;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   169
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   170
    val intr_ts = maps (fn ((_, (_, _, constrs)), rep_set_name) =>
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   171
      map (make_intr rep_set_name (length constrs))
45821
wenzelm
parents: 45743
diff changeset
   172
        ((1 upto length constrs) ~~ constrs)) (descr' ~~ rep_set_names');
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   173
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   174
    val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   175
      thy1
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   176
      |> Sign.map_naming Name_Space.conceal
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   177
      |> Inductive.add_inductive_global
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   178
          {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name,
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   179
           coind = false, no_elim = true, no_ind = false, skip_mono = true, fork_mono = false}
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   180
          (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   181
          (map (fn x => (Attrib.empty_binding, x)) intr_ts) []
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   182
      ||> Sign.restore_naming thy1
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   183
      ||> Theory.checkpoint;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   184
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   185
    (********************************* typedef ********************************)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   186
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   187
    val (typedefs, thy3) = thy2
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   188
      |> Sign.parent_path
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   189
      |> fold_map
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   190
        (fn (((name, mx), tvs), c) =>
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   191
          Typedef.add_typedef_global false NONE (name, tvs, mx)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   192
            (Collect $ Const (c, UnivT')) NONE
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   193
            (rtac exI 1 THEN rtac CollectI 1 THEN
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   194
              QUIET_BREADTH_FIRST (has_fewer_prems 1)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   195
              (resolve_tac rep_intrs 1)))
45878
wenzelm
parents: 45863
diff changeset
   196
        (types_syntax ~~ tyvars ~~ take (length newTs) rep_set_names)
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   197
      ||> Sign.add_path big_name;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   198
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   199
    (*********************** definition of constructors ***********************)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   200
45878
wenzelm
parents: 45863
diff changeset
   201
    val big_rep_name = big_name ^ "_Rep_";
45229
wenzelm
parents: 44241
diff changeset
   202
    val rep_names = map (prefix "Rep_") new_type_names;
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   203
    val rep_names' = map (fn i => big_rep_name ^ string_of_int i) (1 upto length (flat (tl descr)));
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   204
    val all_rep_names =
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   205
      map (Sign.intern_const thy3) rep_names @
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   206
      map (Sign.full_bname thy3) rep_names';
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   207
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   208
    (* isomorphism declarations *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   209
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   210
    val iso_decls = map (fn (T, s) => (Binding.name s, T --> Univ_elT, NoSyn))
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   211
      (oldTs ~~ rep_names');
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   212
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   213
    (* constructor definitions *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   214
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   215
    fun make_constr_def tname T n ((cname, cargs), (cname', mx)) (thy, defs, eqns, i) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   216
      let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   217
        fun constr_arg dt (j, l_args, r_args) =
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   218
          let
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   219
            val T = Datatype_Aux.typ_of_dtyp descr' dt;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   220
            val free_t = Datatype_Aux.mk_Free "x" T j;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   221
          in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   222
            (case (Datatype_Aux.strip_dtyp dt, strip_type T) of
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   223
              ((_, Datatype_Aux.DtRec m), (Us, U)) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   224
                (j + 1, free_t :: l_args, mk_lim
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   225
                  (Const (nth all_rep_names m, U --> Univ_elT) $
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   226
                    Datatype_Aux.app_bnds free_t (length Us)) Us :: r_args)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   227
            | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args))
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   228
          end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   229
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   230
        val (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []);
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   231
        val constrT = map (Datatype_Aux.typ_of_dtyp descr') cargs ---> T;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   232
        val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   233
        val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   234
        val lhs = list_comb (Const (cname, constrT), l_args);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   235
        val rhs = mk_univ_inj r_args n i;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   236
        val def = Logic.mk_equals (lhs, Const (abs_name, Univ_elT --> T) $ rhs);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   237
        val def_name = Long_Name.base_name cname ^ "_def";
45821
wenzelm
parents: 45743
diff changeset
   238
        val eqn =
wenzelm
parents: 45743
diff changeset
   239
          HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (rep_name, T --> Univ_elT) $ lhs, rhs));
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   240
        val ([def_thm], thy') =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   241
          thy
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   242
          |> Sign.add_consts_i [(cname', constrT, mx)]
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39302
diff changeset
   243
          |> (Global_Theory.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)];
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   244
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   245
      in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   246
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   247
    (* constructor definitions for datatype *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   248
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   249
    fun dt_constr_defs ((((_, (_, _, constrs)), tname), T), constr_syntax)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   250
        (thy, defs, eqns, rep_congs, dist_lemmas) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   251
      let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   252
        val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   253
        val rep_const =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   254
          cterm_of thy (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT));
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   255
        val cong' =
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33969
diff changeset
   256
          Drule.export_without_context
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33969
diff changeset
   257
            (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   258
        val dist =
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33969
diff changeset
   259
          Drule.export_without_context
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33969
diff changeset
   260
            (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   261
        val (thy', defs', eqns', _) =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   262
          fold ((make_constr_def tname T) (length constrs))
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   263
            (constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   264
      in
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   265
        (Sign.parent_path thy', defs', eqns @ [eqns'],
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   266
          rep_congs @ [cong'], dist_lemmas @ [dist])
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   267
      end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   268
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   269
    val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   270
      fold dt_constr_defs
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   271
        (hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   272
        (thy3 |> Sign.add_consts_i iso_decls |> Sign.parent_path, [], [], [], []);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   273
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   274
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   275
    (*********** isomorphisms for new types (introduced by typedef) ***********)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   276
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   277
    val _ = Datatype_Aux.message config "Proving isomorphism properties ...";
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   278
35994
9cc3df9a606e Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents: 35842
diff changeset
   279
    val newT_iso_axms = map (fn (_, (_, td)) =>
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   280
      (collect_simp (#Abs_inverse td), #Rep_inverse td,
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   281
       collect_simp (#Rep td))) typedefs;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   282
35994
9cc3df9a606e Typedef.info: separate global and local part, only the latter is transformed by morphisms;
wenzelm
parents: 35842
diff changeset
   283
    val newT_iso_inj_thms = map (fn (_, (_, td)) =>
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   284
      (collect_simp (#Abs_inject td) RS iffD1, #Rep_inject td RS iffD1)) typedefs;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   285
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   286
    (********* isomorphisms between existing types and "unfolded" types *******)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   287
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   288
    (*---------------------------------------------------------------------*)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   289
    (* isomorphisms are defined using primrec-combinators:                 *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   290
    (* generate appropriate functions for instantiating primrec-combinator *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   291
    (*                                                                     *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   292
    (*   e.g.  dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y))    *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   293
    (*                                                                     *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   294
    (* also generate characteristic equations for isomorphisms             *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   295
    (*                                                                     *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   296
    (*   e.g.  dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   297
    (*---------------------------------------------------------------------*)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   298
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   299
    fun make_iso_def k ks n (cname, cargs) (fs, eqns, i) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   300
      let
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   301
        val argTs = map (Datatype_Aux.typ_of_dtyp descr') cargs;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   302
        val T = nth recTs k;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   303
        val rep_name = nth all_rep_names k;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   304
        val rep_const = Const (rep_name, T --> Univ_elT);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   305
        val constr = Const (cname, argTs ---> T);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   306
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   307
        fun process_arg ks' dt (i2, i2', ts, Ts) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   308
          let
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   309
            val T' = Datatype_Aux.typ_of_dtyp descr' dt;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   310
            val (Us, U) = strip_type T'
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   311
          in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   312
            (case Datatype_Aux.strip_dtyp dt of
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   313
              (_, Datatype_Aux.DtRec j) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   314
                if member (op =) ks' j then
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   315
                  (i2 + 1, i2' + 1, ts @ [mk_lim (Datatype_Aux.app_bnds
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   316
                     (Datatype_Aux.mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us],
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   317
                   Ts @ [Us ---> Univ_elT])
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   318
                else
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   319
                  (i2 + 1, i2', ts @ [mk_lim
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   320
                     (Const (nth all_rep_names j, U --> Univ_elT) $
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   321
                        Datatype_Aux.app_bnds (Datatype_Aux.mk_Free "x" T' i2) (length Us)) Us], Ts)
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   322
            | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (Datatype_Aux.mk_Free "x" T' i2)], Ts))
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   323
          end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   324
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   325
        val (i2, i2', ts, Ts) = fold (process_arg ks) cargs (1, 1, [], []);
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   326
        val xs = map (uncurry (Datatype_Aux.mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   327
        val ys = map (uncurry (Datatype_Aux.mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
45878
wenzelm
parents: 45863
diff changeset
   328
        val f = fold_rev lambda (xs @ ys) (mk_univ_inj ts n i);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   329
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   330
        val (_, _, ts', _) = fold (process_arg []) cargs (1, 1, [], []);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   331
        val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   332
          (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i))
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   333
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   334
      in (fs @ [f], eqns @ [eqn], i + 1) end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   335
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   336
    (* define isomorphisms for all mutually recursive datatypes in list ds *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   337
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   338
    fun make_iso_defs ds (thy, char_thms) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   339
      let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   340
        val ks = map fst ds;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   341
        val (_, (tname, _, _)) = hd ds;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   342
        val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   343
45878
wenzelm
parents: 45863
diff changeset
   344
        fun process_dt (k, (_, _, constrs)) (fs, eqns, isos) =
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   345
          let
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   346
            val (fs', eqns', _) = fold (make_iso_def k ks (length constrs)) constrs (fs, eqns, 1);
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   347
            val iso = (nth recTs k, nth all_rep_names k);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   348
          in (fs', eqns', isos @ [iso]) end;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   349
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   350
        val (fs, eqns, isos) = fold process_dt ds ([], [], []);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   351
        val fTs = map fastype_of fs;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   352
        val defs =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   353
          map (fn (rec_name, (T, iso_name)) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   354
            (Binding.name (Long_Name.base_name iso_name ^ "_def"),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   355
              Logic.mk_equals (Const (iso_name, T --> Univ_elT),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   356
                list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   357
        val (def_thms, thy') =
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39302
diff changeset
   358
          apsnd Theory.checkpoint ((Global_Theory.add_defs false o map Thm.no_attributes) defs thy);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   359
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   360
        (* prove characteristic equations *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   361
45878
wenzelm
parents: 45863
diff changeset
   362
        val rewrites = def_thms @ map mk_meta_eq rec_rewrites;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   363
        val char_thms' =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   364
          map (fn eqn => Skip_Proof.prove_global thy' [] [] eqn
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   365
            (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   366
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   367
      in (thy', char_thms' @ char_thms) end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   368
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   369
    val (thy5, iso_char_thms) =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   370
      apfst Theory.checkpoint (fold_rev make_iso_defs (tl descr) (Sign.add_path big_name thy4, []));
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   371
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   372
    (* prove isomorphism properties *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   373
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   374
    fun mk_funs_inv thy thm =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   375
      let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   376
        val prop = Thm.prop_of thm;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   377
        val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   378
          (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.legacy_freeze prop;
45738
0430f9123e43 eliminated some legacy operations;
wenzelm
parents: 45735
diff changeset
   379
        val used = Term.add_tfree_names a [];
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   380
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   381
        fun mk_thm i =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   382
          let
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   383
            val Ts = map (TFree o rpair HOLogic.typeS) (Name.variant_list used (replicate i "'t"));
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   384
            val f = Free ("f", Ts ---> U);
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   385
          in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   386
            Skip_Proof.prove_global thy [] []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   387
              (Logic.mk_implies
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   388
                (HOLogic.mk_Trueprop (HOLogic.list_all
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   389
                   (map (pair "x") Ts, S $ Datatype_Aux.app_bnds f i)),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   390
                 HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   391
                   r $ (a $ Datatype_Aux.app_bnds f i)), f))))
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   392
              (fn _ => EVERY [REPEAT_DETERM_N i (rtac ext 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   393
                 REPEAT (etac allE 1), rtac thm 1, atac 1])
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   394
          end
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   395
      in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   396
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   397
    (* prove  inj dt_Rep_i  and  dt_Rep_i x : dt_rep_set_i *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   398
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   399
    val fun_congs =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   400
      map (fn T => make_elim (Drule.instantiate' [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   401
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   402
    fun prove_iso_thms ds (inj_thms, elem_thms) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   403
      let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   404
        val (_, (tname, _, _)) = hd ds;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   405
        val induct = #induct (the (Symtab.lookup dt_info tname));
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   406
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   407
        fun mk_ind_concl (i, _) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   408
          let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   409
            val T = nth recTs i;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   410
            val Rep_t = Const (nth all_rep_names i, T --> Univ_elT);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   411
            val rep_set_name = nth rep_set_names i;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   412
          in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   413
            (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   414
                HOLogic.mk_eq (Rep_t $ Datatype_Aux.mk_Free "x" T i, Rep_t $ Bound 0) $
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   415
                  HOLogic.mk_eq (Datatype_Aux.mk_Free "x" T i, Bound 0)),
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   416
              Const (rep_set_name, UnivT') $ (Rep_t $ Datatype_Aux.mk_Free "x" T i))
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   417
          end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   418
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   419
        val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   420
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   421
        val rewrites = map mk_meta_eq iso_char_thms;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   422
        val inj_thms' = map snd newT_iso_inj_thms @ map (fn r => r RS @{thm injD}) inj_thms;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   423
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   424
        val inj_thm =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   425
          Skip_Proof.prove_global thy5 [] []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   426
            (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj ind_concl1))
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   427
            (fn _ => EVERY
45735
7d7d7af647a9 tuned signature;
wenzelm
parents: 45701
diff changeset
   428
              [(Datatype_Aux.ind_tac induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1,
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   429
               REPEAT (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   430
                 [rtac allI 1, rtac impI 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   431
                  Datatype_Aux.exh_tac (exh_thm_of dt_info) 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   432
                  REPEAT (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   433
                    [hyp_subst_tac 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   434
                     rewrite_goals_tac rewrites,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   435
                     REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   436
                     (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   437
                     ORELSE (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   438
                       [REPEAT (eresolve_tac (Scons_inject ::
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   439
                          map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   440
                        REPEAT (cong_tac 1), rtac refl 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   441
                        REPEAT (atac 1 ORELSE (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   442
                          [REPEAT (rtac ext 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   443
                           REPEAT (eresolve_tac (mp :: allE ::
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   444
                             map make_elim (Suml_inject :: Sumr_inject ::
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   445
                               Lim_inject :: inj_thms') @ fun_congs) 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   446
                           atac 1]))])])])]);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   447
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   448
        val inj_thms'' = map (fn r => r RS datatype_injI) (Datatype_Aux.split_conj_thm inj_thm);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   449
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   450
        val elem_thm =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   451
          Skip_Proof.prove_global thy5 [] []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   452
            (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj ind_concl2))
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   453
            (fn _ =>
45735
7d7d7af647a9 tuned signature;
wenzelm
parents: 45701
diff changeset
   454
              EVERY [(Datatype_Aux.ind_tac induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1,
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   455
                rewrite_goals_tac rewrites,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   456
                REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   457
                  ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   458
45878
wenzelm
parents: 45863
diff changeset
   459
      in (inj_thms'' @ inj_thms, elem_thms @ Datatype_Aux.split_conj_thm elem_thm) end;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   460
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   461
    val (iso_inj_thms_unfolded, iso_elem_thms) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   462
      fold_rev prove_iso_thms (tl descr) ([], map #3 newT_iso_axms);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   463
    val iso_inj_thms =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   464
      map snd newT_iso_inj_thms @ map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   465
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   466
    (* prove  dt_rep_set_i x --> x : range dt_Rep_i *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   467
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   468
    fun mk_iso_t (((set_name, iso_name), i), T) =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   469
      let val isoT = T --> Univ_elT in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   470
        HOLogic.imp $
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   471
          (Const (set_name, UnivT') $ Datatype_Aux.mk_Free "x" Univ_elT i) $
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45738
diff changeset
   472
            (if i < length newTs then @{term True}
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   473
             else HOLogic.mk_mem (Datatype_Aux.mk_Free "x" Univ_elT i,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   474
               Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   475
                 Const (iso_name, isoT) $ Const (@{const_abbrev UNIV}, HOLogic.mk_setT T)))
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   476
      end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   477
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   478
    val iso_t = HOLogic.mk_Trueprop (Datatype_Aux.mk_conj (map mk_iso_t
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   479
      (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs)));
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   480
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   481
    (* all the theorems are proved by one single simultaneous induction *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   482
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   483
    val range_eqs = map (fn r => mk_meta_eq (r RS @{thm range_ex1_eq})) iso_inj_thms_unfolded;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   484
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   485
    val iso_thms =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   486
      if length descr = 1 then []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   487
      else
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   488
        drop (length newTs) (Datatype_Aux.split_conj_thm
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   489
          (Skip_Proof.prove_global thy5 [] [] iso_t (fn _ => EVERY
45735
7d7d7af647a9 tuned signature;
wenzelm
parents: 45701
diff changeset
   490
             [(Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1,
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   491
              REPEAT (rtac TrueI 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   492
              rewrite_goals_tac (mk_meta_eq @{thm choice_eq} ::
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   493
                Thm.symmetric (mk_meta_eq @{thm fun_eq_iff}) :: range_eqs),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   494
              rewrite_goals_tac (map Thm.symmetric range_eqs),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   495
              REPEAT (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   496
                [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   497
                   maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   498
                 TRY (hyp_subst_tac 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   499
                 rtac (sym RS range_eqI) 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   500
                 resolve_tac iso_char_thms 1])])));
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   501
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   502
    val Abs_inverse_thms' =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   503
      map #1 newT_iso_axms @
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   504
      map2 (fn r_inj => fn r => @{thm f_the_inv_into_f} OF [r_inj, r RS mp])
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   505
        iso_inj_thms_unfolded iso_thms;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   506
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   507
    val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms';
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   508
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   509
    (******************* freeness theorems for constructors *******************)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   510
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   511
    val _ = Datatype_Aux.message config "Proving freeness of constructors ...";
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   512
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   513
    (* prove theorem  Rep_i (Constr_j ...) = Inj_j ...  *)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   514
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   515
    fun prove_constr_rep_thm eqn =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   516
      let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   517
        val inj_thms = map fst newT_iso_inj_thms;
45878
wenzelm
parents: 45863
diff changeset
   518
        val rewrites = @{thm o_def} :: constr_defs @ map (mk_meta_eq o #2) newT_iso_axms;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   519
      in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   520
        Skip_Proof.prove_global thy5 [] [] eqn
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   521
        (fn _ => EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   522
          [resolve_tac inj_thms 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   523
           rewrite_goals_tac rewrites,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   524
           rtac refl 3,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   525
           resolve_tac rep_intrs 2,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   526
           REPEAT (resolve_tac iso_elem_thms 1)])
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   527
      end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   528
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   529
    (*--------------------------------------------------------------*)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   530
    (* constr_rep_thms and rep_congs are used to prove distinctness *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   531
    (* of constructors.                                             *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   532
    (*--------------------------------------------------------------*)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   533
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   534
    val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   535
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   536
    val dist_rewrites =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   537
      map (fn (rep_thms, dist_lemma) =>
45878
wenzelm
parents: 45863
diff changeset
   538
        dist_lemma :: (rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   539
          (constr_rep_thms ~~ dist_lemmas);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   540
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   541
    fun prove_distinct_thms dist_rewrites' (k, ts) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   542
      let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   543
        fun prove [] = []
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   544
          | prove (t :: ts) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   545
              let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   546
                val dist_thm = Skip_Proof.prove_global thy5 [] [] t (fn _ =>
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   547
                  EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1])
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 33969
diff changeset
   548
              in dist_thm :: Drule.export_without_context (dist_thm RS not_sym) :: prove ts end;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   549
      in prove ts end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   550
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   551
    val distinct_thms =
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   552
      map2 (prove_distinct_thms) dist_rewrites (Datatype_Prop.make_distincts descr);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   553
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   554
    (* prove injectivity of constructors *)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   555
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   556
    fun prove_constr_inj_thm rep_thms t =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   557
      let
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   558
        val inj_thms = Scons_inject ::
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   559
          map make_elim
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   560
            (iso_inj_thms @
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   561
              [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   562
               Lim_inject, Suml_inject, Sumr_inject])
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   563
      in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   564
        Skip_Proof.prove_global thy5 [] [] t
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   565
          (fn _ => EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   566
            [rtac iffI 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   567
             REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   568
             dresolve_tac rep_congs 1, dtac box_equals 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   569
             REPEAT (resolve_tac rep_thms 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   570
             REPEAT (eresolve_tac inj_thms 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   571
             REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   572
               REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   573
               atac 1]))])
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   574
      end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   575
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   576
    val constr_inject =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   577
      map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts)
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   578
        (Datatype_Prop.make_injs descr ~~ constr_rep_thms);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   579
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   580
    val ((constr_inject', distinct_thms'), thy6) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   581
      thy5
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   582
      |> Sign.parent_path
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   583
      |> Datatype_Aux.store_thmss "inject" new_type_names constr_inject
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   584
      ||>> Datatype_Aux.store_thmss "distinct" new_type_names distinct_thms;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   585
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   586
    (*************************** induction theorem ****************************)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   587
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   588
    val _ = Datatype_Aux.message config "Proving induction rule for datatypes ...";
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   589
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   590
    val Rep_inverse_thms =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   591
      map (fn (_, iso, _) => iso RS subst) newT_iso_axms @
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   592
      map (fn r => r RS @{thm the_inv_f_f} RS subst) iso_inj_thms_unfolded;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   593
    val Rep_inverse_thms' = map (fn r => r RS @{thm the_inv_f_f}) iso_inj_thms_unfolded;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   594
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   595
    fun mk_indrule_lemma ((i, _), T) (prems, concls) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   596
      let
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   597
        val Rep_t =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   598
          Const (nth all_rep_names i, T --> Univ_elT) $ Datatype_Aux.mk_Free "x" T i;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   599
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   600
        val Abs_t =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   601
          if i < length newTs then
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   602
            Const (Sign.intern_const thy6 ("Abs_" ^ nth new_type_names i), Univ_elT --> T)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   603
          else Const (@{const_name the_inv_into},
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   604
              [HOLogic.mk_setT T, T --> Univ_elT, Univ_elT] ---> T) $
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   605
            HOLogic.mk_UNIV T $ Const (nth all_rep_names i, T --> Univ_elT);
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   606
      in
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   607
        (prems @
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   608
          [HOLogic.imp $
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   609
            (Const (nth rep_set_names i, UnivT') $ Rep_t) $
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   610
              (Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   611
         concls @
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   612
          [Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ Datatype_Aux.mk_Free "x" T i])
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   613
      end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   614
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   615
    val (indrule_lemma_prems, indrule_lemma_concls) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   616
      fold mk_indrule_lemma (descr' ~~ recTs) ([], []);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   617
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   618
    val cert = cterm_of thy6;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   619
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   620
    val indrule_lemma =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   621
      Skip_Proof.prove_global thy6 [] []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   622
        (Logic.mk_implies
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   623
          (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj indrule_lemma_prems),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   624
           HOLogic.mk_Trueprop (Datatype_Aux.mk_conj indrule_lemma_concls)))
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   625
        (fn _ =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   626
          EVERY
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   627
           [REPEAT (etac conjE 1),
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   628
            REPEAT (EVERY
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   629
              [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1,
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   630
               etac mp 1, resolve_tac iso_elem_thms 1])]);
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   631
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   632
    val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   633
    val frees =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   634
      if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))]
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   635
      else map (Free o apfst fst o dest_Var) Ps;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   636
    val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   637
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   638
    val dt_induct_prop = Datatype_Prop.make_ind descr;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   639
    val dt_induct =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   640
      Skip_Proof.prove_global thy6 []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   641
      (Logic.strip_imp_prems dt_induct_prop)
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   642
      (Logic.strip_imp_concl dt_induct_prop)
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   643
      (fn {prems, ...} =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   644
        EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   645
          [rtac indrule_lemma' 1,
45735
7d7d7af647a9 tuned signature;
wenzelm
parents: 45701
diff changeset
   646
           (Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1,
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   647
           EVERY (map (fn (prem, r) => (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   648
             [REPEAT (eresolve_tac Abs_inverse_thms 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   649
              simp_tac (HOL_basic_ss addsimps (Thm.symmetric r :: Rep_inverse_thms')) 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   650
              DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
45878
wenzelm
parents: 45863
diff changeset
   651
                  (prems ~~ (constr_defs @ map mk_meta_eq iso_char_thms)))]);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   652
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   653
    val ([dt_induct'], thy7) =
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   654
      thy6
40929
7ff03a5e044f theorem names generated by the (rep_)datatype command now have mandatory qualifiers
huffman
parents: 40719
diff changeset
   655
      |> Global_Theory.add_thms
7ff03a5e044f theorem names generated by the (rep_)datatype command now have mandatory qualifiers
huffman
parents: 40719
diff changeset
   656
        [((Binding.qualify true big_name (Binding.name "induct"), dt_induct),
7ff03a5e044f theorem names generated by the (rep_)datatype command now have mandatory qualifiers
huffman
parents: 40719
diff changeset
   657
          [case_names_induct])]
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   658
      ||> Theory.checkpoint;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   659
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   660
  in
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   661
    ((constr_inject', distinct_thms', dt_induct'), thy7)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   662
  end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   663
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   664
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   665
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   666
(** datatype definition **)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   667
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   668
(* specifications *)
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   669
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   670
type spec = (binding * (string * sort) list * mixfix) * (binding * typ list * mixfix) list;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   671
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   672
type spec_cmd =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   673
  (binding * (string * string option) list * mixfix) * (binding * string list * mixfix) list;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   674
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   675
local
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   676
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   677
fun parse_spec ctxt ((b, args, mx), constrs) =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   678
  ((b, map (apsnd (Typedecl.read_constraint ctxt)) args, mx),
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   679
    constrs |> map (fn (c, Ts, mx') => (c, map (Syntax.parse_typ ctxt) Ts, mx')));
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   680
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   681
fun check_specs ctxt (specs: spec list) =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   682
  let
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   683
    fun prep_spec ((tname, args, mx), constrs) tys =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   684
      let
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   685
        val (args', tys1) = chop (length args) tys;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   686
        val (constrs', tys3) = (constrs, tys1) |-> fold_map (fn (cname, cargs, mx') => fn tys2 =>
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   687
          let val (cargs', tys3) = chop (length cargs) tys2;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   688
          in ((cname, cargs', mx'), tys3) end);
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   689
      in (((tname, map dest_TFree args', mx), constrs'), tys3) end;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   690
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   691
    val all_tys =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   692
      specs |> maps (fn ((_, args, _), cs) => map TFree args @ maps #2 cs)
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   693
      |> Syntax.check_typs ctxt;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   694
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   695
  in #1 (fold_map prep_spec specs all_tys) end;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   696
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   697
fun prep_specs parse raw_specs thy =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   698
  let
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   699
    val ctxt = thy
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   700
      |> Theory.copy
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   701
      |> Sign.add_types_global (map (fn ((b, args, mx), _) => (b, length args, mx)) raw_specs)
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   702
      |> Proof_Context.init_global
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   703
      |> fold (fn ((_, args, _), _) => fold (fn (a, _) =>
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   704
          Variable.declare_typ (TFree (a, dummyS))) args) raw_specs;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   705
    val specs = check_specs ctxt (map (parse ctxt) raw_specs);
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   706
  in (specs, ctxt) end;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   707
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   708
in
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   709
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   710
val read_specs = prep_specs parse_spec;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   711
val check_specs = prep_specs (K I);
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   712
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   713
end;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   714
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   715
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   716
(* main commands *)
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   717
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   718
fun gen_add_datatype prep_specs config raw_specs thy =
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   719
  let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   720
    val _ = Theory.requires thy "Datatype" "datatype definitions";
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   721
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   722
    val (dts, spec_ctxt) = prep_specs raw_specs thy;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   723
    val ((_, tyvars, _), _) :: _ = dts;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   724
    val string_of_tyvar = Syntax.string_of_typ spec_ctxt o TFree;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   725
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   726
    val (new_dts, types_syntax) = dts |> map (fn ((tname, tvs, mx), _) =>
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   727
      let val full_tname = Sign.full_name thy tname in
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   728
        (case duplicates (op =) tvs of
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   729
          [] =>
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   730
            if eq_set (op =) (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   731
            else error "Mutually recursive datatypes must have same type parameters"
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   732
        | dups =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   733
            error ("Duplicate parameter(s) for datatype " ^ Binding.print tname ^
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   734
              " : " ^ commas (map string_of_tyvar dups)))
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   735
      end) |> split_list;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   736
    val dt_names = map fst new_dts;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   737
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   738
    val _ =
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   739
      (case duplicates (op =) (map fst new_dts) of
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   740
        [] => ()
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   741
      | dups => error ("Duplicate datatypes: " ^ commas_quote dups));
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   742
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   743
    fun prep_dt_spec ((tname, tvs, mx), constrs) (dts', constr_syntax, i) =
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   744
      let
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   745
        fun prep_constr (cname, cargs, mx') (constrs, constr_syntax') =
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   746
          let
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   747
            val _ =
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   748
              (case subtract (op =) tvs (fold Term.add_tfreesT cargs []) of
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   749
                [] => ()
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   750
              | vs => error ("Extra type variables on rhs: " ^ commas (map string_of_tyvar vs)));
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   751
            val c = Sign.full_name_path thy (Binding.name_of tname) cname;
35129
ed24ba6f69aa discontinued unnamed infix syntax;
wenzelm
parents: 35021
diff changeset
   752
          in
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   753
            (constrs @ [(c, map (Datatype_Aux.dtyp_of_typ new_dts) cargs)],
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   754
              constr_syntax' @ [(cname, mx')])
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   755
          end handle ERROR msg =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   756
            cat_error msg ("The error above occurred in constructor " ^ Binding.print cname ^
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   757
              " of datatype " ^ Binding.print tname);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   758
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   759
        val (constrs', constr_syntax') = fold prep_constr constrs ([], []);
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   760
      in
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   761
        (case duplicates (op =) (map fst constrs') of
35129
ed24ba6f69aa discontinued unnamed infix syntax;
wenzelm
parents: 35021
diff changeset
   762
          [] =>
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   763
            (dts' @ [(i, (Sign.full_name thy tname, map Datatype_Aux.DtTFree tvs, constrs'))],
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   764
              constr_syntax @ [constr_syntax'], i + 1)
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42375
diff changeset
   765
        | dups =>
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   766
            error ("Duplicate constructors " ^ commas_quote dups ^
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   767
              " in datatype " ^ Binding.print tname))
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   768
      end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   769
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   770
    val (dts', constr_syntax, i) = fold prep_dt_spec dts ([], [], 0);
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   771
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   772
    val dt_info = Datatype_Data.get_all thy;
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   773
    val (descr, _) = Datatype_Aux.unfold_datatypes spec_ctxt dts' dt_info dts' i;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   774
    val _ =
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   775
      Datatype_Aux.check_nonempty descr
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   776
        handle (exn as Datatype_Aux.Datatype_Empty s) =>
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   777
          if #strict config then error ("Nonemptiness check failed for datatype " ^ quote s)
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   778
          else reraise exn;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   779
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   780
    val _ =
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   781
      Datatype_Aux.message config
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   782
        ("Constructing datatype(s) " ^ commas_quote (map (Binding.name_of o #1 o #1) dts));
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   783
  in
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   784
    thy
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   785
    |> representation_proofs config dt_info descr types_syntax constr_syntax
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   786
      (Datatype_Data.mk_case_names_induct (flat descr))
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   787
    |-> (fn (inject, distinct, induct) =>
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   788
      Datatype_Data.derive_datatype_props config dt_names descr induct inject distinct)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   789
  end;
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   790
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   791
val add_datatype = gen_add_datatype check_specs;
45863
afdb92130f5a tuned signature;
wenzelm
parents: 45839
diff changeset
   792
val add_datatype_cmd = gen_add_datatype read_specs;
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   793
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   794
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   795
(* outer syntax *)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   796
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   797
val spec_cmd =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   798
  Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   799
  (Parse.$$$ "=" |-- Parse.enum1 "|" (Parse.binding -- Scan.repeat Parse.typ -- Parse.opt_mixfix))
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   800
  >> (fn (((vs, t), mx), cons) => ((t, vs, mx), map Parse.triple1 cons));
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   801
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   802
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36945
diff changeset
   803
  Outer_Syntax.command "datatype" "define inductive datatypes" Keyword.thy_decl
45863
afdb92130f5a tuned signature;
wenzelm
parents: 45839
diff changeset
   804
    (Parse.and_list1 spec_cmd
afdb92130f5a tuned signature;
wenzelm
parents: 45839
diff changeset
   805
      >> (Toplevel.theory o (snd oo add_datatype_cmd Datatype_Aux.default_config)));
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   806
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   807
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   808
open Datatype_Data;
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   809
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   810
end;