src/HOL/Tools/Datatype/datatype.ML
author haftmann
Tue, 13 Oct 2009 14:08:01 +0200
changeset 32922 8e40cd05de7a
parent 32915 a7a97960054b
child 33037 b22e44496dc2
permissions -rw-r--r--
deactivated Datatype.distinct_simproc
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31689
diff changeset
     1
(*  Title:      HOL/Tools/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
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
     4
Datatype package for Isabelle/HOL.
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     5
*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     6
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31689
diff changeset
     7
signature DATATYPE =
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
     8
sig
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
     9
  include DATATYPE_COMMON
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    10
  val add_datatype : config -> string list -> (string list * binding * mixfix *
31783
cfbe9609ceb1 add_datatypes does not yield particular rules any longer
haftmann
parents: 31781
diff changeset
    11
    (binding * typ list * mixfix) list) list -> theory -> string list * theory
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    12
  val datatype_cmd : string list -> (string list * binding * mixfix *
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    13
    (binding * string list * mixfix) list) list -> theory -> theory
31781
861e675f01e6 add_datatype interface yields type names and less rules
haftmann
parents: 31775
diff changeset
    14
  val rep_datatype : config -> (string list -> Proof.context -> Proof.context)
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    15
    -> string list option -> term list -> theory -> Proof.state
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    16
  val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
    17
  val get_info : theory -> string -> info option
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
    18
  val the_info : theory -> string -> info
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
    19
  val the_descr : theory -> string list
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    20
    -> descr * (string * sort) list * string list
31868
edd1f30c0477 canonical prefix for datatype derivates
haftmann
parents: 31794
diff changeset
    21
      * string * (string list * string list) * (typ list * typ list)
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
    22
  val the_spec : theory -> string -> (string * sort) list * (string * typ list) list
32911
5f7386f7cbe6 dropped dist_ss
haftmann
parents: 32907
diff changeset
    23
  val all_distincts : theory -> typ list -> thm list list
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
    24
  val get_constrs : theory -> string -> (string * typ) list option
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
    25
  val get_all : theory -> info Symtab.table
32896
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    26
  val info_of_constr : theory -> string * typ -> info option
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
    27
  val info_of_case : theory -> string -> info option
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    28
  val interpretation : (config -> string list -> theory -> theory) -> theory -> theory
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    29
  val distinct_simproc : simproc
32671
fbd224850767 adapted configuration for DatatypeCase.make_case
bulwahn
parents: 32374
diff changeset
    30
  val make_case :  Proof.context -> DatatypeCase.config -> string list -> term ->
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    31
    (term * term) list -> term * (term * (int * bool)) list
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    32
  val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    33
  val read_typ: theory ->
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    34
    (typ list * (string * sort) list) * string -> typ list * (string * sort) list
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
    35
  val setup: theory -> theory
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    36
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    37
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31689
diff changeset
    38
structure Datatype : DATATYPE =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    39
struct
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    40
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    41
open DatatypeAux;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    42
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    43
(** theory data **)
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    44
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    45
(* data management *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    46
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    47
structure DatatypesData = TheoryDataFun
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22777
diff changeset
    48
(
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    49
  type T =
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    50
    {types: info Symtab.table,
32896
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    51
     constrs: (string * info) list Symtab.table,
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
    52
     cases: info Symtab.table};
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    53
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    54
  val empty =
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    55
    {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
6556
daa00919502b theory data: copy;
wenzelm
parents: 6479
diff changeset
    56
  val copy = I;
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    57
  val extend = I;
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    58
  fun merge _
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    59
    ({types = types1, constrs = constrs1, cases = cases1},
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    60
     {types = types2, constrs = constrs2, cases = cases2}) =
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    61
    {types = Symtab.merge (K true) (types1, types2),
32896
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    62
     constrs = Symtab.join (K (AList.merge (op =) (K true))) (constrs1, constrs2),
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    63
     cases = Symtab.merge (K true) (cases1, cases2)};
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22777
diff changeset
    64
);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    65
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
    66
val get_all = #types o DatatypesData.get;
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    67
val get_info = Symtab.lookup o get_all;
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    68
fun the_info thy name = (case get_info thy name of
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    69
      SOME info => info
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    70
    | NONE => error ("Unknown datatype " ^ quote name));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    71
32896
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    72
fun info_of_constr thy (c, T) =
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    73
  let
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    74
    val tab = Symtab.lookup_list ((#constrs o DatatypesData.get) thy) c;
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    75
    val hint = case strip_type T of (_, Type (tyco, _)) => SOME tyco | _ => NONE;
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    76
    val default = if null tab then NONE
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    77
      else SOME (snd (Library.last_elem tab))
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    78
        (*conservative wrt. overloaded constructors*);
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    79
  in case hint
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    80
   of NONE => default
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    81
    | SOME tyco => case AList.lookup (op =) tab tyco
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    82
       of NONE => default (*permissive*)
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    83
        | SOME info => SOME info
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    84
  end;
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    85
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    86
val info_of_case = Symtab.lookup o #cases o DatatypesData.get;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    87
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    88
fun register (dt_infos : (string * info) list) =
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    89
  DatatypesData.map (fn {types, constrs, cases} =>
32896
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    90
    {types = types |> fold Symtab.update dt_infos,
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    91
     constrs = constrs |> fold (fn (constr, dtname_info) =>
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    92
         Symtab.map_default (constr, []) (cons dtname_info))
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    93
       (maps (fn (dtname, info as {descr, index, ...}) =>
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    94
          map (rpair (dtname, info) o fst)
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    95
            (#3 (the (AList.lookup op = descr index)))) dt_infos),
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    96
     cases = cases |> fold Symtab.update
99cd75a18b78 lookup for datatype constructors considers type annotations to resolve overloading
haftmann
parents: 32731
diff changeset
    97
       (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)});
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    98
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
    99
(* complex queries *)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   100
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
   101
fun the_spec thy dtco =
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   102
  let
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
   103
    val info as { descr, index, sorts = raw_sorts, ... } = the_info thy dtco;
26093
haftmann
parents: 25888
diff changeset
   104
    val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index;
haftmann
parents: 25888
diff changeset
   105
    val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
haftmann
parents: 25888
diff changeset
   106
      o DatatypeAux.dest_DtTFree) dtys;
haftmann
parents: 25888
diff changeset
   107
    val cos = map
haftmann
parents: 25888
diff changeset
   108
      (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos;
haftmann
parents: 25888
diff changeset
   109
  in (sorts, cos) end;
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   110
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
   111
fun the_descr thy (raw_tycos as raw_tyco :: _) =
31604
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   112
  let
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
   113
    val info = the_info thy raw_tyco;
31604
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   114
    val descr = #descr info;
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   115
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   116
    val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr (#index info);
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   117
    val vs = map ((fn v => (v, (the o AList.lookup (op =) (#sorts info)) v))
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   118
      o dest_DtTFree) dtys;
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   119
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   120
    fun is_DtTFree (DtTFree _) = true
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   121
      | is_DtTFree _ = false
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   122
    val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr;
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   123
    val protoTs as (dataTs, _) = chop k descr
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   124
      |> (pairself o map) (fn (_, (tyco, dTs, _)) => (tyco, map (typ_of_dtyp descr vs) dTs));
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   125
    
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   126
    val tycos = map fst dataTs;
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   127
    val _ = if gen_eq_set (op =) (tycos, raw_tycos) then ()
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   128
      else error ("Type constructors " ^ commas (map quote raw_tycos)
32374
62617ef2c0d0 inserted space into message
haftmann
parents: 32172
diff changeset
   129
        ^ " do not belong exhaustively to one mutual recursive datatype");
31604
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   130
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   131
    val (Ts, Us) = (pairself o map) Type protoTs;
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   132
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   133
    val names = map Long_Name.base_name (the_default tycos (#alt_names info));
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   134
    val (auxnames, _) = Name.make_context names
31868
edd1f30c0477 canonical prefix for datatype derivates
haftmann
parents: 31794
diff changeset
   135
      |> fold_map (yield_singleton Name.variants o name_of_typ) Us;
edd1f30c0477 canonical prefix for datatype derivates
haftmann
parents: 31794
diff changeset
   136
    val prefix = space_implode "_" names;
31604
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   137
31868
edd1f30c0477 canonical prefix for datatype derivates
haftmann
parents: 31794
diff changeset
   138
  in (descr, vs, tycos, prefix, (names, auxnames), (Ts, Us)) end;
31604
eb2f9d709296 separate directory for datatype package
haftmann
parents: 31361
diff changeset
   139
32911
5f7386f7cbe6 dropped dist_ss
haftmann
parents: 32907
diff changeset
   140
fun all_distincts thy Ts =
5f7386f7cbe6 dropped dist_ss
haftmann
parents: 32907
diff changeset
   141
  let
5f7386f7cbe6 dropped dist_ss
haftmann
parents: 32907
diff changeset
   142
    fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts
5f7386f7cbe6 dropped dist_ss
haftmann
parents: 32907
diff changeset
   143
      | add_tycos _ = I;
5f7386f7cbe6 dropped dist_ss
haftmann
parents: 32907
diff changeset
   144
    val tycos = fold add_tycos Ts [];
5f7386f7cbe6 dropped dist_ss
haftmann
parents: 32907
diff changeset
   145
  in map_filter (Option.map #distinct o get_info thy) tycos end;
5f7386f7cbe6 dropped dist_ss
haftmann
parents: 32907
diff changeset
   146
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
   147
fun get_constrs thy dtco =
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
   148
  case try (the_spec thy) dtco
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   149
   of SOME (sorts, cos) =>
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   150
        let
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   151
          fun subst (v, sort) = TVar ((v, 0), sort);
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   152
          fun subst_ty (TFree v) = subst v
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   153
            | subst_ty ty = ty;
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   154
          val dty = Type (dtco, map subst sorts);
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   155
          fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   156
        in SOME (map mk_co cos) end
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   157
    | NONE => NONE;
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   158
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   159
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   160
(** various auxiliary **)
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   161
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   162
(* simplification procedure for showing distinctness of constructors *)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   163
7060
berghofe
parents: 7049
diff changeset
   164
fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
berghofe
parents: 7049
diff changeset
   165
  | stripT p = p;
berghofe
parents: 7049
diff changeset
   166
berghofe
parents: 7049
diff changeset
   167
fun stripC (i, f $ x) = stripC (i + 1, f)
berghofe
parents: 7049
diff changeset
   168
  | stripC p = p;
berghofe
parents: 7049
diff changeset
   169
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   170
val distinctN = "constr_distinct";
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   171
32900
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32896
diff changeset
   172
fun distinct_rule thy ss tname eq_t = Goal.prove (Simplifier.the_context ss) [] [] eq_t
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32896
diff changeset
   173
  (K (EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
dc883c6126d4 dropped simproc_dist formally
haftmann
parents: 32896
diff changeset
   174
    atac 2, resolve_tac (#distinct (the_info thy tname)) 1, etac FalseE 1]));
25537
55017c8b0a24 interface and distinct simproc tuned
haftmann
parents: 25495
diff changeset
   175
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   176
fun get_constr thy dtco =
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   177
  get_info thy dtco
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   178
  |> Option.map (fn { descr, index, ... } => (#3 o the o AList.lookup (op =) descr) index);
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   179
19539
5b37bb0ad964 ThyInfo.the_theory;
wenzelm
parents: 19346
diff changeset
   180
fun distinct_proc thy ss (t as Const ("op =", _) $ t1 $ t2) =
7060
berghofe
parents: 7049
diff changeset
   181
  (case (stripC (0, t1), stripC (0, t2)) of
berghofe
parents: 7049
diff changeset
   182
     ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
berghofe
parents: 7049
diff changeset
   183
         (case (stripT (0, T1), stripT (0, T2)) of
berghofe
parents: 7049
diff changeset
   184
            ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
berghofe
parents: 7049
diff changeset
   185
                if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   186
                   (case get_constr thy tname1 of
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   187
                      SOME constrs => let val cnames = map fst constrs
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   188
                        in if cname1 mem cnames andalso cname2 mem cnames then
25537
55017c8b0a24 interface and distinct simproc tuned
haftmann
parents: 25495
diff changeset
   189
                             SOME (distinct_rule thy ss tname1
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   190
                               (Logic.mk_equals (t, HOLogic.false_const)))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   191
                           else NONE
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   192
                        end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   193
                    | NONE => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   194
                else NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   195
          | _ => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   196
   | _ => NONE)
20054
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   197
  | distinct_proc _ _ _ = NONE;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   198
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13340
diff changeset
   199
val distinct_simproc =
29064
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
   200
  Simplifier.simproc @{theory HOL} distinctN ["s = t"] distinct_proc;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   201
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   202
val simproc_setup =
26496
49ae9456eba9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26343
diff changeset
   203
  Simplifier.map_simpset (fn ss => ss addsimprocs [distinct_simproc]);
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   204
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   205
(* prepare datatype specifications *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   206
27277
7b7ce2d7fafe export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27261
diff changeset
   207
fun read_typ thy ((Ts, sorts), str) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   208
  let
27277
7b7ce2d7fafe export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27261
diff changeset
   209
    val ctxt = ProofContext.init thy
7b7ce2d7fafe export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27261
diff changeset
   210
      |> fold (Variable.declare_typ o TFree) sorts;
7b7ce2d7fafe export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27261
diff changeset
   211
    val T = Syntax.read_typ ctxt str;
7b7ce2d7fafe export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27261
diff changeset
   212
  in (Ts @ [T], Term.add_tfreesT T sorts) end;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   213
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   214
fun cert_typ sign ((Ts, sorts), raw_T) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   215
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   216
    val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   217
      TYPE (msg, _, _) => error msg;
27277
7b7ce2d7fafe export read_typ/cert_typ -- version with regular context operations;
wenzelm
parents: 27261
diff changeset
   218
    val sorts' = Term.add_tfreesT T sorts;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   219
  in (Ts @ [T],
18964
67f572e03236 renamed gen_duplicates to duplicates;
wenzelm
parents: 18963
diff changeset
   220
      case duplicates (op =) (map fst sorts') of
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   221
         [] => sorts'
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   222
       | dups => error ("Inconsistent sort constraints for " ^ commas dups))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   223
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   224
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   225
(* case names *)
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   226
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   227
local
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   228
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   229
fun dt_recs (DtTFree _) = []
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   230
  | dt_recs (DtType (_, dts)) = maps dt_recs dts
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   231
  | dt_recs (DtRec i) = [i];
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   232
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   233
fun dt_cases (descr: descr) (_, args, constrs) =
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   234
  let
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   235
    fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   236
    val bnames = map the_bname (distinct (op =) (maps dt_recs args));
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   237
  in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   238
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   239
fun induct_cases descr =
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   240
  DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   241
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   242
fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   243
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   244
in
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   245
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   246
fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   247
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   248
fun mk_case_names_exhausts descr new =
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   249
  map (RuleCases.case_names o exhaust_cases descr o #1)
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   250
    (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   251
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   252
end;
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   253
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   254
(* translation rules for case *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   255
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   256
fun make_case ctxt = DatatypeCase.make_case
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   257
  (info_of_constr (ProofContext.theory_of ctxt)) ctxt;
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   258
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   259
fun strip_case ctxt = DatatypeCase.strip_case
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   260
  (info_of_case (ProofContext.theory_of ctxt));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   261
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   262
fun add_case_tr' case_names thy =
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   263
  Sign.add_advanced_trfuns ([], [],
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   264
    map (fn case_name =>
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   265
      let val case_name' = Sign.const_syntax_name thy case_name
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   266
      in (case_name', DatatypeCase.case_tr' info_of_case case_name')
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   267
      end) case_names, []) thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   268
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   269
val trfun_setup =
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   270
  Sign.add_advanced_trfuns ([],
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   271
    [("_case_syntax", DatatypeCase.case_tr true info_of_constr)],
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   272
    [], []);
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   273
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   274
(* document antiquotation *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   275
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   276
val _ = ThyOutput.antiquotation "datatype" Args.tyname
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   277
  (fn {source = src, context = ctxt, ...} => fn dtco =>
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   278
    let
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   279
      val thy = ProofContext.theory_of ctxt;
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   280
      val (vs, cos) = the_spec thy dtco;
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   281
      val ty = Type (dtco, map TFree vs);
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   282
      fun pretty_typ_bracket (ty as Type (_, _ :: _)) =
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   283
            Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty]
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   284
        | pretty_typ_bracket ty =
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   285
            Syntax.pretty_typ ctxt ty;
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   286
      fun pretty_constr (co, tys) =
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   287
        (Pretty.block o Pretty.breaks)
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   288
          (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   289
            map pretty_typ_bracket tys);
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   290
      val pretty_datatype =
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   291
        Pretty.block
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   292
          (Pretty.command "datatype" :: Pretty.brk 1 ::
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   293
           Syntax.pretty_typ ctxt ty ::
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   294
           Pretty.str " =" :: Pretty.brk 1 ::
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   295
           flat (separate [Pretty.brk 1, Pretty.str "| "]
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   296
             (map (single o pretty_constr) cos)));
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   297
    in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   298
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   299
32721
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   300
(** abstract theory extensions relative to a datatype characterisation **)
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   301
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   302
structure DatatypeInterpretation = InterpretationFun
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   303
  (type T = config * string list val eq: T * T -> bool = eq_snd op =);
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   304
fun interpretation f = DatatypeInterpretation.interpretation (uncurry f);
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   305
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   306
fun make_dt_info alt_names descr sorts induct inducts rec_names rec_rewrites
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   307
    (index, (((((((((((_, (tname, _, _))), inject), distinct),
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   308
      exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong), (split, split_asm))) =
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   309
  (tname,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   310
   {index = index,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   311
    alt_names = alt_names,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   312
    descr = descr,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   313
    sorts = sorts,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   314
    inject = inject,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   315
    distinct = distinct,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   316
    induct = induct,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   317
    inducts = inducts,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   318
    exhaust = exhaust,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   319
    nchotomy = nchotomy,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   320
    rec_names = rec_names,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   321
    rec_rewrites = rec_rewrites,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   322
    case_name = case_name,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   323
    case_rewrites = case_rewrites,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   324
    case_cong = case_cong,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   325
    weak_case_cong = weak_case_cong,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   326
    split = split,
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   327
    split_asm = split_asm});
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   328
32728
2c55fc50f670 further unification of datatype and rep_datatype
haftmann
parents: 32727
diff changeset
   329
fun derive_datatype_props config dt_names alt_names descr sorts
32907
0300f8dd63ea dropped rule duplicates
haftmann
parents: 32904
diff changeset
   330
    induct inject distinct thy1 =
32721
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   331
  let
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   332
    val thy2 = thy1 |> Theory.checkpoint;
32728
2c55fc50f670 further unification of datatype and rep_datatype
haftmann
parents: 32727
diff changeset
   333
    val flat_descr = flat descr;
32719
36cae240b46c simplified rep_datatype
haftmann
parents: 32718
diff changeset
   334
    val new_type_names = map Long_Name.base_name (the_default dt_names alt_names);
32728
2c55fc50f670 further unification of datatype and rep_datatype
haftmann
parents: 32727
diff changeset
   335
    val _ = message config ("Deriving properties for datatype(s) " ^ commas_quote new_type_names);
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   336
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   337
    val (exhaust, thy3) = DatatypeAbsProofs.prove_casedist_thms config new_type_names
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   338
      descr sorts induct (mk_case_names_exhausts flat_descr dt_names) thy2;
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   339
    val (nchotomys, thy4) = DatatypeAbsProofs.prove_nchotomys config new_type_names
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   340
      descr sorts exhaust thy3;
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   341
    val ((rec_names, rec_rewrites), thy5) = DatatypeAbsProofs.prove_primrec_thms
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   342
      config new_type_names descr sorts (#inject o the o Symtab.lookup (get_all thy4))
32915
a7a97960054b more appropriate abstraction over distinctness rules
haftmann
parents: 32911
diff changeset
   343
      inject (distinct, all_distincts thy2 (get_rec_types flat_descr sorts))
32911
5f7386f7cbe6 dropped dist_ss
haftmann
parents: 32907
diff changeset
   344
      induct thy4;
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   345
    val ((case_rewrites, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   346
      config new_type_names descr sorts rec_names rec_rewrites thy5;
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   347
    val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   348
      descr sorts nchotomys case_rewrites thy6;
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   349
    val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   350
      descr sorts thy7;
32731
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   351
    val (splits, thy9) = DatatypeAbsProofs.prove_split_thms
32907
0300f8dd63ea dropped rule duplicates
haftmann
parents: 32904
diff changeset
   352
      config new_type_names descr sorts inject distinct exhaust case_rewrites thy8;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   353
32731
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   354
    val inducts = Project_Rule.projections (ProofContext.init thy2) induct;
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   355
    val dt_infos = map_index (make_dt_info alt_names flat_descr sorts induct inducts rec_names rec_rewrites)
32907
0300f8dd63ea dropped rule duplicates
haftmann
parents: 32904
diff changeset
   356
      (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~
32731
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   357
        case_names ~~ case_rewrites ~~ case_congs ~~ weak_case_congs ~~ splits);
31781
861e675f01e6 add_datatype interface yields type names and less rules
haftmann
parents: 31775
diff changeset
   358
    val dt_names = map fst dt_infos;
32731
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   359
    val prfx = Binding.qualify true (space_implode "_" new_type_names);
32907
0300f8dd63ea dropped rule duplicates
haftmann
parents: 32904
diff changeset
   360
    val simps = flat (inject @ distinct @ case_rewrites) @ rec_rewrites;
32731
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   361
    val named_rules = flat (map_index (fn (index, tname) =>
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   362
      [((Binding.empty, [nth inducts index]), [Induct.induct_type tname]),
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   363
       ((Binding.empty, [nth exhaust index]), [Induct.cases_type tname])]) dt_names);
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   364
    val unnamed_rules = map (fn induct =>
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   365
      ((Binding.empty, [induct]), [Thm.kind_internal, Induct.induct_type ""]))
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   366
        (Library.drop (length dt_names, inducts));
32720
fc32e6771749 streamlined rep_datatype further
haftmann
parents: 32719
diff changeset
   367
  in
fc32e6771749 streamlined rep_datatype further
haftmann
parents: 32719
diff changeset
   368
    thy9
32731
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   369
    |> PureThy.add_thmss ([((prfx (Binding.name "simps"), simps), []),
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   370
        ((prfx (Binding.name "inducts"), inducts), []),
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   371
        ((prfx (Binding.name "splits"), maps (fn (x, y) => [x, y]) splits), []),
32907
0300f8dd63ea dropped rule duplicates
haftmann
parents: 32904
diff changeset
   372
        ((Binding.empty, flat case_rewrites @ flat distinct @ rec_rewrites),
32731
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   373
          [Simplifier.simp_add]),
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   374
        ((Binding.empty, rec_rewrites), [Code.add_default_eqn_attribute]),
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   375
        ((Binding.empty, flat inject), [iff_add]),
32907
0300f8dd63ea dropped rule duplicates
haftmann
parents: 32904
diff changeset
   376
        ((Binding.empty, map (fn th => th RS notE) (flat distinct)),
32731
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   377
          [Classical.safe_elim NONE]),
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   378
        ((Binding.empty, weak_case_congs), [Simplifier.attrib (op addcongs)])]
f7ed14d60818 less auxiliary functions
haftmann
parents: 32730
diff changeset
   379
        @ named_rules @ unnamed_rules)
32720
fc32e6771749 streamlined rep_datatype further
haftmann
parents: 32719
diff changeset
   380
    |> snd
32730
haftmann
parents: 32729
diff changeset
   381
    |> add_case_tr' case_names
haftmann
parents: 32729
diff changeset
   382
    |> register dt_infos
32720
fc32e6771749 streamlined rep_datatype further
haftmann
parents: 32719
diff changeset
   383
    |> DatatypeInterpretation.data (config, dt_names)
fc32e6771749 streamlined rep_datatype further
haftmann
parents: 32719
diff changeset
   384
    |> pair dt_names
fc32e6771749 streamlined rep_datatype further
haftmann
parents: 32719
diff changeset
   385
  end;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   386
32721
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   387
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   388
(** declare existing type as datatype **)
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   389
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   390
fun prove_rep_datatype config dt_names alt_names descr sorts raw_inject half_distinct raw_induct thy1 =
32721
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   391
  let
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   392
    val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   393
    val new_type_names = map Long_Name.base_name (the_default dt_names alt_names);
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   394
    val (((inject, distinct), [induct]), thy2) =
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   395
      thy1
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   396
      |> store_thmss "inject" new_type_names raw_inject
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   397
      ||>> store_thmss "distinct" new_type_names raw_distinct
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   398
      ||> Sign.add_path (space_implode "_" new_type_names)
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   399
      ||>> PureThy.add_thms [((Binding.name "induct", raw_induct), [mk_case_names_induct descr])]
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   400
      ||> Sign.restore_naming thy1;
32722
ad04cda866be explicit pointless checkpoint
haftmann
parents: 32721
diff changeset
   401
  in
ad04cda866be explicit pointless checkpoint
haftmann
parents: 32721
diff changeset
   402
    thy2
32728
2c55fc50f670 further unification of datatype and rep_datatype
haftmann
parents: 32727
diff changeset
   403
    |> derive_datatype_props config dt_names alt_names [descr] sorts
32907
0300f8dd63ea dropped rule duplicates
haftmann
parents: 32904
diff changeset
   404
         induct inject distinct
32722
ad04cda866be explicit pointless checkpoint
haftmann
parents: 32721
diff changeset
   405
 end;
32721
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   406
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   407
fun gen_rep_datatype prep_term config after_qed alt_names raw_ts thy =
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   408
  let
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   409
    fun constr_of_term (Const (c, T)) = (c, T)
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   410
      | constr_of_term t =
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   411
          error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   412
    fun no_constr (c, T) = error ("Bad constructor: "
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   413
      ^ Sign.extern_const thy c ^ "::"
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   414
      ^ Syntax.string_of_typ_global thy T);
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   415
    fun type_of_constr (cT as (_, T)) =
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   416
      let
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29064
diff changeset
   417
        val frees = OldTerm.typ_tfrees T;
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   418
        val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   419
          handle TYPE _ => no_constr cT
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   420
        val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   421
        val _ = if length frees <> length vs then no_constr cT else ();
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   422
      in (tyco, (vs, cT)) end;
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   423
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   424
    val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   425
    val _ = case map_filter (fn (tyco, _) =>
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
   426
        if Symtab.defined (get_all thy) tyco then SOME tyco else NONE) raw_cs
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   427
     of [] => ()
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   428
      | tycos => error ("Type(s) " ^ commas (map quote tycos)
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   429
          ^ " already represented inductivly");
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   430
    val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   431
    val ms = case distinct (op =) (map length raw_vss)
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   432
     of [n] => 0 upto n - 1
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   433
      | _ => error ("Different types in given constructors");
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   434
    fun inter_sort m = map (fn xs => nth xs m) raw_vss
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   435
      |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy))
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   436
    val sorts = map inter_sort ms;
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   437
    val vs = Name.names Name.context Name.aT sorts;
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   438
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   439
    fun norm_constr (raw_vs, (c, T)) = (c, map_atyps
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   440
      (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   441
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   442
    val cs = map (apsnd (map norm_constr)) raw_cs;
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   443
    val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs))
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   444
      o fst o strip_type;
32719
36cae240b46c simplified rep_datatype
haftmann
parents: 32718
diff changeset
   445
    val dt_names = map fst cs;
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   446
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   447
    fun mk_spec (i, (tyco, constr)) = (i, (tyco,
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   448
      map (DtTFree o fst) vs,
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   449
      (map o apsnd) dtyps_of_typ constr))
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   450
    val descr = map_index mk_spec cs;
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   451
    val injs = DatatypeProp.make_injs [descr] vs;
27300
4cb3101d2bf7 DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
haftmann
parents: 27277
diff changeset
   452
    val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs);
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   453
    val ind = DatatypeProp.make_ind [descr] vs;
27300
4cb3101d2bf7 DatatypeProp.make_distincts: only one half of each symmetric pair is constructed
haftmann
parents: 27277
diff changeset
   454
    val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   455
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   456
    fun after_qed' raw_thms =
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   457
      let
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   458
        val [[[raw_induct]], raw_inject, half_distinct] =
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   459
          unflat rules (map Drule.zero_var_indexes_list raw_thms);
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   460
            (*FIXME somehow dubious*)
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   461
      in
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   462
        ProofContext.theory_result
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   463
          (prove_rep_datatype config dt_names alt_names descr vs raw_inject half_distinct raw_induct)
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   464
        #-> after_qed
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   465
      end;
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   466
  in
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   467
    thy
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   468
    |> ProofContext.init
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   469
    |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   470
  end;
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   471
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   472
val rep_datatype = gen_rep_datatype Sign.cert_term;
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   473
val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global default_config (K I);
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   474
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   475
32717
0e787c721fa3 re-established reasonable inner outline for module
haftmann
parents: 32712
diff changeset
   476
(** definitional introduction of datatypes **)
11958
2ece34b9fd8e Isar: fixed rep_datatype args;
wenzelm
parents: 11805
diff changeset
   477
32721
a5fcc7960681 emerging common infrastructure for datatype and rep_datatype
haftmann
parents: 32720
diff changeset
   478
fun gen_add_datatype prep_typ config new_type_names dts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   479
  let
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20715
diff changeset
   480
    val _ = Theory.requires thy "Datatype" "datatype definitions";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   481
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   482
    (* this theory is used just for parsing *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   483
    val tmp_thy = thy |>
5892
e5e44cc54eb2 Attribute.tthms_of;
wenzelm
parents: 5720
diff changeset
   484
      Theory.copy |>
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24711
diff changeset
   485
      Sign.add_types (map (fn (tvs, tname, mx, _) =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   486
        (tname, length tvs, mx)) dts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   487
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   488
    val (tyvars, _, _, _)::_ = dts;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   489
    val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   490
      let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname)
18964
67f572e03236 renamed gen_duplicates to duplicates;
wenzelm
parents: 18963
diff changeset
   491
      in (case duplicates (op =) tvs of
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   492
            [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   493
                  else error ("Mutually recursive datatypes must have same type parameters")
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   494
          | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   495
              " : " ^ commas dups))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   496
      end) dts);
32728
2c55fc50f670 further unification of datatype and rep_datatype
haftmann
parents: 32727
diff changeset
   497
    val dt_names = map fst new_dts;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   498
18964
67f572e03236 renamed gen_duplicates to duplicates;
wenzelm
parents: 18963
diff changeset
   499
    val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   500
      [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   501
31361
3e900a2acaed Fixed broken code dealing with alternative names.
berghofe
parents: 31247
diff changeset
   502
    fun prep_dt_spec ((tvs, tname, mx, constrs), tname') (dts', constr_syntax, sorts, i) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   503
      let
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   504
        fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
5279
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   505
          let
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   506
            val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
29270
0eade173f77e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29064
diff changeset
   507
            val _ = (case fold (curry OldTerm.add_typ_tfree_names) cargs' [] \\ tvs of
5279
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   508
                [] => ()
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   509
              | vs => error ("Extra type variables on rhs: " ^ commas vs))
32124
954321008785 dropped ancient flat_names option
haftmann
parents: 31868
diff changeset
   510
          in (constrs @ [(Sign.full_name_path tmp_thy tname'
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   511
                  (Binding.map_name (Syntax.const_name mx') cname),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   512
                   map (dtyp_of_typ new_dts) cargs')],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   513
              constr_syntax' @ [(cname, mx')], sorts'')
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   514
          end handle ERROR msg => cat_error msg
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   515
           ("The error above occured in constructor " ^ quote (Binding.str_of cname) ^
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   516
            " of datatype " ^ quote (Binding.str_of tname));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   517
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   518
        val (constrs', constr_syntax', sorts') =
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   519
          fold prep_constr constrs ([], [], sorts)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   520
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   521
      in
18964
67f572e03236 renamed gen_duplicates to duplicates;
wenzelm
parents: 18963
diff changeset
   522
        case duplicates (op =) (map fst constrs') of
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   523
           [] =>
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   524
             (dts' @ [(i, (Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   525
                map DtTFree tvs, constrs'))],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   526
              constr_syntax @ [constr_syntax'], sorts', i + 1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   527
         | dups => error ("Duplicate constructors " ^ commas dups ^
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   528
             " in datatype " ^ quote (Binding.str_of tname))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   529
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   530
31361
3e900a2acaed Fixed broken code dealing with alternative names.
berghofe
parents: 31247
diff changeset
   531
    val (dts', constr_syntax, sorts', i) =
3e900a2acaed Fixed broken code dealing with alternative names.
berghofe
parents: 31247
diff changeset
   532
      fold prep_dt_spec (dts ~~ new_type_names) ([], [], [], 0);
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   533
    val sorts = sorts' @ (map (rpair (Sign.defaultS tmp_thy)) (tyvars \\ map fst sorts'));
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
   534
    val dt_info = get_all thy;
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   535
    val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
14887
4938ce4ef295 Added exception Datatype_Empty.
berghofe
parents: 14799
diff changeset
   536
    val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
31668
a616e56a5ec8 datatype packages: record datatype_config for configuration flags; less verbose signatures
haftmann
parents: 31604
diff changeset
   537
      if #strict config then error ("Nonemptiness check failed for datatype " ^ s)
14887
4938ce4ef295 Added exception Datatype_Empty.
berghofe
parents: 14799
diff changeset
   538
      else raise exn;
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   539
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   540
    val _ = message config ("Constructing datatype(s) " ^ commas_quote new_type_names);
32907
0300f8dd63ea dropped rule duplicates
haftmann
parents: 32904
diff changeset
   541
    val ((inject, distinct, induct), thy') = thy |>
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   542
      DatatypeRepProofs.representation_proofs config dt_info new_type_names descr sorts
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   543
        types_syntax constr_syntax (mk_case_names_induct (flat descr));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   544
  in
32729
1441cf4ddc1a shared code between rep_datatype and datatype
haftmann
parents: 32728
diff changeset
   545
    derive_datatype_props config dt_names (SOME new_type_names) descr sorts
32907
0300f8dd63ea dropped rule duplicates
haftmann
parents: 32904
diff changeset
   546
      induct inject distinct thy'
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   547
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   548
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   549
val add_datatype = gen_add_datatype cert_typ;
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   550
val datatype_cmd = snd ooo gen_add_datatype read_typ default_config;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   551
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   552
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   553
(** package setup **)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   554
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   555
(* setup theory *)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   556
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
   557
val setup =
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
   558
  trfun_setup #>
24711
e8bba7723858 simplified interpretation setup;
wenzelm
parents: 24699
diff changeset
   559
  DatatypeInterpretation.init;
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   560
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   561
(* outer syntax *)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   562
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   563
local
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   564
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   565
structure P = OuterParse and K = OuterKeyword
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   566
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   567
fun prep_datatype_decls args =
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   568
  let
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   569
    val names = map
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   570
      (fn ((((NONE, _), t), _), _) => Binding.name_of t | ((((SOME t, _), _), _), _) => t) args;
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12708
diff changeset
   571
    val specs = map (fn ((((_, vs), t), mx), cons) =>
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12708
diff changeset
   572
      (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   573
  in (names, specs) end;
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   574
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   575
val parse_datatype_decl =
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   576
  (Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_infix --
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   577
    (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix)));
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   578
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   579
val parse_datatype_decls = P.and_list1 parse_datatype_decl >> prep_datatype_decls;
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   580
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   581
in
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   582
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24830
diff changeset
   583
val _ =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   584
  OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
31735
a00292a5587d tuned interface
haftmann
parents: 31723
diff changeset
   585
    (parse_datatype_decls >> (fn (names, specs) => Toplevel.theory (datatype_cmd names specs)));
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   586
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24830
diff changeset
   587
val _ =
27104
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   588
  OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   589
    (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   590
      >> (fn (alt_names, ts) => Toplevel.print
791607529f6d rep_datatype command now takes list of constructors as input arguments
haftmann
parents: 27097
diff changeset
   591
           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
   592
30391
d930432adb13 adapted to simplified ThyOutput.antiquotation interface;
wenzelm
parents: 30364
diff changeset
   593
end;
d930432adb13 adapted to simplified ThyOutput.antiquotation interface;
wenzelm
parents: 30364
diff changeset
   594
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   595
end;
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   596