src/HOL/Tools/Datatype/datatype.ML
author wenzelm
Sat, 17 Dec 2011 15:09:11 +0100
changeset 45910 566c34b64f70
parent 45909 6fe61da4c467
child 46218 ecf6375e2abb
permissions -rw-r--r--
eliminated Drule.export_without_context which is not really required here;
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
45909
6fe61da4c467 tuned signature;
wenzelm
parents: 45901
diff changeset
    13
  val distinct_lemma: thm
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    14
  type spec =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    15
    (binding * (string * sort) list * mixfix) *
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    16
    (binding * typ list * mixfix) list
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    17
  type spec_cmd =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    18
    (binding * (string * string option) list * mixfix) *
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    19
    (binding * string list * mixfix) list
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    20
  val read_specs: spec_cmd list -> theory -> spec list * Proof.context
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    21
  val check_specs: spec list -> theory -> spec list * Proof.context
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
    22
  val add_datatype: config -> spec list -> theory -> string list * theory
45863
afdb92130f5a tuned signature;
wenzelm
parents: 45839
diff changeset
    23
  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
    24
  val spec_cmd: spec_cmd parser
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    25
end;
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    26
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
    27
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
    28
struct
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    29
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
    30
(** 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
    31
45909
6fe61da4c467 tuned signature;
wenzelm
parents: 45901
diff changeset
    32
val distinct_lemma = @{lemma "f x \<noteq> f y ==> x \<noteq> y" by iprover};
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
    33
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
    34
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    35
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
    36
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
    37
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
    38
  #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
    39
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 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
    41
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
    42
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
    43
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
    44
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
    45
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
    46
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
    47
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
    48
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
    49
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
    50
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
    51
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
    52
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
    53
40719
acb830207103 keep private things private, without comments;
wenzelm
parents: 40237
diff changeset
    54
val datatype_injI =
acb830207103 keep private things private, without comments;
wenzelm
parents: 40237
diff changeset
    55
  @{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
    56
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
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
    58
(** 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
    59
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
    60
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
    61
    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
    62
  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
    63
    val descr' = flat descr;
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
    64
    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
    65
    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
    66
    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
    67
    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
    68
    val rep_set_names' =
45743
857b7fcb0365 misc tuning;
wenzelm
parents: 45742
diff changeset
    69
      if length descr' = 1 then [big_rec_name]
45821
wenzelm
parents: 45743
diff changeset
    70
      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
    71
    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
    72
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
    73
    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
    74
    val leafTs' = Datatype_Aux.get_nonrec_types descr';
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
    75
    val branchTs = Datatype_Aux.get_branching_types descr';
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
    76
    val branchT =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
    77
      if null branchTs then HOLogic.unitT
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 36960
diff changeset
    78
      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
    79
    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
    80
    val unneeded_vars =
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
    81
      subtract (op =) (fold Term.add_tfreesT (leafTs' @ branchTs) []) (hd tyvars);
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
    82
    val leafTs = leafTs' @ map TFree unneeded_vars;
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
    83
    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
    84
    val (newTs, oldTs) = chop (length (hd descr)) recTs;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
    85
    val sumT =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
    86
      if null leafTs then HOLogic.unitT
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 36960
diff changeset
    87
      else Balanced_Tree.make (fn (T, U) => Type (@{type_name Sum_Type.sum}, [T, U])) leafTs;
45878
wenzelm
parents: 45863
diff changeset
    88
    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
    89
    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
    90
    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
    91
    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
    92
45878
wenzelm
parents: 45863
diff changeset
    93
    val In0 = Const (@{const_name Datatype.In0}, Univ_elT --> Univ_elT);
wenzelm
parents: 45863
diff changeset
    94
    val In1 = Const (@{const_name Datatype.In1}, Univ_elT --> Univ_elT);
wenzelm
parents: 45863
diff changeset
    95
    val Leaf = Const (@{const_name Datatype.Leaf}, sumT --> Univ_elT);
wenzelm
parents: 45863
diff changeset
    96
    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
    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
    (* 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
    99
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' 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
   101
      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
   102
        fun mk_inj' T n i =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   103
          if n = 1 then x
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   104
          else
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   105
            let
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   106
              val n2 = n div 2;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   107
              val Type (_, [T1, T2]) = T;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   108
            in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   109
              if i <= n2
45821
wenzelm
parents: 45743
diff changeset
   110
              then Const (@{const_name Inl}, T1 --> T) $ mk_inj' T1 n2 i
wenzelm
parents: 45743
diff changeset
   111
              else Const (@{const_name Inr}, T2 --> T) $ mk_inj' T2 (n - n2) (i - n2)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   112
            end;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   113
      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
   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
    (* 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
   116
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
    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
   118
      {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
   119
        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
   120
        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
   121
          if ts = [] then Const (@{const_name undefined}, Univ_elT)
45878
wenzelm
parents: 45863
diff changeset
   122
          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
   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
    (* 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
   125
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_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
   127
      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
   128
        fun mk_inj T n i =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   129
          if n = 1 then x
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   130
          else
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   131
            let
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   132
              val n2 = n div 2;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   133
              val Type (_, [T1, T2]) = T;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   134
              fun mkT U = (U --> Univ_elT) --> T --> Univ_elT;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   135
            in
45878
wenzelm
parents: 45863
diff changeset
   136
              if i <= n2 then Const (@{const_name Sum_Type.Suml}, mkT T1) $ mk_inj T1 n2 i
wenzelm
parents: 45863
diff changeset
   137
              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
   138
            end;
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   139
      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
   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
    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
   142
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   143
    (************** 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
   144
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   145
    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
   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
    (* 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
   148
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 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
   150
      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
   151
        fun mk_prem dt (j, prems, ts) =
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   152
          (case Datatype_Aux.strip_dtyp dt of
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   153
            (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
   154
              let
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   155
                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
   156
                val free_t =
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   157
                  Datatype_Aux.app_bnds (Datatype_Aux.mk_Free "x" (Ts ---> Univ_elT) j) (length Ts)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   158
              in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   159
                (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
   160
                  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
   161
                    (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
   162
                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
   163
              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
   164
          | _ =>
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   165
              let val T = Datatype_Aux.typ_of_dtyp descr' dt
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   166
              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
   167
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   168
        val (_, prems, ts) = fold_rev mk_prem cargs (1, [], []);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   169
        val concl = HOLogic.mk_Trueprop (Free (s, UnivT') $ mk_univ_inj ts n i);
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   170
      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
   171
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   172
    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
   173
      map (make_intr rep_set_name (length constrs))
45821
wenzelm
parents: 45743
diff changeset
   174
        ((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
   175
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
    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
   177
      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
   178
      |> 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
   179
      |> 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
   180
          {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
   181
           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
   182
          (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
   183
          (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
   184
      ||> 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
   185
      ||> 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
   186
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   187
    (********************************* 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
   188
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   189
    val (typedefs, thy3) = thy2
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   190
      |> Sign.parent_path
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   191
      |> fold_map
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   192
        (fn (((name, mx), tvs), c) =>
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   193
          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
   194
            (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
   195
            (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
   196
              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
   197
              (resolve_tac rep_intrs 1)))
45878
wenzelm
parents: 45863
diff changeset
   198
        (types_syntax ~~ tyvars ~~ take (length newTs) rep_set_names)
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   199
      ||> 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
   200
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   201
    (*********************** 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
   202
45878
wenzelm
parents: 45863
diff changeset
   203
    val big_rep_name = big_name ^ "_Rep_";
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   204
    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
   205
    val all_rep_names =
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   206
      map (#Rep_name o #1 o #2) typedefs @
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
   207
      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
   208
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
    (* 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
   210
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
    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
   212
      (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
   213
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
    (* 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
   215
45889
4ff377493dbb misc tuning and simplification;
wenzelm
parents: 45879
diff changeset
   216
    fun make_constr_def (typedef: Typedef.info) T n
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   217
        ((cname, cargs), (cname', mx)) (thy, defs, eqns, 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
   218
      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
   219
        fun constr_arg dt (j, l_args, r_args) =
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   220
          let
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   221
            val T = Datatype_Aux.typ_of_dtyp descr' dt;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   222
            val free_t = Datatype_Aux.mk_Free "x" T j;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   223
          in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   224
            (case (Datatype_Aux.strip_dtyp dt, strip_type T) of
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   225
              ((_, Datatype_Aux.DtRec m), (Us, U)) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   226
                (j + 1, free_t :: l_args, mk_lim
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   227
                  (Const (nth all_rep_names m, U --> Univ_elT) $
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   228
                    Datatype_Aux.app_bnds free_t (length Us)) Us :: r_args)
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   229
            | _ => (j + 1, free_t :: l_args, (Leaf $ mk_inj T free_t) :: 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
   230
          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
   231
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 (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []);
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   233
        val constrT = map (Datatype_Aux.typ_of_dtyp descr') cargs ---> T;
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   234
        val ({Abs_name, Rep_name, ...}, _) = typedef;
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
   235
        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
   236
        val rhs = mk_univ_inj r_args n i;
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   237
        val def = Logic.mk_equals (lhs, Const (Abs_name, Univ_elT --> T) $ 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
   238
        val def_name = Long_Name.base_name cname ^ "_def";
45821
wenzelm
parents: 45743
diff changeset
   239
        val eqn =
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   240
          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
   241
        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
   242
          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
   243
          |> 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
   244
          |> (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
   245
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
      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
   247
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
    (* 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
   249
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   250
    fun dt_constr_defs (((((_, (_, _, constrs)), tname), typedef: Typedef.info), T), 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
   251
        (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
   252
      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
   253
        val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   254
        val rep_const = cterm_of thy (Const (#Rep_name (#1 typedef), T --> Univ_elT));
45910
566c34b64f70 eliminated Drule.export_without_context which is not really required here;
wenzelm
parents: 45909
diff changeset
   255
        val cong' = cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong;
566c34b64f70 eliminated Drule.export_without_context which is not really required here;
wenzelm
parents: 45909
diff changeset
   256
        val dist = cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   257
        val (thy', defs', eqns', _) =
45889
4ff377493dbb misc tuning and simplification;
wenzelm
parents: 45879
diff changeset
   258
          fold (make_constr_def typedef T (length constrs))
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   259
            (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
   260
      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
   261
        (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
   262
          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
   263
      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
   264
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
    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
   266
      fold dt_constr_defs
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   267
        (hd descr ~~ new_type_names ~~ map #2 typedefs ~~ newTs ~~ 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
   268
        (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
   269
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
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
    (*********** 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
   272
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   273
    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
   274
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   275
    val newT_iso_axms = typedefs |> map (fn (_, (_, {Abs_inverse, Rep_inverse, Rep, ...})) =>
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   276
      (collect_simp Abs_inverse, Rep_inverse, collect_simp Rep));
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
   277
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   278
    val newT_iso_inj_thms = typedefs |> map (fn (_, (_, {Abs_inject, Rep_inject, ...})) =>
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   279
      (collect_simp Abs_inject RS iffD1, Rep_inject RS iffD1));
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
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
    (********* 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
   282
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   283
    (*---------------------------------------------------------------------*)
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
    (* 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
   285
    (* 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
   286
    (*                                                                     *)
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
    (*   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
   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
    (* 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
   290
    (*                                                                     *)
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
    (*   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
   292
    (*---------------------------------------------------------------------*)
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
    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
   295
      let
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   296
        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
   297
        val T = nth recTs k;
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   298
        val rep_const = Const (nth all_rep_names k, 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
   299
        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
   300
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   301
        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
   302
          let
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   303
            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
   304
            val (Us, U) = strip_type T'
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   305
          in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   306
            (case Datatype_Aux.strip_dtyp dt of
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   307
              (_, Datatype_Aux.DtRec j) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   308
                if member (op =) ks' j then
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   309
                  (i2 + 1, i2' + 1, ts @ [mk_lim (Datatype_Aux.app_bnds
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   310
                     (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
   311
                   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
   312
                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
   313
                  (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
   314
                     (Const (nth all_rep_names j, U --> Univ_elT) $
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   315
                        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
   316
            | _ => (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
   317
          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
   318
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
        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
   320
        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
   321
        val ys = map (uncurry (Datatype_Aux.mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
45878
wenzelm
parents: 45863
diff changeset
   322
        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
   323
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
        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
   325
        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
   326
          (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
   327
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   328
      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
   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
    (* 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
   331
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
    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
   333
      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
   334
        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
   335
        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
   336
        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
   337
45878
wenzelm
parents: 45863
diff changeset
   338
        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
   339
          let
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   340
            val (fs', eqns', _) = fold (make_iso_def k ks (length constrs)) constrs (fs, eqns, 1);
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   341
            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
   342
          in (fs', eqns', isos @ [iso]) end;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   343
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
   344
        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
   345
        val fTs = map fastype_of fs;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   346
        val defs =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   347
          map (fn (rec_name, (T, iso_name)) =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   348
            (Binding.name (Long_Name.base_name iso_name ^ "_def"),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   349
              Logic.mk_equals (Const (iso_name, T --> Univ_elT),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   350
                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
   351
        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
   352
          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
   353
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   354
        (* 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
   355
45878
wenzelm
parents: 45863
diff changeset
   356
        val rewrites = def_thms @ map mk_meta_eq rec_rewrites;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   357
        val char_thms' =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   358
          map (fn eqn => Skip_Proof.prove_global thy' [] [] eqn
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   359
            (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
   360
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
      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
   362
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   363
    val (thy5, iso_char_thms) =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   364
      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
   365
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
    (* 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
   367
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
    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
   369
      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
   370
        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
   371
        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
   372
          (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.legacy_freeze prop;
45738
0430f9123e43 eliminated some legacy operations;
wenzelm
parents: 45735
diff changeset
   373
        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
   374
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
        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
   376
          let
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   377
            val Ts = map (TFree o rpair HOLogic.typeS) (Name.variant_list used (replicate i "'t"));
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   378
            val f = Free ("f", Ts ---> U);
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   379
          in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   380
            Skip_Proof.prove_global thy [] []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   381
              (Logic.mk_implies
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   382
                (HOLogic.mk_Trueprop (HOLogic.list_all
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   383
                   (map (pair "x") Ts, S $ Datatype_Aux.app_bnds f i)),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   384
                 HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   385
                   r $ (a $ Datatype_Aux.app_bnds f i)), f))))
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   386
              (fn _ => EVERY [REPEAT_DETERM_N i (rtac ext 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   387
                 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
   388
          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
   389
      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
   390
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   391
    (* 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
   392
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   393
    val fun_congs =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   394
      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
   395
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
    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
   397
      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
   398
        val (_, (tname, _, _)) = hd ds;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   399
        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
   400
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
        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
   402
          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
   403
            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
   404
            val Rep_t = Const (nth all_rep_names i, T --> Univ_elT);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   405
            val rep_set_name = nth rep_set_names i;
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   406
            val concl1 =
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   407
              HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   408
                HOLogic.mk_eq (Rep_t $ Datatype_Aux.mk_Free "x" T i, Rep_t $ Bound 0) $
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   409
                  HOLogic.mk_eq (Datatype_Aux.mk_Free "x" T i, Bound 0));
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   410
            val concl2 = Const (rep_set_name, UnivT') $ (Rep_t $ Datatype_Aux.mk_Free "x" T i);
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   411
          in (concl1, concl2) 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
   412
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   413
        val (ind_concl1, ind_concl2) = split_list (map mk_ind_concl ds);
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
   414
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   415
        val rewrites = map mk_meta_eq iso_char_thms;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   416
        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
   417
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   418
        val inj_thm =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   419
          Skip_Proof.prove_global thy5 [] []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   420
            (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj ind_concl1))
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   421
            (fn _ => EVERY
45735
7d7d7af647a9 tuned signature;
wenzelm
parents: 45701
diff changeset
   422
              [(Datatype_Aux.ind_tac induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1,
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   423
               REPEAT (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   424
                 [rtac allI 1, rtac impI 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   425
                  Datatype_Aux.exh_tac (exh_thm_of dt_info) 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   426
                  REPEAT (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   427
                    [hyp_subst_tac 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   428
                     rewrite_goals_tac rewrites,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   429
                     REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   430
                     (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   431
                     ORELSE (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   432
                       [REPEAT (eresolve_tac (Scons_inject ::
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   433
                          map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   434
                        REPEAT (cong_tac 1), rtac refl 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   435
                        REPEAT (atac 1 ORELSE (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   436
                          [REPEAT (rtac ext 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   437
                           REPEAT (eresolve_tac (mp :: allE ::
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   438
                             map make_elim (Suml_inject :: Sumr_inject ::
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   439
                               Lim_inject :: inj_thms') @ fun_congs) 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   440
                           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
   441
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   442
        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
   443
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   444
        val elem_thm =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   445
          Skip_Proof.prove_global thy5 [] []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   446
            (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj ind_concl2))
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   447
            (fn _ =>
45735
7d7d7af647a9 tuned signature;
wenzelm
parents: 45701
diff changeset
   448
              EVERY [(Datatype_Aux.ind_tac induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1,
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   449
                rewrite_goals_tac rewrites,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   450
                REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   451
                  ((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
   452
45878
wenzelm
parents: 45863
diff changeset
   453
      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
   454
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   455
    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
   456
      fold_rev prove_iso_thms (tl descr) ([], map #3 newT_iso_axms);
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   457
    val iso_inj_thms =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   458
      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
   459
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
    (* 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
   461
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
    fun mk_iso_t (((set_name, iso_name), i), T) =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   463
      let val isoT = T --> Univ_elT in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   464
        HOLogic.imp $
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   465
          (Const (set_name, UnivT') $ Datatype_Aux.mk_Free "x" Univ_elT i) $
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45738
diff changeset
   466
            (if i < length newTs then @{term True}
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   467
             else HOLogic.mk_mem (Datatype_Aux.mk_Free "x" Univ_elT i,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   468
               Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   469
                 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
   470
      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
   471
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   472
    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
   473
      (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
   474
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   475
    (* 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
   476
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   477
    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
   478
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   479
    val iso_thms =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   480
      if length descr = 1 then []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   481
      else
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   482
        drop (length newTs) (Datatype_Aux.split_conj_thm
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   483
          (Skip_Proof.prove_global thy5 [] [] iso_t (fn _ => EVERY
45735
7d7d7af647a9 tuned signature;
wenzelm
parents: 45701
diff changeset
   484
             [(Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1,
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   485
              REPEAT (rtac TrueI 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   486
              rewrite_goals_tac (mk_meta_eq @{thm choice_eq} ::
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   487
                Thm.symmetric (mk_meta_eq @{thm fun_eq_iff}) :: range_eqs),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   488
              rewrite_goals_tac (map Thm.symmetric range_eqs),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   489
              REPEAT (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   490
                [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   491
                   maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   492
                 TRY (hyp_subst_tac 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   493
                 rtac (sym RS range_eqI) 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   494
                 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
   495
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   496
    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
   497
      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
   498
      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
   499
        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
   500
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
    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
   502
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
    (******************* 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
   504
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   505
    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
   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
    (* prove theorem  Rep_i (Constr_j ...) = Inj_j ...  *)
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   508
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
   509
    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
   510
      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
   511
        val inj_thms = map fst newT_iso_inj_thms;
45878
wenzelm
parents: 45863
diff changeset
   512
        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
   513
      in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   514
        Skip_Proof.prove_global thy5 [] [] eqn
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   515
        (fn _ => EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   516
          [resolve_tac inj_thms 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   517
           rewrite_goals_tac rewrites,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   518
           rtac refl 3,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   519
           resolve_tac rep_intrs 2,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   520
           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
   521
      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
   522
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   523
    (*--------------------------------------------------------------*)
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   524
    (* 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
   525
    (* 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
   526
    (*--------------------------------------------------------------*)
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
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
    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
   529
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   530
    val dist_rewrites =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   531
      map (fn (rep_thms, dist_lemma) =>
45878
wenzelm
parents: 45863
diff changeset
   532
        dist_lemma :: (rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   533
          (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
   534
45889
4ff377493dbb misc tuning and simplification;
wenzelm
parents: 45879
diff changeset
   535
    fun prove_distinct_thms dist_rewrites' =
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
   536
      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
   537
        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
   538
          | 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
   539
              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
   540
                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
   541
                  EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1])
45910
566c34b64f70 eliminated Drule.export_without_context which is not really required here;
wenzelm
parents: 45909
diff changeset
   542
              in dist_thm :: Drule.zero_var_indexes (dist_thm RS not_sym) :: prove ts end;
45889
4ff377493dbb misc tuning and simplification;
wenzelm
parents: 45879
diff changeset
   543
      in prove 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
   544
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   545
    val distinct_thms =
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   546
      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
   547
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   548
    (* 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
   549
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
    fun prove_constr_inj_thm rep_thms t =
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   551
      let
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   552
        val inj_thms = Scons_inject ::
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   553
          map make_elim
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   554
            (iso_inj_thms @
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   555
              [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   556
               Lim_inject, Suml_inject, Sumr_inject])
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   557
      in
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   558
        Skip_Proof.prove_global thy5 [] [] t
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   559
          (fn _ => EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   560
            [rtac iffI 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   561
             REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   562
             dresolve_tac rep_congs 1, dtac box_equals 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   563
             REPEAT (resolve_tac rep_thms 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   564
             REPEAT (eresolve_tac inj_thms 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   565
             REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   566
               REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   567
               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
   568
      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
   569
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   570
    val constr_inject =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   571
      map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts)
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   572
        (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
   573
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
    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
   575
      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
   576
      |> Sign.parent_path
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   577
      |> Datatype_Aux.store_thmss "inject" new_type_names constr_inject
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   578
      ||>> 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
   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
    (*************************** 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
   581
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   582
    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
   583
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   584
    val Rep_inverse_thms =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   585
      map (fn (_, iso, _) => iso RS subst) newT_iso_axms @
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   586
      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
   587
    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
   588
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   589
    fun mk_indrule_lemma (i, _) 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
   590
      let
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   591
        val Rep_t = Const (nth all_rep_names i, T --> Univ_elT) $ Datatype_Aux.mk_Free "x" T i;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   592
        val Abs_t =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   593
          if i < length newTs then
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   594
            Const (#Abs_name (#1 (#2 (nth typedefs i))), Univ_elT --> T)
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   595
          else
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   596
            Const (@{const_name the_inv_into},
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
   597
              [HOLogic.mk_setT T, T --> Univ_elT, Univ_elT] ---> T) $
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   598
            HOLogic.mk_UNIV T $ Const (nth all_rep_names i, T --> Univ_elT);
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   599
        val prem =
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   600
          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
   601
            (Const (nth rep_set_names i, UnivT') $ Rep_t) $
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   602
              (Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t));
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   603
        val concl =
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   604
          Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ Datatype_Aux.mk_Free "x" T i;
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   605
      in (prem, 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
   606
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   607
    val (indrule_lemma_prems, indrule_lemma_concls) =
45879
71b8d0d170b1 avoid fragile Sign.intern_const -- pass internal names directly;
wenzelm
parents: 45878
diff changeset
   608
      split_list (map2 mk_indrule_lemma descr' recTs);
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
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   610
    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
   611
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   612
    val indrule_lemma =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   613
      Skip_Proof.prove_global thy6 [] []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   614
        (Logic.mk_implies
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   615
          (HOLogic.mk_Trueprop (Datatype_Aux.mk_conj indrule_lemma_prems),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   616
           HOLogic.mk_Trueprop (Datatype_Aux.mk_conj indrule_lemma_concls)))
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   617
        (fn _ =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   618
          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
   619
           [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
   620
            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
   621
              [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
   622
               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
   623
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   624
    val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   625
    val frees =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   626
      if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))]
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   627
      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
   628
    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
   629
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   630
    val dt_induct_prop = Datatype_Prop.make_ind descr;
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   631
    val dt_induct =
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   632
      Skip_Proof.prove_global thy6 []
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   633
      (Logic.strip_imp_prems dt_induct_prop)
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   634
      (Logic.strip_imp_concl dt_induct_prop)
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   635
      (fn {prems, ...} =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   636
        EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   637
          [rtac indrule_lemma' 1,
45735
7d7d7af647a9 tuned signature;
wenzelm
parents: 45701
diff changeset
   638
           (Datatype_Aux.ind_tac rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac) 1,
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   639
           EVERY (map (fn (prem, r) => (EVERY
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   640
             [REPEAT (eresolve_tac Abs_inverse_thms 1),
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   641
              simp_tac (HOL_basic_ss addsimps (Thm.symmetric r :: Rep_inverse_thms')) 1,
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   642
              DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
45878
wenzelm
parents: 45863
diff changeset
   643
                  (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
   644
45901
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45890
diff changeset
   645
    val ([(_, [dt_induct'])], thy7) =
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
   646
      thy6
45901
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45890
diff changeset
   647
      |> Global_Theory.note_thmss ""
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45890
diff changeset
   648
        [((Binding.qualify true big_name (Binding.name "induct"), [case_names_induct]),
cea7cd0c7e99 eliminated old-fashioned Global_Theory.add_thms(s);
wenzelm
parents: 45890
diff changeset
   649
          [([dt_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
   650
      ||> 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
   651
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
  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
   653
    ((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
   654
  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
   655
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   656
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   657
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   658
(** 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
   659
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   660
(* specifications *)
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   661
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   662
type spec = (binding * (string * sort) list * mixfix) * (binding * typ list * mixfix) list;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   663
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   664
type spec_cmd =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   665
  (binding * (string * string option) list * mixfix) * (binding * string list * mixfix) list;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   666
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   667
local
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   668
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   669
fun parse_spec ctxt ((b, args, mx), constrs) =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   670
  ((b, map (apsnd (Typedecl.read_constraint ctxt)) args, mx),
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   671
    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
   672
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   673
fun check_specs ctxt (specs: spec list) =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   674
  let
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   675
    fun prep_spec ((tname, args, mx), constrs) tys =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   676
      let
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   677
        val (args', tys1) = chop (length args) tys;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   678
        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
   679
          let val (cargs', tys3) = chop (length cargs) tys2;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   680
          in ((cname, cargs', mx'), tys3) end);
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   681
      in (((tname, map dest_TFree args', mx), constrs'), tys3) end;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   682
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   683
    val all_tys =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   684
      specs |> maps (fn ((_, args, _), cs) => map TFree args @ maps #2 cs)
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   685
      |> Syntax.check_typs ctxt;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   686
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   687
  in #1 (fold_map prep_spec specs all_tys) end;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   688
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   689
fun prep_specs parse raw_specs thy =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   690
  let
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   691
    val ctxt = thy
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   692
      |> Theory.copy
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   693
      |> 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
   694
      |> Proof_Context.init_global
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   695
      |> fold (fn ((_, args, _), _) => fold (fn (a, _) =>
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   696
          Variable.declare_typ (TFree (a, dummyS))) args) raw_specs;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   697
    val specs = check_specs ctxt (map (parse ctxt) raw_specs);
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   698
  in (specs, ctxt) end;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   699
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   700
in
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   701
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   702
val read_specs = prep_specs parse_spec;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   703
val check_specs = prep_specs (K I);
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   704
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   705
end;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   706
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
(* main commands *)
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
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
   711
  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
   712
    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
   713
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   714
    val (dts, spec_ctxt) = prep_specs raw_specs thy;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   715
    val ((_, tyvars, _), _) :: _ = dts;
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   716
    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
   717
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   718
    val (new_dts, types_syntax) = dts |> map (fn ((tname, tvs, mx), _) =>
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   719
      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
   720
        (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
   721
          [] =>
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   722
            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
   723
            else error "Mutually recursive datatypes must have same type parameters"
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   724
        | dups =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   725
            error ("Duplicate parameter(s) for datatype " ^ Binding.print tname ^
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   726
              " : " ^ commas (map string_of_tyvar dups)))
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   727
      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
   728
    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
   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
    val _ =
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   731
      (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
   732
        [] => ()
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   733
      | 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
   734
45889
4ff377493dbb misc tuning and simplification;
wenzelm
parents: 45879
diff changeset
   735
    fun prep_dt_spec ((tname, tvs, _), 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
   736
      let
45889
4ff377493dbb misc tuning and simplification;
wenzelm
parents: 45879
diff changeset
   737
        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
   738
          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
   739
            val _ =
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   740
              (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
   741
                [] => ()
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   742
              | 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
   743
            val c = Sign.full_name_path thy (Binding.name_of tname) cname;
35129
ed24ba6f69aa discontinued unnamed infix syntax;
wenzelm
parents: 35021
diff changeset
   744
          in
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   745
            (constrs @ [(c, map (Datatype_Aux.dtyp_of_typ new_dts) cargs)],
45889
4ff377493dbb misc tuning and simplification;
wenzelm
parents: 45879
diff changeset
   746
              constr_syntax' @ [(cname, mx)])
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   747
          end handle ERROR msg =>
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   748
            cat_error msg ("The error above occurred in constructor " ^ Binding.print cname ^
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   749
              " 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
   750
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   751
        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
   752
      in
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   753
        (case duplicates (op =) (map fst constrs') of
35129
ed24ba6f69aa discontinued unnamed infix syntax;
wenzelm
parents: 35021
diff changeset
   754
          [] =>
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   755
            (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
   756
              constr_syntax @ [constr_syntax'], i + 1)
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42375
diff changeset
   757
        | dups =>
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   758
            error ("Duplicate constructors " ^ commas_quote dups ^
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   759
              " 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
   760
      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
   761
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   762
    val (dts', constr_syntax, i) = fold prep_dt_spec dts ([], [], 0);
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   763
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
   764
    val dt_info = Datatype_Data.get_all thy;
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   765
    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
   766
    val _ =
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   767
      Datatype_Aux.check_nonempty descr
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   768
        handle (exn as Datatype_Aux.Datatype_Empty s) =>
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   769
          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
   770
          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
   771
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   772
    val _ =
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   773
      Datatype_Aux.message config
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   774
        ("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
   775
  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
   776
    thy
45822
843dc212f69e datatype dtyp with explicit sort information;
wenzelm
parents: 45821
diff changeset
   777
    |> representation_proofs config dt_info descr types_syntax constr_syntax
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   778
      (Datatype_Data.mk_case_names_induct (flat descr))
45700
9dcbf6a1829c misc tuning;
wenzelm
parents: 45229
diff changeset
   779
    |-> (fn (inject, distinct, induct) =>
45890
5f70aaecae26 separate rep_datatype.ML;
wenzelm
parents: 45889
diff changeset
   780
      Rep_Datatype.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
   781
  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
   782
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   783
val add_datatype = gen_add_datatype check_specs;
45863
afdb92130f5a tuned signature;
wenzelm
parents: 45839
diff changeset
   784
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
   785
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   786
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   787
(* 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
   788
45839
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   789
val spec_cmd =
43a5b86bc102 'datatype' specifications allow explicit sort constraints;
wenzelm
parents: 45822
diff changeset
   790
  Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
45701
615da8b8d758 discontinued obsolete datatype "alt_names";
wenzelm
parents: 45700
diff changeset
   791
  (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
   792
  >> (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
   793
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   794
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36945
diff changeset
   795
  Outer_Syntax.command "datatype" "define inductive datatypes" Keyword.thy_decl
45863
afdb92130f5a tuned signature;
wenzelm
parents: 45839
diff changeset
   796
    (Parse.and_list1 spec_cmd
afdb92130f5a tuned signature;
wenzelm
parents: 45839
diff changeset
   797
      >> (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
   798
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   799
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   800
open Datatype_Data;
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40929
diff changeset
   801
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
   802
end;