src/HOL/Tools/Datatype/datatype_data.ML
author wenzelm
Wed, 23 Nov 2011 22:59:39 +0100
changeset 45620 f2a587696afb
parent 45430 b8eb7a791dac
child 45700 9dcbf6a1829c
permissions -rw-r--r--
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33970
haftmann
parents: 33969
diff changeset
     1
(*  Title:      HOL/Tools/Datatype/datatype_data.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
33970
haftmann
parents: 33969
diff changeset
     4
Datatype package: bookkeeping; interpretation of existing types as datatypes.
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     5
*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     6
33970
haftmann
parents: 33969
diff changeset
     7
signature DATATYPE_DATA =
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
     8
sig
33970
haftmann
parents: 33969
diff changeset
     9
  include DATATYPE_COMMON
haftmann
parents: 33969
diff changeset
    10
  val derive_datatype_props : config -> string list -> string list option
haftmann
parents: 33969
diff changeset
    11
    -> descr list -> (string * sort) list -> thm -> thm list list -> thm list list
haftmann
parents: 33969
diff changeset
    12
    -> theory -> string list * theory
haftmann
parents: 33969
diff changeset
    13
  val rep_datatype : config -> (string list -> Proof.context -> Proof.context)
haftmann
parents: 33969
diff changeset
    14
    -> string list option -> term list -> theory -> Proof.state
haftmann
parents: 33969
diff changeset
    15
  val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state
haftmann
parents: 33969
diff changeset
    16
  val get_info : theory -> string -> info option
haftmann
parents: 33969
diff changeset
    17
  val the_info : theory -> string -> info
haftmann
parents: 33969
diff changeset
    18
  val the_descr : theory -> string list
haftmann
parents: 33969
diff changeset
    19
    -> descr * (string * sort) list * string list
haftmann
parents: 33969
diff changeset
    20
      * string * (string list * string list) * (typ list * typ list)
haftmann
parents: 33969
diff changeset
    21
  val the_spec : theory -> string -> (string * sort) list * (string * typ list) list
haftmann
parents: 33969
diff changeset
    22
  val all_distincts : theory -> typ list -> thm list list
haftmann
parents: 33969
diff changeset
    23
  val get_constrs : theory -> string -> (string * typ) list option
haftmann
parents: 33969
diff changeset
    24
  val get_all : theory -> info Symtab.table
haftmann
parents: 33969
diff changeset
    25
  val info_of_constr : theory -> string * typ -> info option
43580
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
    26
  val info_of_constr_permissive : theory -> string * typ -> info option
33970
haftmann
parents: 33969
diff changeset
    27
  val info_of_case : theory -> string -> info option
haftmann
parents: 33969
diff changeset
    28
  val interpretation : (config -> string list -> theory -> theory) -> theory -> theory
haftmann
parents: 33969
diff changeset
    29
  val make_case :  Proof.context -> Datatype_Case.config -> string list -> term ->
43253
fa3c61dc9f2c removed generation of instantiated pattern set, which is never actually used
krauss
parents: 42361
diff changeset
    30
    (term * term) list -> term
33970
haftmann
parents: 33969
diff changeset
    31
  val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
haftmann
parents: 33969
diff changeset
    32
  val read_typ: theory -> string -> (string * sort) list -> typ * (string * sort) list
haftmann
parents: 33969
diff changeset
    33
  val cert_typ: theory -> typ -> (string * sort) list -> typ * (string * sort) list
haftmann
parents: 33969
diff changeset
    34
  val mk_case_names_induct: descr -> attribute
haftmann
parents: 33969
diff changeset
    35
  val setup: theory -> theory
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    36
end;
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    37
33970
haftmann
parents: 33969
diff changeset
    38
structure Datatype_Data: DATATYPE_DATA =
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
    39
struct
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    40
33970
haftmann
parents: 33969
diff changeset
    41
(** theory data **)
haftmann
parents: 33969
diff changeset
    42
haftmann
parents: 33969
diff changeset
    43
(* data management *)
haftmann
parents: 33969
diff changeset
    44
haftmann
parents: 33969
diff changeset
    45
structure DatatypesData = Theory_Data
haftmann
parents: 33969
diff changeset
    46
(
haftmann
parents: 33969
diff changeset
    47
  type T =
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
    48
    {types: Datatype_Aux.info Symtab.table,
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
    49
     constrs: (string * Datatype_Aux.info) list Symtab.table,
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
    50
     cases: Datatype_Aux.info Symtab.table};
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
    51
33970
haftmann
parents: 33969
diff changeset
    52
  val empty =
haftmann
parents: 33969
diff changeset
    53
    {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
haftmann
parents: 33969
diff changeset
    54
  val extend = I;
haftmann
parents: 33969
diff changeset
    55
  fun merge
haftmann
parents: 33969
diff changeset
    56
    ({types = types1, constrs = constrs1, cases = cases1},
haftmann
parents: 33969
diff changeset
    57
     {types = types2, constrs = constrs2, cases = cases2}) : T =
haftmann
parents: 33969
diff changeset
    58
    {types = Symtab.merge (K true) (types1, types2),
haftmann
parents: 33969
diff changeset
    59
     constrs = Symtab.join (K (AList.merge (op =) (K true))) (constrs1, constrs2),
haftmann
parents: 33969
diff changeset
    60
     cases = Symtab.merge (K true) (cases1, cases2)};
haftmann
parents: 33969
diff changeset
    61
);
haftmann
parents: 33969
diff changeset
    62
haftmann
parents: 33969
diff changeset
    63
val get_all = #types o DatatypesData.get;
haftmann
parents: 33969
diff changeset
    64
val get_info = Symtab.lookup o get_all;
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
33970
haftmann
parents: 33969
diff changeset
    66
fun the_info thy name =
haftmann
parents: 33969
diff changeset
    67
  (case get_info thy name of
haftmann
parents: 33969
diff changeset
    68
    SOME info => info
haftmann
parents: 33969
diff changeset
    69
  | NONE => error ("Unknown datatype " ^ quote 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
    70
33970
haftmann
parents: 33969
diff changeset
    71
fun info_of_constr thy (c, T) =
haftmann
parents: 33969
diff changeset
    72
  let
haftmann
parents: 33969
diff changeset
    73
    val tab = Symtab.lookup_list ((#constrs o DatatypesData.get) thy) c;
43580
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
    74
  in
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
    75
    case body_type T of
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
    76
      Type (tyco, _) => AList.lookup (op =) tab tyco
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
    77
    | _ => NONE
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
    78
  end;
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
    79
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
    80
fun info_of_constr_permissive thy (c, T) =
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
    81
  let
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
    82
    val tab = Symtab.lookup_list ((#constrs o DatatypesData.get) thy) c;
40844
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 39557
diff changeset
    83
    val hint = case body_type T of Type (tyco, _) => SOME tyco | _ => NONE;
41489
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41423
diff changeset
    84
    val default =
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41423
diff changeset
    85
      if null tab then NONE
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 41423
diff changeset
    86
      else SOME (snd (List.last tab))
33970
haftmann
parents: 33969
diff changeset
    87
        (*conservative wrt. overloaded constructors*);
haftmann
parents: 33969
diff changeset
    88
  in case hint
haftmann
parents: 33969
diff changeset
    89
   of NONE => default
haftmann
parents: 33969
diff changeset
    90
    | SOME tyco => case AList.lookup (op =) tab tyco
haftmann
parents: 33969
diff changeset
    91
       of NONE => default (*permissive*)
haftmann
parents: 33969
diff changeset
    92
        | SOME info => SOME info
haftmann
parents: 33969
diff changeset
    93
  end;
haftmann
parents: 33969
diff changeset
    94
haftmann
parents: 33969
diff changeset
    95
val info_of_case = Symtab.lookup o #cases o DatatypesData.get;
haftmann
parents: 33969
diff changeset
    96
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
    97
fun register (dt_infos : (string * Datatype_Aux.info) list) =
33970
haftmann
parents: 33969
diff changeset
    98
  DatatypesData.map (fn {types, constrs, cases} =>
haftmann
parents: 33969
diff changeset
    99
    {types = types |> fold Symtab.update dt_infos,
haftmann
parents: 33969
diff changeset
   100
     constrs = constrs |> fold (fn (constr, dtname_info) =>
haftmann
parents: 33969
diff changeset
   101
         Symtab.map_default (constr, []) (cons dtname_info))
haftmann
parents: 33969
diff changeset
   102
       (maps (fn (dtname, info as {descr, index, ...}) =>
haftmann
parents: 33969
diff changeset
   103
          map (rpair (dtname, info) o fst)
haftmann
parents: 33969
diff changeset
   104
            (#3 (the (AList.lookup op = descr index)))) dt_infos),
haftmann
parents: 33969
diff changeset
   105
     cases = cases |> fold Symtab.update
haftmann
parents: 33969
diff changeset
   106
       (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)});
haftmann
parents: 33969
diff changeset
   107
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
   108
33970
haftmann
parents: 33969
diff changeset
   109
(* complex queries *)
haftmann
parents: 33969
diff changeset
   110
haftmann
parents: 33969
diff changeset
   111
fun the_spec thy dtco =
haftmann
parents: 33969
diff changeset
   112
  let
haftmann
parents: 33969
diff changeset
   113
    val { descr, index, sorts = raw_sorts, ... } = the_info thy dtco;
haftmann
parents: 33969
diff changeset
   114
    val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index;
haftmann
parents: 33969
diff changeset
   115
    val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
haftmann
parents: 33969
diff changeset
   116
      o Datatype_Aux.dest_DtTFree) dtys;
haftmann
parents: 33969
diff changeset
   117
    val cos = map
haftmann
parents: 33969
diff changeset
   118
      (fn (co, tys) => (co, map (Datatype_Aux.typ_of_dtyp descr sorts) tys)) raw_cos;
haftmann
parents: 33969
diff changeset
   119
  in (sorts, cos) end;
haftmann
parents: 33969
diff changeset
   120
haftmann
parents: 33969
diff changeset
   121
fun the_descr thy (raw_tycos as raw_tyco :: _) =
haftmann
parents: 33969
diff changeset
   122
  let
haftmann
parents: 33969
diff changeset
   123
    val info = the_info thy raw_tyco;
haftmann
parents: 33969
diff changeset
   124
    val descr = #descr info;
haftmann
parents: 33969
diff changeset
   125
haftmann
parents: 33969
diff changeset
   126
    val SOME (_, dtys, _) = AList.lookup (op =) descr (#index info);
haftmann
parents: 33969
diff changeset
   127
    val vs = map ((fn v => (v, (the o AList.lookup (op =) (#sorts info)) v))
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   128
      o Datatype_Aux.dest_DtTFree) dtys;
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
   129
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   130
    fun is_DtTFree (Datatype_Aux.DtTFree _) = true
33970
haftmann
parents: 33969
diff changeset
   131
      | is_DtTFree _ = false
haftmann
parents: 33969
diff changeset
   132
    val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   133
    val protoTs as (dataTs, _) =
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   134
      chop k descr
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   135
      |> (pairself o map)
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   136
        (fn (_, (tyco, dTs, _)) => (tyco, map (Datatype_Aux.typ_of_dtyp descr vs) dTs));
33970
haftmann
parents: 33969
diff changeset
   137
    
haftmann
parents: 33969
diff changeset
   138
    val tycos = map fst dataTs;
45430
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   139
    val _ =
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   140
      if eq_set (op =) (tycos, raw_tycos) then ()
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   141
      else error ("Type constructors " ^ commas_quote raw_tycos ^
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   142
        " do not belong exhaustively to one mutual recursive datatype");
33970
haftmann
parents: 33969
diff changeset
   143
haftmann
parents: 33969
diff changeset
   144
    val (Ts, Us) = (pairself o map) Type protoTs;
haftmann
parents: 33969
diff changeset
   145
haftmann
parents: 33969
diff changeset
   146
    val names = map Long_Name.base_name (the_default tycos (#alt_names info));
haftmann
parents: 33969
diff changeset
   147
    val (auxnames, _) = Name.make_context names
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 43253
diff changeset
   148
      |> fold_map (Name.variant o Datatype_Aux.name_of_typ) Us;
33970
haftmann
parents: 33969
diff changeset
   149
    val prefix = space_implode "_" names;
haftmann
parents: 33969
diff changeset
   150
haftmann
parents: 33969
diff changeset
   151
  in (descr, vs, tycos, prefix, (names, auxnames), (Ts, Us)) end;
haftmann
parents: 33969
diff changeset
   152
haftmann
parents: 33969
diff changeset
   153
fun all_distincts thy Ts =
haftmann
parents: 33969
diff changeset
   154
  let
haftmann
parents: 33969
diff changeset
   155
    fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts
haftmann
parents: 33969
diff changeset
   156
      | add_tycos _ = I;
haftmann
parents: 33969
diff changeset
   157
    val tycos = fold add_tycos Ts [];
haftmann
parents: 33969
diff changeset
   158
  in map_filter (Option.map #distinct o get_info thy) tycos end;
haftmann
parents: 33969
diff changeset
   159
haftmann
parents: 33969
diff changeset
   160
fun get_constrs thy dtco =
haftmann
parents: 33969
diff changeset
   161
  case try (the_spec thy) dtco
haftmann
parents: 33969
diff changeset
   162
   of SOME (sorts, cos) =>
haftmann
parents: 33969
diff changeset
   163
        let
haftmann
parents: 33969
diff changeset
   164
          fun subst (v, sort) = TVar ((v, 0), sort);
haftmann
parents: 33969
diff changeset
   165
          fun subst_ty (TFree v) = subst v
haftmann
parents: 33969
diff changeset
   166
            | subst_ty ty = ty;
haftmann
parents: 33969
diff changeset
   167
          val dty = Type (dtco, map subst sorts);
haftmann
parents: 33969
diff changeset
   168
          fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
haftmann
parents: 33969
diff changeset
   169
        in SOME (map mk_co cos) end
haftmann
parents: 33969
diff changeset
   170
    | NONE => NONE;
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
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
33970
haftmann
parents: 33969
diff changeset
   174
(** various auxiliary **)
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
33970
haftmann
parents: 33969
diff changeset
   176
(* prepare datatype specifications *)
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
   177
33970
haftmann
parents: 33969
diff changeset
   178
fun read_typ thy str sorts =
haftmann
parents: 33969
diff changeset
   179
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   180
    val ctxt = Proof_Context.init_global thy
33970
haftmann
parents: 33969
diff changeset
   181
      |> fold (Variable.declare_typ o TFree) sorts;
haftmann
parents: 33969
diff changeset
   182
    val T = Syntax.read_typ ctxt str;
haftmann
parents: 33969
diff changeset
   183
  in (T, Term.add_tfreesT T sorts) 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
   184
33970
haftmann
parents: 33969
diff changeset
   185
fun cert_typ sign raw_T sorts =
haftmann
parents: 33969
diff changeset
   186
  let
haftmann
parents: 33969
diff changeset
   187
    val T = Type.no_tvars (Sign.certify_typ sign raw_T)
haftmann
parents: 33969
diff changeset
   188
      handle TYPE (msg, _, _) => error msg;
haftmann
parents: 33969
diff changeset
   189
    val sorts' = Term.add_tfreesT T sorts;
haftmann
parents: 33969
diff changeset
   190
    val _ =
haftmann
parents: 33969
diff changeset
   191
      case duplicates (op =) (map fst sorts') of
haftmann
parents: 33969
diff changeset
   192
        [] => ()
haftmann
parents: 33969
diff changeset
   193
      | dups => error ("Inconsistent sort constraints for " ^ commas dups)
haftmann
parents: 33969
diff changeset
   194
  in (T, sorts') 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
   195
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
33970
haftmann
parents: 33969
diff changeset
   197
(* case names *)
haftmann
parents: 33969
diff changeset
   198
haftmann
parents: 33969
diff changeset
   199
local
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
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   201
fun dt_recs (Datatype_Aux.DtTFree _) = []
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   202
  | dt_recs (Datatype_Aux.DtType (_, dts)) = maps dt_recs dts
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   203
  | dt_recs (Datatype_Aux.DtRec i) = [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
   204
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   205
fun dt_cases (descr: Datatype_Aux.descr) (_, args, constrs) =
33970
haftmann
parents: 33969
diff changeset
   206
  let
haftmann
parents: 33969
diff changeset
   207
    fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
haftmann
parents: 33969
diff changeset
   208
    val bnames = map the_bname (distinct (op =) (maps dt_recs args));
haftmann
parents: 33969
diff changeset
   209
  in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs 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
   210
33970
haftmann
parents: 33969
diff changeset
   211
fun induct_cases descr =
haftmann
parents: 33969
diff changeset
   212
  Datatype_Prop.indexify_names (maps (dt_cases descr) (map #2 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
   213
33970
haftmann
parents: 33969
diff changeset
   214
fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
haftmann
parents: 33969
diff changeset
   215
haftmann
parents: 33969
diff changeset
   216
in
haftmann
parents: 33969
diff changeset
   217
haftmann
parents: 33969
diff changeset
   218
fun mk_case_names_induct descr = Rule_Cases.case_names (induct_cases 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
   219
33970
haftmann
parents: 33969
diff changeset
   220
fun mk_case_names_exhausts descr new =
haftmann
parents: 33969
diff changeset
   221
  map (Rule_Cases.case_names o exhaust_cases descr o #1)
haftmann
parents: 33969
diff changeset
   222
    (filter (fn ((_, (name, _, _))) => member (op =) new name) 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
   223
33970
haftmann
parents: 33969
diff changeset
   224
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
   225
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   226
33970
haftmann
parents: 33969
diff changeset
   227
(* translation rules for case *)
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
   228
33970
haftmann
parents: 33969
diff changeset
   229
fun make_case ctxt = Datatype_Case.make_case
43580
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
   230
  (info_of_constr_permissive (Proof_Context.theory_of ctxt)) ctxt;
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
   231
33970
haftmann
parents: 33969
diff changeset
   232
fun strip_case ctxt = Datatype_Case.strip_case
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   233
  (info_of_case (Proof_Context.theory_of ctxt));
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
   234
33970
haftmann
parents: 33969
diff changeset
   235
fun add_case_tr' case_names thy =
haftmann
parents: 33969
diff changeset
   236
  Sign.add_advanced_trfuns ([], [],
haftmann
parents: 33969
diff changeset
   237
    map (fn case_name =>
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42289
diff changeset
   238
      let val case_name' = Lexicon.mark_const case_name
33970
haftmann
parents: 33969
diff changeset
   239
      in (case_name', Datatype_Case.case_tr' info_of_case case_name')
haftmann
parents: 33969
diff changeset
   240
      end) case_names, []) 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
   241
33970
haftmann
parents: 33969
diff changeset
   242
val trfun_setup =
haftmann
parents: 33969
diff changeset
   243
  Sign.add_advanced_trfuns ([],
43580
023a1d1f97bd new Datatype.info_of_constr with strict behaviour wrt. to overloaded constructors -- side effect: function package correctly identifies 0::int as a non-constructor;
krauss
parents: 43329
diff changeset
   244
    [(@{syntax_const "_case_syntax"}, Datatype_Case.case_tr true info_of_constr_permissive)],
33970
haftmann
parents: 33969
diff changeset
   245
    [], []);
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
   246
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   247
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
33970
haftmann
parents: 33969
diff changeset
   249
(** document antiquotation **)
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
   250
43564
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   251
val antiq_setup =
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   252
  Thy_Output.antiquotation @{binding datatype} (Args.type_name true)
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   253
    (fn {source = src, context = ctxt, ...} => fn dtco =>
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   254
      let
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   255
        val thy = Proof_Context.theory_of ctxt;
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   256
        val (vs, cos) = the_spec thy dtco;
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   257
        val ty = Type (dtco, map TFree vs);
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   258
        val pretty_typ_bracket = Syntax.pretty_typ (Config.put pretty_priority 1001 ctxt);
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   259
        fun pretty_constr (co, tys) =
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   260
          Pretty.block (Pretty.breaks
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   261
            (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   262
              map pretty_typ_bracket tys));
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   263
        val pretty_datatype =
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   264
          Pretty.block
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   265
           (Pretty.command "datatype" :: Pretty.brk 1 ::
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   266
            Syntax.pretty_typ ctxt ty ::
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   267
            Pretty.str " =" :: Pretty.brk 1 ::
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   268
            flat (separate [Pretty.brk 1, Pretty.str "| "] (map (single o pretty_constr) cos)));
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   269
      in
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   270
        Thy_Output.output ctxt
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   271
          (Thy_Output.maybe_pretty_source (K (K pretty_datatype)) ctxt src [()])
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   272
      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
   273
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   274
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   275
33970
haftmann
parents: 33969
diff changeset
   276
(** abstract theory extensions relative to a datatype characterisation **)
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
33970
haftmann
parents: 33969
diff changeset
   278
structure Datatype_Interpretation = Interpretation
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   279
  (type T = Datatype_Aux.config * string list val eq: T * T -> bool = eq_snd op =);
33970
haftmann
parents: 33969
diff changeset
   280
fun interpretation f = Datatype_Interpretation.interpretation (uncurry f);
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
   281
33970
haftmann
parents: 33969
diff changeset
   282
fun make_dt_info alt_names descr sorts induct inducts rec_names rec_rewrites
haftmann
parents: 33969
diff changeset
   283
    (index, (((((((((((_, (tname, _, _))), inject), distinct),
haftmann
parents: 33969
diff changeset
   284
      exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong),
haftmann
parents: 33969
diff changeset
   285
        (split, split_asm))) =
haftmann
parents: 33969
diff changeset
   286
  (tname,
haftmann
parents: 33969
diff changeset
   287
   {index = index,
haftmann
parents: 33969
diff changeset
   288
    alt_names = alt_names,
haftmann
parents: 33969
diff changeset
   289
    descr = descr,
haftmann
parents: 33969
diff changeset
   290
    sorts = sorts,
haftmann
parents: 33969
diff changeset
   291
    inject = inject,
haftmann
parents: 33969
diff changeset
   292
    distinct = distinct,
haftmann
parents: 33969
diff changeset
   293
    induct = induct,
haftmann
parents: 33969
diff changeset
   294
    inducts = inducts,
haftmann
parents: 33969
diff changeset
   295
    exhaust = exhaust,
haftmann
parents: 33969
diff changeset
   296
    nchotomy = nchotomy,
haftmann
parents: 33969
diff changeset
   297
    rec_names = rec_names,
haftmann
parents: 33969
diff changeset
   298
    rec_rewrites = rec_rewrites,
haftmann
parents: 33969
diff changeset
   299
    case_name = case_name,
haftmann
parents: 33969
diff changeset
   300
    case_rewrites = case_rewrites,
haftmann
parents: 33969
diff changeset
   301
    case_cong = case_cong,
haftmann
parents: 33969
diff changeset
   302
    weak_case_cong = weak_case_cong,
haftmann
parents: 33969
diff changeset
   303
    split = split,
haftmann
parents: 33969
diff changeset
   304
    split_asm = split_asm});
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
   305
33970
haftmann
parents: 33969
diff changeset
   306
fun derive_datatype_props config dt_names alt_names descr sorts
haftmann
parents: 33969
diff changeset
   307
    induct inject distinct thy1 =
haftmann
parents: 33969
diff changeset
   308
  let
haftmann
parents: 33969
diff changeset
   309
    val thy2 = thy1 |> Theory.checkpoint;
haftmann
parents: 33969
diff changeset
   310
    val flat_descr = flat descr;
haftmann
parents: 33969
diff changeset
   311
    val new_type_names = map Long_Name.base_name (the_default dt_names alt_names);
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   312
    val _ =
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   313
      Datatype_Aux.message config
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   314
        ("Deriving properties for datatype(s) " ^ commas_quote new_type_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
   315
33970
haftmann
parents: 33969
diff changeset
   316
    val (exhaust, thy3) = Datatype_Abs_Proofs.prove_casedist_thms config new_type_names
haftmann
parents: 33969
diff changeset
   317
      descr sorts induct (mk_case_names_exhausts flat_descr dt_names) thy2;
haftmann
parents: 33969
diff changeset
   318
    val (nchotomys, thy4) = Datatype_Abs_Proofs.prove_nchotomys config new_type_names
haftmann
parents: 33969
diff changeset
   319
      descr sorts exhaust thy3;
haftmann
parents: 33969
diff changeset
   320
    val ((rec_names, rec_rewrites), thy5) = Datatype_Abs_Proofs.prove_primrec_thms
haftmann
parents: 33969
diff changeset
   321
      config new_type_names descr sorts (#inject o the o Symtab.lookup (get_all thy4))
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   322
      inject (distinct, all_distincts thy2 (Datatype_Aux.get_rec_types flat_descr sorts))
33970
haftmann
parents: 33969
diff changeset
   323
      induct thy4;
haftmann
parents: 33969
diff changeset
   324
    val ((case_rewrites, case_names), thy6) = Datatype_Abs_Proofs.prove_case_thms
haftmann
parents: 33969
diff changeset
   325
      config new_type_names descr sorts rec_names rec_rewrites thy5;
haftmann
parents: 33969
diff changeset
   326
    val (case_congs, thy7) = Datatype_Abs_Proofs.prove_case_congs new_type_names
haftmann
parents: 33969
diff changeset
   327
      descr sorts nchotomys case_rewrites thy6;
haftmann
parents: 33969
diff changeset
   328
    val (weak_case_congs, thy8) = Datatype_Abs_Proofs.prove_weak_case_congs new_type_names
haftmann
parents: 33969
diff changeset
   329
      descr sorts thy7;
haftmann
parents: 33969
diff changeset
   330
    val (splits, thy9) = Datatype_Abs_Proofs.prove_split_thms
haftmann
parents: 33969
diff changeset
   331
      config new_type_names descr sorts inject distinct exhaust case_rewrites thy8;
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
   332
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   333
    val inducts = Project_Rule.projections (Proof_Context.init_global thy2) induct;
33970
haftmann
parents: 33969
diff changeset
   334
    val dt_infos = map_index
haftmann
parents: 33969
diff changeset
   335
      (make_dt_info alt_names flat_descr sorts induct inducts rec_names rec_rewrites)
haftmann
parents: 33969
diff changeset
   336
      (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~
haftmann
parents: 33969
diff changeset
   337
        case_names ~~ case_rewrites ~~ case_congs ~~ weak_case_congs ~~ splits);
haftmann
parents: 33969
diff changeset
   338
    val dt_names = map fst dt_infos;
haftmann
parents: 33969
diff changeset
   339
    val prfx = Binding.qualify true (space_implode "_" new_type_names);
haftmann
parents: 33969
diff changeset
   340
    val simps = flat (inject @ distinct @ case_rewrites) @ rec_rewrites;
haftmann
parents: 33969
diff changeset
   341
    val named_rules = flat (map_index (fn (index, tname) =>
haftmann
parents: 33969
diff changeset
   342
      [((Binding.empty, [nth inducts index]), [Induct.induct_type tname]),
haftmann
parents: 33969
diff changeset
   343
       ((Binding.empty, [nth exhaust index]), [Induct.cases_type tname])]) dt_names);
haftmann
parents: 33969
diff changeset
   344
    val unnamed_rules = map (fn induct =>
haftmann
parents: 33969
diff changeset
   345
      ((Binding.empty, [induct]), [Rule_Cases.inner_rule, Induct.induct_type ""]))
haftmann
parents: 33969
diff changeset
   346
        (drop (length dt_names) inducts);
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
   347
  in
33970
haftmann
parents: 33969
diff changeset
   348
    thy9
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 38767
diff changeset
   349
    |> Global_Theory.add_thmss ([((prfx (Binding.name "simps"), simps), []),
33970
haftmann
parents: 33969
diff changeset
   350
        ((prfx (Binding.name "inducts"), inducts), []),
haftmann
parents: 33969
diff changeset
   351
        ((prfx (Binding.name "splits"), maps (fn (x, y) => [x, y]) splits), []),
haftmann
parents: 33969
diff changeset
   352
        ((Binding.empty, flat case_rewrites @ flat distinct @ rec_rewrites),
haftmann
parents: 33969
diff changeset
   353
          [Simplifier.simp_add]),
haftmann
parents: 33969
diff changeset
   354
        ((Binding.empty, rec_rewrites), [Code.add_default_eqn_attribute]),
haftmann
parents: 33969
diff changeset
   355
        ((Binding.empty, flat inject), [iff_add]),
haftmann
parents: 33969
diff changeset
   356
        ((Binding.empty, map (fn th => th RS notE) (flat distinct)),
haftmann
parents: 33969
diff changeset
   357
          [Classical.safe_elim NONE]),
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45430
diff changeset
   358
        ((Binding.empty, weak_case_congs), [Simplifier.cong_add]),
36602
0cbcdfd9e527 slightly more standard induct_simp_add/del attributes;
wenzelm
parents: 36323
diff changeset
   359
        ((Binding.empty, flat (distinct @ inject)), [Induct.induct_simp_add])] @
0cbcdfd9e527 slightly more standard induct_simp_add/del attributes;
wenzelm
parents: 36323
diff changeset
   360
          named_rules @ unnamed_rules)
33970
haftmann
parents: 33969
diff changeset
   361
    |> snd
haftmann
parents: 33969
diff changeset
   362
    |> add_case_tr' case_names
haftmann
parents: 33969
diff changeset
   363
    |> register dt_infos
haftmann
parents: 33969
diff changeset
   364
    |> Datatype_Interpretation.data (config, dt_names)
haftmann
parents: 33969
diff changeset
   365
    |> pair dt_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
   366
  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
   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
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
33970
haftmann
parents: 33969
diff changeset
   370
(** declare existing type as datatype **)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   371
33970
haftmann
parents: 33969
diff changeset
   372
fun prove_rep_datatype config dt_names alt_names descr sorts
haftmann
parents: 33969
diff changeset
   373
    raw_inject half_distinct raw_induct thy1 =
haftmann
parents: 33969
diff changeset
   374
  let
haftmann
parents: 33969
diff changeset
   375
    val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
haftmann
parents: 33969
diff changeset
   376
    val new_type_names = map Long_Name.base_name (the_default dt_names alt_names);
40929
7ff03a5e044f theorem names generated by the (rep_)datatype command now have mandatory qualifiers
huffman
parents: 40844
diff changeset
   377
    val prfx = Binding.qualify true (space_implode "_" new_type_names);
33970
haftmann
parents: 33969
diff changeset
   378
    val (((inject, distinct), [induct]), thy2) =
haftmann
parents: 33969
diff changeset
   379
      thy1
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   380
      |> Datatype_Aux.store_thmss "inject" new_type_names raw_inject
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   381
      ||>> Datatype_Aux.store_thmss "distinct" new_type_names raw_distinct
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 38767
diff changeset
   382
      ||>> Global_Theory.add_thms
40929
7ff03a5e044f theorem names generated by the (rep_)datatype command now have mandatory qualifiers
huffman
parents: 40844
diff changeset
   383
        [((prfx (Binding.name "induct"), raw_induct),
7ff03a5e044f theorem names generated by the (rep_)datatype command now have mandatory qualifiers
huffman
parents: 40844
diff changeset
   384
          [mk_case_names_induct descr])];
33970
haftmann
parents: 33969
diff changeset
   385
  in
haftmann
parents: 33969
diff changeset
   386
    thy2
haftmann
parents: 33969
diff changeset
   387
    |> derive_datatype_props config dt_names alt_names [descr] sorts
haftmann
parents: 33969
diff changeset
   388
         induct inject distinct
haftmann
parents: 33969
diff changeset
   389
 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
   390
33970
haftmann
parents: 33969
diff changeset
   391
fun gen_rep_datatype prep_term config after_qed alt_names raw_ts thy =
haftmann
parents: 33969
diff changeset
   392
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   393
    val ctxt = Proof_Context.init_global thy;
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42297
diff changeset
   394
33970
haftmann
parents: 33969
diff changeset
   395
    fun constr_of_term (Const (c, T)) = (c, T)
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42297
diff changeset
   396
      | constr_of_term t = error ("Not a constant: " ^ Syntax.string_of_term ctxt t);
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42297
diff changeset
   397
    fun no_constr (c, T) =
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   398
      error ("Bad constructor: " ^ Proof_Context.extern_const ctxt c ^ "::" ^
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42297
diff changeset
   399
        Syntax.string_of_typ ctxt T);
33970
haftmann
parents: 33969
diff changeset
   400
    fun type_of_constr (cT as (_, 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
   401
      let
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43581
diff changeset
   402
        val frees = Misc_Legacy.typ_tfrees T;
40844
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 39557
diff changeset
   403
        val (tyco, vs) = (apsnd o map) dest_TFree (dest_Type (body_type T))
33970
haftmann
parents: 33969
diff changeset
   404
          handle TYPE _ => no_constr cT
haftmann
parents: 33969
diff changeset
   405
        val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
haftmann
parents: 33969
diff changeset
   406
        val _ = if length frees <> length vs then no_constr cT else ();
haftmann
parents: 33969
diff changeset
   407
      in (tyco, (vs, cT)) end;
haftmann
parents: 33969
diff changeset
   408
haftmann
parents: 33969
diff changeset
   409
    val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
45430
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   410
    val _ =
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   411
      (case map_filter (fn (tyco, _) =>
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   412
          if Symtab.defined (get_all thy) tyco then SOME tyco else NONE) raw_cs of
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   413
        [] => ()
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   414
      | tycos => error ("Type(s) " ^ commas_quote tycos ^ " already represented inductivly"));
33970
haftmann
parents: 33969
diff changeset
   415
    val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
45430
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   416
    val ms =
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   417
      (case distinct (op =) (map length raw_vss) of
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   418
         [n] => 0 upto n - 1
b8eb7a791dac misc tuning;
wenzelm
parents: 44121
diff changeset
   419
      | _ => error "Different types in given constructors");
33970
haftmann
parents: 33969
diff changeset
   420
    fun inter_sort m = map (fn xs => nth xs m) raw_vss
haftmann
parents: 33969
diff changeset
   421
      |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy))
haftmann
parents: 33969
diff changeset
   422
    val sorts = map inter_sort ms;
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43326
diff changeset
   423
    val vs = Name.invent_names Name.context Name.aT sorts;
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
   424
33970
haftmann
parents: 33969
diff changeset
   425
    fun norm_constr (raw_vs, (c, T)) = (c, map_atyps
haftmann
parents: 33969
diff changeset
   426
      (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
haftmann
parents: 33969
diff changeset
   427
haftmann
parents: 33969
diff changeset
   428
    val cs = map (apsnd (map norm_constr)) raw_cs;
40844
5895c525739d more direct use of binder_types/body_type;
wenzelm
parents: 39557
diff changeset
   429
    val dtyps_of_typ =
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   430
      map (Datatype_Aux.dtyp_of_typ (map (rpair (map fst vs) o fst) cs)) o binder_types;
33970
haftmann
parents: 33969
diff changeset
   431
    val dt_names = map fst cs;
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
   432
33970
haftmann
parents: 33969
diff changeset
   433
    fun mk_spec (i, (tyco, constr)) = (i, (tyco,
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   434
      map (Datatype_Aux.DtTFree o fst) vs,
33970
haftmann
parents: 33969
diff changeset
   435
      (map o apsnd) dtyps_of_typ constr))
haftmann
parents: 33969
diff changeset
   436
    val descr = map_index mk_spec cs;
haftmann
parents: 33969
diff changeset
   437
    val injs = Datatype_Prop.make_injs [descr] vs;
haftmann
parents: 33969
diff changeset
   438
    val half_distincts = map snd (Datatype_Prop.make_distincts [descr] vs);
haftmann
parents: 33969
diff changeset
   439
    val ind = Datatype_Prop.make_ind [descr] vs;
haftmann
parents: 33969
diff changeset
   440
    val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
haftmann
parents: 33969
diff changeset
   441
haftmann
parents: 33969
diff changeset
   442
    fun after_qed' raw_thms =
haftmann
parents: 33969
diff changeset
   443
      let
haftmann
parents: 33969
diff changeset
   444
        val [[[raw_induct]], raw_inject, half_distinct] =
haftmann
parents: 33969
diff changeset
   445
          unflat rules (map Drule.zero_var_indexes_list raw_thms);
haftmann
parents: 33969
diff changeset
   446
            (*FIXME somehow dubious*)
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33963
diff changeset
   447
      in
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42359
diff changeset
   448
        Proof_Context.background_theory_result
33970
haftmann
parents: 33969
diff changeset
   449
          (prove_rep_datatype config dt_names alt_names descr vs
haftmann
parents: 33969
diff changeset
   450
            raw_inject half_distinct raw_induct)
haftmann
parents: 33969
diff changeset
   451
        #-> after_qed
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
      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
   453
  in
42359
6ca5407863ed prefer local name spaces;
wenzelm
parents: 42297
diff changeset
   454
    ctxt
36323
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 35360
diff changeset
   455
    |> Proof.theorem NONE after_qed' ((map o map) (rpair []) (flat rules))
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
   456
  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
   457
33970
haftmann
parents: 33969
diff changeset
   458
val rep_datatype = gen_rep_datatype Sign.cert_term;
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   459
val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global Datatype_Aux.default_config (K I);
33970
haftmann
parents: 33969
diff changeset
   460
haftmann
parents: 33969
diff changeset
   461
haftmann
parents: 33969
diff changeset
   462
haftmann
parents: 33969
diff changeset
   463
(** package setup **)
haftmann
parents: 33969
diff changeset
   464
haftmann
parents: 33969
diff changeset
   465
(* setup theory *)
haftmann
parents: 33969
diff changeset
   466
haftmann
parents: 33969
diff changeset
   467
val setup =
haftmann
parents: 33969
diff changeset
   468
  trfun_setup #>
43564
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43329
diff changeset
   469
  antiq_setup #>
33970
haftmann
parents: 33969
diff changeset
   470
  Datatype_Interpretation.init;
haftmann
parents: 33969
diff changeset
   471
haftmann
parents: 33969
diff changeset
   472
haftmann
parents: 33969
diff changeset
   473
(* 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
   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
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   476
  Outer_Syntax.command "rep_datatype" "represent existing types inductively" Keyword.thy_goal
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   477
    (Scan.option (Parse.$$$ "(" |-- Scan.repeat1 Parse.name --| Parse.$$$ ")") --
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   478
      Scan.repeat1 Parse.term
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   479
      >> (fn (alt_names, ts) =>
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36610
diff changeset
   480
        Toplevel.print o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   481
41423
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   482
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   483
open Datatype_Aux;
25df154b8ffc do not open auxiliary ML structures;
wenzelm
parents: 40957
diff changeset
   484
30391
d930432adb13 adapted to simplified ThyOutput.antiquotation interface;
wenzelm
parents: 30364
diff changeset
   485
end;