src/HOL/Tools/datatype_package.ML
author haftmann
Tue, 25 Sep 2007 12:16:08 +0200
changeset 24699 c6674504103f
parent 24626 85eceef2edc7
child 24711 e8bba7723858
permissions -rw-r--r--
datatype interpretators for size and datatype_realizer
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Tools/datatype_package.ML
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     2
    ID:         $Id$
11539
0f17da240450 tuned headers;
wenzelm
parents: 11345
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     4
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
     5
Datatype package for Isabelle/HOL.
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     6
*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     7
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
     8
signature BASIC_DATATYPE_PACKAGE =
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
     9
sig
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    10
  val induct_tac : string -> int -> tactic
9747
043098ba5098 introduced induct_thm_tac
nipkow
parents: 9739
diff changeset
    11
  val induct_thm_tac : thm -> string -> int -> tactic
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
    12
  val case_tac : string -> int -> tactic
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
    13
  val distinct_simproc : simproc
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    14
end;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    15
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    16
signature DATATYPE_PACKAGE =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    17
sig
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    18
  include BASIC_DATATYPE_PACKAGE
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    19
  val quiet_mode : bool ref
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
    20
  val add_datatype_i : bool -> bool -> string list -> (string list * bstring * mixfix *
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
    21
    (bstring * typ list * mixfix) list) list -> theory ->
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
    22
      {distinct : thm list list,
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
    23
       inject : thm list list,
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
    24
       exhaustion : thm list,
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
    25
       rec_thms : thm list,
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
    26
       case_thms : thm list list,
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
    27
       split_thms : (thm * thm) list,
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
    28
       induction : thm,
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
    29
       simps : thm list} * theory
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    30
  val add_datatype : bool -> string list -> (string list * bstring * mixfix *
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    31
    (bstring * string list * mixfix) list) list -> theory ->
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    32
      {distinct : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    33
       inject : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    34
       exhaustion : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    35
       rec_thms : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    36
       case_thms : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    37
       split_thms : (thm * thm) list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    38
       induction : thm,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    39
       simps : thm list} * theory
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    40
  val rep_datatype_i : string list option -> (thm list * attribute list) list list ->
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    41
    (thm list * attribute list) list list -> (thm list * attribute list) ->
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    42
    theory ->
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    43
      {distinct : thm list list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    44
       inject : thm list list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    45
       exhaustion : thm list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    46
       rec_thms : thm list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    47
       case_thms : thm list list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    48
       split_thms : (thm * thm) list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    49
       induction : thm,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    50
       simps : thm list} * theory
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15661
diff changeset
    51
  val rep_datatype : string list option -> (thmref * Attrib.src list) list list ->
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    52
    (thmref * Attrib.src list) list list -> thmref * Attrib.src list -> theory ->
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    53
      {distinct : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    54
       inject : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    55
       exhaustion : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    56
       rec_thms : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    57
       case_thms : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    58
       split_thms : (thm * thm) list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    59
       induction : thm,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    60
       simps : thm list} * theory
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    61
  val get_datatypes : theory -> DatatypeAux.datatype_info Symtab.table
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
    62
  val get_datatype : theory -> string -> DatatypeAux.datatype_info option
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
    63
  val the_datatype : theory -> string -> DatatypeAux.datatype_info
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    64
  val datatype_of_constr : theory -> string -> DatatypeAux.datatype_info option
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    65
  val datatype_of_case : theory -> string -> DatatypeAux.datatype_info option
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
    66
  val get_datatype_spec : theory -> string -> ((string * sort) list * (string * typ list) list) option
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
    67
  val get_datatype_constrs : theory -> string -> (string * typ) list option
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
    68
  val add_interpretator: (string list -> theory -> theory) -> theory -> theory
6441
268aa3c4a059 print_datatypes;
wenzelm
parents: 6427
diff changeset
    69
  val print_datatypes : theory -> unit
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    70
  val make_case :  Proof.context -> bool -> string list -> term ->
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    71
    (term * term) list -> term * (term * (int * bool)) list
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    72
  val strip_case: Proof.context -> bool ->
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    73
    term -> (term * (term * term) list) option
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
    74
  val setup: theory -> theory
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    75
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    76
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    77
structure DatatypePackage : DATATYPE_PACKAGE =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    78
struct
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    79
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    80
open DatatypeAux;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    81
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    82
val quiet_mode = quiet_mode;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    83
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    84
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22777
diff changeset
    85
(* theory data *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    86
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    87
structure DatatypesData = TheoryDataFun
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22777
diff changeset
    88
(
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    89
  type T =
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    90
    {types: datatype_info Symtab.table,
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    91
     constrs: datatype_info Symtab.table,
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    92
     cases: datatype_info Symtab.table};
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    93
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    94
  val empty =
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    95
    {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
6556
daa00919502b theory data: copy;
wenzelm
parents: 6479
diff changeset
    96
  val copy = I;
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    97
  val extend = I;
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    98
  fun merge _
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
    99
    ({types = types1, constrs = constrs1, cases = cases1},
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   100
     {types = types2, constrs = constrs2, cases = cases2}) =
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   101
    {types = Symtab.merge (K true) (types1, types2),
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   102
     constrs = Symtab.merge (K true) (constrs1, constrs2),
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   103
     cases = Symtab.merge (K true) (cases1, cases2)};
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22777
diff changeset
   104
);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   105
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   106
val get_datatypes = #types o DatatypesData.get;
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   107
val map_datatypes = DatatypesData.map;
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22777
diff changeset
   108
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22777
diff changeset
   109
fun print_datatypes thy =
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22777
diff changeset
   110
  Pretty.writeln (Pretty.strs ("datatypes:" ::
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22777
diff changeset
   111
    map #1 (NameSpace.extern_table (Sign.type_space thy, get_datatypes thy))));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   112
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   113
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   114
(** theory information about datatypes **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   115
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   116
fun put_dt_infos (dt_infos : (string * datatype_info) list) =
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   117
  map_datatypes (fn {types, constrs, cases} =>
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   118
    {types = fold Symtab.update dt_infos types,
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   119
     constrs = fold Symtab.update
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   120
       (maps (fn (_, info as {descr, index, ...}) => map (rpair info o fst)
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   121
          (#3 (the (AList.lookup op = descr index)))) dt_infos) constrs,
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   122
     cases = fold Symtab.update
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   123
       (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   124
       cases});
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   125
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   126
val get_datatype = Symtab.lookup o get_datatypes;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   127
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   128
fun the_datatype thy name = (case get_datatype thy name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   129
      SOME info => info
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   130
    | NONE => error ("Unknown datatype " ^ quote name));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   131
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   132
val datatype_of_constr = Symtab.lookup o #constrs o DatatypesData.get;
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   133
val datatype_of_case = Symtab.lookup o #cases o DatatypesData.get;
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   134
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   135
fun get_datatype_descr thy dtco =
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   136
  get_datatype thy dtco
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   137
  |> Option.map (fn info as { descr, index, ... } => 
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   138
       (info, (((fn SOME (_, dtys, cos) => (dtys, cos)) o AList.lookup (op =) descr) index)));
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   139
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   140
fun get_datatype_spec thy dtco =
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   141
  let
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   142
    fun mk_cons typ_of_dtyp (co, tys) =
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   143
      (co, map typ_of_dtyp tys);
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   144
    fun mk_dtyp ({ sorts = raw_sorts, descr, ... } : DatatypeAux.datatype_info, (dtys, cos)) =
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   145
      let
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   146
        val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   147
          o DatatypeAux.dest_DtTFree) dtys;
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   148
        val typ_of_dtyp = DatatypeAux.typ_of_dtyp descr sorts;
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   149
        val tys = map typ_of_dtyp dtys;
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   150
      in (sorts, map (mk_cons typ_of_dtyp) cos) end;
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   151
  in Option.map mk_dtyp (get_datatype_descr thy dtco) end;
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   152
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   153
fun get_datatype_constrs thy dtco =
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   154
  case get_datatype_spec thy dtco
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   155
   of SOME (sorts, cos) =>
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   156
        let
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   157
          fun subst (v, sort) = TVar ((v, 0), sort);
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   158
          fun subst_ty (TFree v) = subst v
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   159
            | subst_ty ty = ty;
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   160
          val dty = Type (dtco, map subst sorts);
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   161
          fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   162
        in SOME (map mk_co cos) end
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   163
    | NONE => NONE;
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   164
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   165
fun find_tname var Bi =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   166
  let val frees = map dest_Free (term_frees Bi)
14412
109cc0dc706b find_tname now handles parameter renaming properly ("as they are printed").
berghofe
parents: 14174
diff changeset
   167
      val params = rename_wrt_term Bi (Logic.strip_params Bi);
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   168
  in case AList.lookup (op =) (frees @ params) var of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   169
       NONE => error ("No such variable in subgoal: " ^ quote var)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   170
     | SOME(Type (tn, _)) => tn
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   171
     | _ => error ("Cannot determine type of " ^ quote var)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   172
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   173
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   174
fun infer_tname state i aterm =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   175
  let
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 22480
diff changeset
   176
    val sign = Thm.theory_of_thm state;
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   177
    val (_, _, Bi, _) = Thm.dest_state (state, i)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   178
    val params = Logic.strip_params Bi;   (*params of subgoal i*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   179
    val params = rev (rename_wrt_term Bi params);   (*as they are printed*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   180
    val (types, sorts) = types_sorts state;
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   181
    fun types' (a, ~1) = (case AList.lookup (op =) params a of NONE => types(a, ~1) | sm => sm)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   182
      | types' ixn = types ixn;
22709
9ab51bac6287 removed obsolete TypeInfer.logicT -- use dummyT;
wenzelm
parents: 22675
diff changeset
   183
    val ([ct], _) = Thm.read_def_cterms (sign, types', sorts) [] false [(aterm, dummyT)];
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   184
  in case #T (rep_cterm ct) of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   185
       Type (tn, _) => tn
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   186
     | _ => error ("Cannot determine type of " ^ quote aterm)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   187
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   188
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   189
(*Warn if the (induction) variable occurs Free among the premises, which
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   190
  usually signals a mistake.  But calls the tactic either way!*)
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   191
fun occs_in_prems tacf vars =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   192
  SUBGOAL (fn (Bi, i) =>
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   193
           (if exists (fn (a, _) => member (op =) vars a)
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   194
                      (fold Term.add_frees (#2 (strip_context Bi)) [])
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   195
             then warning "Induction variable occurs also among premises!"
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   196
             else ();
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   197
            tacf i));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   198
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   199
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   200
(* generic induction tactic for datatypes *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   201
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   202
local
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   203
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   204
fun prep_var (Var (ixn, _), SOME x) = SOME (ixn, x)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   205
  | prep_var _ = NONE;
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   206
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   207
fun prep_inst (concl, xs) = (*exception Library.UnequalLengths*)
11725
d0c37d04906b HOLogic.dest_concls, InductAttrib.vars_of;
wenzelm
parents: 11539
diff changeset
   208
  let val vs = InductAttrib.vars_of concl
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   209
  in map_filter prep_var (Library.drop (length vs - length xs, vs) ~~ xs) end;
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   210
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   211
in
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   212
18751
38dc4ff2a32b bugfix: induct_tac no longer raises THM "dest_state"
paulson
parents: 18728
diff changeset
   213
fun gen_induct_tac inst_tac (varss, opt_rule) i state = 
38dc4ff2a32b bugfix: induct_tac no longer raises THM "dest_state"
paulson
parents: 18728
diff changeset
   214
  SUBGOAL (fn (Bi,_) =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   215
  let
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   216
    val (rule, rule_name) =
18751
38dc4ff2a32b bugfix: induct_tac no longer raises THM "dest_state"
paulson
parents: 18728
diff changeset
   217
      case opt_rule of
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   218
          SOME r => (r, "Induction rule")
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   219
        | NONE =>
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   220
            let val tn = find_tname (hd (map_filter I (flat varss))) Bi
22598
f31a869077f0 rep_thm/cterm/ctyp: removed obsolete sign field;
wenzelm
parents: 22578
diff changeset
   221
                val thy = Thm.theory_of_thm state
f31a869077f0 rep_thm/cterm/ctyp: removed obsolete sign field;
wenzelm
parents: 22578
diff changeset
   222
            in (#induction (the_datatype thy tn), "Induction rule for type " ^ tn) 
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   223
            end
11725
d0c37d04906b HOLogic.dest_concls, InductAttrib.vars_of;
wenzelm
parents: 11539
diff changeset
   224
    val concls = HOLogic.dest_concls (Thm.concl_of rule);
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   225
    val insts = maps prep_inst (concls ~~ varss) handle Library.UnequalLengths =>
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   226
      error (rule_name ^ " has different numbers of variables");
18751
38dc4ff2a32b bugfix: induct_tac no longer raises THM "dest_state"
paulson
parents: 18728
diff changeset
   227
  in occs_in_prems (inst_tac insts rule) (map #2 insts) i end)
38dc4ff2a32b bugfix: induct_tac no longer raises THM "dest_state"
paulson
parents: 18728
diff changeset
   228
  i state;
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   229
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   230
fun induct_tac s =
15444
4f14c151d9f1 induct_tac and case_tac no longer depend on Syntax.string_of_vname.
berghofe
parents: 14981
diff changeset
   231
  gen_induct_tac Tactic.res_inst_tac'
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   232
    (map (single o SOME) (Syntax.read_idents s), NONE);
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   233
9747
043098ba5098 introduced induct_thm_tac
nipkow
parents: 9739
diff changeset
   234
fun induct_thm_tac th s =
15444
4f14c151d9f1 induct_tac and case_tac no longer depend on Syntax.string_of_vname.
berghofe
parents: 14981
diff changeset
   235
  gen_induct_tac Tactic.res_inst_tac'
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   236
    ([map SOME (Syntax.read_idents s)], SOME th);
9747
043098ba5098 introduced induct_thm_tac
nipkow
parents: 9739
diff changeset
   237
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   238
end;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   239
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   240
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   241
(* generic case tactic for datatypes *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   242
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   243
fun case_inst_tac inst_tac t rule i state =
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   244
  let
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   245
    val _ $ Var (ixn, _) $ _ = HOLogic.dest_Trueprop
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   246
      (hd (Logic.strip_assums_hyp (hd (Thm.prems_of rule))));
15444
4f14c151d9f1 induct_tac and case_tac no longer depend on Syntax.string_of_vname.
berghofe
parents: 14981
diff changeset
   247
  in inst_tac [(ixn, t)] rule i state end;
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   248
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   249
fun gen_case_tac inst_tac (t, SOME rule) i state =
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   250
      case_inst_tac inst_tac t rule i state
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   251
  | gen_case_tac inst_tac (t, NONE) i state =
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   252
      let val tn = infer_tname state i t in
15444
4f14c151d9f1 induct_tac and case_tac no longer depend on Syntax.string_of_vname.
berghofe
parents: 14981
diff changeset
   253
        if tn = HOLogic.boolN then inst_tac [(("P", 0), t)] case_split_thm i state
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   254
        else case_inst_tac inst_tac t
22578
b0eb5652f210 removed obsolete sign_of/sign_of_thm;
wenzelm
parents: 22480
diff changeset
   255
               (#exhaustion (the_datatype (Thm.theory_of_thm state) tn))
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   256
               i state
14471
5688b05b2575 case_tac no longer raises THM exception if goal number is out of range.
berghofe
parents: 14412
diff changeset
   257
      end handle THM _ => Seq.empty;
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   258
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   259
fun case_tac t = gen_case_tac Tactic.res_inst_tac' (t, NONE);
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   260
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   261
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   262
8541
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   263
(** Isar tactic emulations **)
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   264
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   265
local
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   266
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   267
val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.$$$ ":");
18988
d6e5fa2ba8b8 Args/Attrib syntax: Context.generic;
wenzelm
parents: 18964
diff changeset
   268
val opt_rule = Scan.option (rule_spec |-- Attrib.thm);
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   269
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15661
diff changeset
   270
val varss =
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15661
diff changeset
   271
  Args.and_list (Scan.repeat (Scan.unless rule_spec (Scan.lift (Args.maybe Args.name))));
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   272
20328
5b240a4216b0 RuleInsts.bires_inst_tac;
wenzelm
parents: 20218
diff changeset
   273
val inst_tac = RuleInsts.bires_inst_tac false;
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   274
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   275
fun induct_meth ctxt (varss, opt_rule) =
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   276
  gen_induct_tac (inst_tac ctxt) (varss, opt_rule);
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   277
fun case_meth ctxt (varss, opt_rule) =
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   278
  gen_case_tac (inst_tac ctxt) (varss, opt_rule);
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   279
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   280
in
8541
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   281
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   282
val tactic_emulations =
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   283
 [("induct_tac", Method.goal_args_ctxt' (varss -- opt_rule) induct_meth,
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   284
    "induct_tac emulation (dynamic instantiation)"),
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   285
  ("case_tac", Method.goal_args_ctxt' (Scan.lift Args.name -- opt_rule) case_meth,
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   286
    "case_tac emulation (dynamic instantiation)")];
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   287
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   288
end;
8541
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   289
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   290
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   291
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   292
(** induct method setup **)
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   293
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   294
(* case names *)
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   295
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   296
local
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   297
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   298
fun dt_recs (DtTFree _) = []
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   299
  | dt_recs (DtType (_, dts)) = maps dt_recs dts
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   300
  | dt_recs (DtRec i) = [i];
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   301
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   302
fun dt_cases (descr: descr) (_, args, constrs) =
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   303
  let
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   304
    fun the_bname i = Sign.base_name (#1 (the (AList.lookup (op =) descr i)));
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   305
    val bnames = map the_bname (distinct (op =) (maps dt_recs args));
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   306
  in map (fn (c, _) => space_implode "_" (Sign.base_name c :: bnames)) constrs end;
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   307
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   308
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   309
fun induct_cases descr =
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   310
  DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   311
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   312
fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   313
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   314
in
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   315
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   316
fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   317
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   318
fun mk_case_names_exhausts descr new =
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   319
  map (RuleCases.case_names o exhaust_cases descr o #1)
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   320
    (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   321
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   322
end;
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   323
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   324
fun add_rules simps case_thms rec_thms inject distinct
11345
cd605c85e421 improved iff_add_global, new function add_rules factoring out common behaviour
oheimb
parents: 10930
diff changeset
   325
                  weak_case_congs cong_att =
21127
c8e862897d13 cleaned up
haftmann
parents: 21045
diff changeset
   326
  PureThy.add_thmss [(("simps", simps), []),
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   327
    (("", flat case_thms @
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   328
          flat distinct @ rec_thms), [Simplifier.simp_add]),
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   329
    (("", rec_thms), [RecfunCodegen.add_default]),
21127
c8e862897d13 cleaned up
haftmann
parents: 21045
diff changeset
   330
    (("", flat inject), [iff_add]),
21689
58abd6d8abd1 Removal of theorem tagging, which the ATP linkup no longer requires,
paulson
parents: 21646
diff changeset
   331
    (("", map (fn th => th RS notE) (flat distinct)), [Classical.safe_elim NONE]),
21127
c8e862897d13 cleaned up
haftmann
parents: 21045
diff changeset
   332
    (("", weak_case_congs), [cong_att])]
c8e862897d13 cleaned up
haftmann
parents: 21045
diff changeset
   333
  #> snd;
11345
cd605c85e421 improved iff_add_global, new function add_rules factoring out common behaviour
oheimb
parents: 10930
diff changeset
   334
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   335
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   336
(* add_cases_induct *)
8306
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   337
19874
cc4b2b882e4c ProjectRule now context dependent;
wenzelm
parents: 19841
diff changeset
   338
fun add_cases_induct infos induction thy =
8306
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   339
  let
19874
cc4b2b882e4c ProjectRule now context dependent;
wenzelm
parents: 19841
diff changeset
   340
    val inducts = ProjectRule.projections (ProofContext.init thy) induction;
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   341
11805
b110a1ea90da declare projected induction rules stemming from nested recursion;
wenzelm
parents: 11725
diff changeset
   342
    fun named_rules (name, {index, exhaustion, ...}: datatype_info) =
19874
cc4b2b882e4c ProjectRule now context dependent;
wenzelm
parents: 19841
diff changeset
   343
      [(("", nth inducts index), [InductAttrib.induct_type name]),
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   344
       (("", exhaustion), [InductAttrib.cases_type name])];
11805
b110a1ea90da declare projected induction rules stemming from nested recursion;
wenzelm
parents: 11725
diff changeset
   345
    fun unnamed_rule i =
19874
cc4b2b882e4c ProjectRule now context dependent;
wenzelm
parents: 19841
diff changeset
   346
      (("", nth inducts i), [PureThy.kind_internal, InductAttrib.induct_type ""]);
18462
b67d423b5234 *.inducts holds all projected rules;
wenzelm
parents: 18455
diff changeset
   347
  in
19874
cc4b2b882e4c ProjectRule now context dependent;
wenzelm
parents: 19841
diff changeset
   348
    thy |> PureThy.add_thms
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   349
      (maps named_rules infos @
19874
cc4b2b882e4c ProjectRule now context dependent;
wenzelm
parents: 19841
diff changeset
   350
        map unnamed_rule (length infos upto length inducts - 1)) |> snd
cc4b2b882e4c ProjectRule now context dependent;
wenzelm
parents: 19841
diff changeset
   351
    |> PureThy.add_thmss [(("inducts", inducts), [])] |> snd
18462
b67d423b5234 *.inducts holds all projected rules;
wenzelm
parents: 18455
diff changeset
   352
  end;
8306
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   353
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   354
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   355
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   356
(**** simplification procedure for showing distinctness of constructors ****)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   357
7060
berghofe
parents: 7049
diff changeset
   358
fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
berghofe
parents: 7049
diff changeset
   359
  | stripT p = p;
berghofe
parents: 7049
diff changeset
   360
berghofe
parents: 7049
diff changeset
   361
fun stripC (i, f $ x) = stripC (i + 1, f)
berghofe
parents: 7049
diff changeset
   362
  | stripC p = p;
berghofe
parents: 7049
diff changeset
   363
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   364
val distinctN = "constr_distinct";
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   365
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   366
exception ConstrDistinct of term;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   367
19539
5b37bb0ad964 ThyInfo.the_theory;
wenzelm
parents: 19346
diff changeset
   368
fun distinct_proc thy ss (t as Const ("op =", _) $ t1 $ t2) =
7060
berghofe
parents: 7049
diff changeset
   369
  (case (stripC (0, t1), stripC (0, t2)) of
berghofe
parents: 7049
diff changeset
   370
     ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
berghofe
parents: 7049
diff changeset
   371
         (case (stripT (0, T1), stripT (0, T2)) of
berghofe
parents: 7049
diff changeset
   372
            ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
berghofe
parents: 7049
diff changeset
   373
                if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
19539
5b37bb0ad964 ThyInfo.the_theory;
wenzelm
parents: 19346
diff changeset
   374
                   (case (get_datatype_descr thy) tname1 of
19346
c4c003abd830 cleanup in datatype package
haftmann
parents: 19046
diff changeset
   375
                      SOME (_, (_, constrs)) => let val cnames = map fst constrs
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   376
                        in if cname1 mem cnames andalso cname2 mem cnames then
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   377
                             let val eq_t = Logic.mk_equals (t, Const ("False", HOLogic.boolT));
19539
5b37bb0ad964 ThyInfo.the_theory;
wenzelm
parents: 19346
diff changeset
   378
                                 val eq_ct = cterm_of thy eq_t;
5b37bb0ad964 ThyInfo.the_theory;
wenzelm
parents: 19346
diff changeset
   379
                                 val Datatype_thy = ThyInfo.the_theory "Datatype" thy;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   380
                                 val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
16486
1a12cdb6ee6b get_thm(s): Name;
wenzelm
parents: 16430
diff changeset
   381
                                   map (get_thm Datatype_thy o Name)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   382
                                     ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"]
19539
5b37bb0ad964 ThyInfo.the_theory;
wenzelm
parents: 19346
diff changeset
   383
                             in (case (#distinct (the_datatype thy tname1)) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   384
                                 QuickAndDirty => SOME (Thm.invoke_oracle
19539
5b37bb0ad964 ThyInfo.the_theory;
wenzelm
parents: 19346
diff changeset
   385
                                   Datatype_thy distinctN (thy, ConstrDistinct eq_t))
20054
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   386
                               | FewConstrs thms =>
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   387
                                   SOME (Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   388
                                     (EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   389
                                       atac 2, resolve_tac thms 1, etac FalseE 1])))
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   390
                               | ManyConstrs (thm, simpset) =>
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   391
                                   SOME (Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   392
                                     (EVERY [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   393
                                      full_simp_tac (Simplifier.inherit_context ss simpset) 1,
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   394
                                      REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   395
                                      eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   396
                                      etac FalseE 1]))))
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   397
                             end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   398
                           else NONE
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   399
                        end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   400
                    | NONE => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   401
                else NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   402
          | _ => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   403
   | _ => NONE)
20054
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   404
  | distinct_proc _ _ _ = NONE;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   405
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13340
diff changeset
   406
val distinct_simproc =
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   407
  Simplifier.simproc HOL.thy distinctN ["s = t"] distinct_proc;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   408
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   409
val dist_ss = HOL_ss addsimprocs [distinct_simproc];
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   410
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   411
val simproc_setup =
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
   412
  Theory.add_oracle (distinctN, fn (_, ConstrDistinct t) => t) #>
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
   413
  (fn thy => ((change_simpset_of thy) (fn ss => ss addsimprocs [distinct_simproc]); thy));
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   414
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   415
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   416
(**** translation rules for case ****)
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   417
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   418
fun make_case ctxt = DatatypeCase.make_case
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   419
  (datatype_of_constr (ProofContext.theory_of ctxt)) ctxt;
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   420
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   421
fun strip_case ctxt = DatatypeCase.strip_case
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   422
  (datatype_of_case (ProofContext.theory_of ctxt));
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   423
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   424
fun add_case_tr' case_names thy =
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   425
  Theory.add_advanced_trfuns ([], [],
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   426
    map (fn case_name => 
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   427
      let val case_name' = Sign.const_syntax_name thy case_name
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   428
      in (case_name', DatatypeCase.case_tr' datatype_of_case case_name')
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   429
      end) case_names, []) thy;
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   430
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   431
val trfun_setup =
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   432
  Theory.add_advanced_trfuns ([],
24349
0dd8782fb02d Final mods for list comprehension
nipkow
parents: 24346
diff changeset
   433
    [("_case_syntax", DatatypeCase.case_tr true datatype_of_constr)],
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   434
    [], []);
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   435
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   436
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   437
(* prepare types *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   438
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   439
fun read_typ sign ((Ts, sorts), str) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   440
  let
22675
acf10be7dcca cleaned/simplified Sign.read_typ, Thm.read_cterm etc.;
wenzelm
parents: 22598
diff changeset
   441
    val T = Type.no_tvars (Sign.read_def_typ (sign, AList.lookup (op =)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   442
      (map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   443
  in (Ts @ [T], add_typ_tfrees (T, sorts)) end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   444
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   445
fun cert_typ sign ((Ts, sorts), raw_T) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   446
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   447
    val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   448
      TYPE (msg, _, _) => error msg;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   449
    val sorts' = add_typ_tfrees (T, sorts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   450
  in (Ts @ [T],
18964
67f572e03236 renamed gen_duplicates to duplicates;
wenzelm
parents: 18963
diff changeset
   451
      case duplicates (op =) (map fst sorts') of
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   452
         [] => sorts'
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   453
       | dups => error ("Inconsistent sort constraints for " ^ commas dups))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   454
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   455
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   456
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   457
(**** make datatype info ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   458
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   459
fun make_dt_info head_len descr sorts induct reccomb_names rec_thms
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   460
    (((((((((i, (_, (tname, _, _))), case_name), case_thms),
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   461
      exhaustion_thm), distinct_thm), inject), nchotomy), case_cong), weak_case_cong) =
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   462
  (tname,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   463
   {index = i,
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   464
    head_len = head_len,
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   465
    descr = descr,
18319
c52b139ebde0 Added new component "sorts" to record datatype_info.
berghofe
parents: 18314
diff changeset
   466
    sorts = sorts,
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   467
    rec_names = reccomb_names,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   468
    rec_rewrites = rec_thms,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   469
    case_name = case_name,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   470
    case_rewrites = case_thms,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   471
    induction = induct,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   472
    exhaustion = exhaustion_thm,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   473
    distinct = distinct_thm,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   474
    inject = inject,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   475
    nchotomy = nchotomy,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   476
    case_cong = case_cong,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   477
    weak_case_cong = weak_case_cong});
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   478
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   479
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   480
(********************* axiomatic introduction of datatypes ********************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   481
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   482
fun add_axiom label t atts thy =
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   483
  thy
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   484
  |> PureThy.add_axioms_i [((label, t), atts)];
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   485
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   486
fun add_axioms label ts atts thy =
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   487
  thy
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   488
  |> PureThy.add_axiomss_i [((label, ts), atts)];
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   489
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   490
fun add_and_get_axioms_atts label tnames ts attss =
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   491
  fold_map (fn (tname, (atts, t)) => fn thy =>
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   492
    thy
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   493
    |> Theory.add_path tname
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   494
    |> add_axiom label t atts
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   495
    ||> Theory.parent_path
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   496
    |-> (fn [ax] => pair ax)) (tnames ~~ (attss ~~ ts));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   497
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   498
fun add_and_get_axioms label tnames ts =
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   499
  add_and_get_axioms_atts label tnames ts (replicate (length tnames) []);
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   500
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   501
fun add_and_get_axiomss label tnames tss =
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   502
  fold_map (fn (tname, ts) => fn thy =>
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   503
    thy
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   504
    |> Theory.add_path tname
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   505
    |> add_axioms label ts []
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   506
    ||> Theory.parent_path
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   507
    |-> (fn [ax] => pair ax)) (tnames ~~ tss);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   508
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   509
fun gen_specify_consts add args thy =
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   510
  let
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   511
    val specs = map (fn (c, T, mx) =>
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   512
      Const (Sign.full_name thy (Syntax.const_name c mx), T)) args;
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   513
  in
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   514
    thy
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   515
    |> add args
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   516
    |> Theory.add_finals_i false specs
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   517
  end;
19716
52c22fccdaaf add_datatype_axm: finalize specified consts;
wenzelm
parents: 19599
diff changeset
   518
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   519
val specify_consts = gen_specify_consts Sign.add_consts_i;
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   520
val specify_consts_authentic = gen_specify_consts Sign.add_consts_authentic;
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   521
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   522
structure DatatypeInterpretator = TheoryInterpretatorFun(struct type T = string list; val eq = op = end);
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
   523
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   524
fun add_interpretator interpretator = DatatypeInterpretator.add_interpretator (fold interpretator);
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
   525
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   526
fun add_datatype_axm flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   527
    case_names_induct case_names_exhausts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   528
  let
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   529
    val descr' = flat descr;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   530
    val recTs = get_rec_types descr' sorts;
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   531
    val used = map fst (fold Term.add_tfreesT recTs []);
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   532
    val newTs = Library.take (length (hd descr), recTs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   533
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   534
    (**** declare new types and constants ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   535
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   536
    val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   537
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   538
    val constr_decls = map (fn (((_, (_, _, constrs)), T), constr_syntax') =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   539
      map (fn ((_, cargs), (cname, mx)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   540
        (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   541
          (constrs ~~ constr_syntax')) ((hd descr) ~~ newTs ~~ constr_syntax);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   542
15457
1fbd4aba46e3 Adapted to modified interface of PureThy.get_thm(s).
berghofe
parents: 15444
diff changeset
   543
    val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   544
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   545
    val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   546
    val reccomb_names = if length descr' = 1 then [big_reccomb_name] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   547
      (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   548
        (1 upto (length descr')));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   549
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 20054
diff changeset
   550
    val freeT = TFree (Name.variant used "'t", HOLogic.typeS);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   551
    val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   552
      map (fn (_, cargs) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   553
        let val Ts = map (typ_of_dtyp descr' sorts) cargs
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   554
        in Ts ---> freeT end) constrs) (hd descr);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   555
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   556
    val case_names = map (fn s => (s ^ "_case")) new_type_names;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   557
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   558
    val thy2' = thy
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   559
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   560
      (** new types **)
21419
809e7520234a added instance for class size
haftmann
parents: 21350
diff changeset
   561
      |> fold2 (fn (name, mx) => fn tvs => TypedefPackage.add_typedecls [(name, tvs, mx)])
809e7520234a added instance for class size
haftmann
parents: 21350
diff changeset
   562
           types_syntax tyvars
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   563
      |> add_path flat_names (space_implode "_" new_type_names)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   564
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   565
      (** primrec combinators **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   566
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   567
      |> specify_consts (map (fn ((name, T), T') =>
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   568
           (name, reccomb_fn_Ts @ [T] ---> T', NoSyn)) (reccomb_names ~~ recTs ~~ rec_result_Ts))
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   569
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   570
      (** case combinators **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   571
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   572
      |> specify_consts_authentic (map (fn ((name, T), Ts) =>
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   573
           (name, Ts @ [T] ---> freeT, NoSyn)) (case_names ~~ newTs ~~ case_fn_Ts));
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   574
19716
52c22fccdaaf add_datatype_axm: finalize specified consts;
wenzelm
parents: 19599
diff changeset
   575
    val reccomb_names' = map (Sign.full_name thy2') reccomb_names;
52c22fccdaaf add_datatype_axm: finalize specified consts;
wenzelm
parents: 19599
diff changeset
   576
    val case_names' = map (Sign.full_name thy2') case_names;
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   577
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   578
    val thy2 = thy2'
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   579
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   580
      (** constructors **)
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   581
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   582
      |> parent_path flat_names
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   583
      |> fold (fn ((((_, (_, _, constrs)), T), tname),
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   584
        constr_syntax') =>
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   585
          add_path flat_names tname #>
19716
52c22fccdaaf add_datatype_axm: finalize specified consts;
wenzelm
parents: 19599
diff changeset
   586
            specify_consts (map (fn ((_, cargs), (cname, mx)) =>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   587
              (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   588
                (constrs ~~ constr_syntax')) #>
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   589
          parent_path flat_names)
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   590
            (hd descr ~~ newTs ~~ new_type_names ~~ constr_syntax);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   591
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   592
    (**** introduction of axioms ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   593
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   594
    val rec_axs = DatatypeProp.make_primrecs new_type_names descr sorts thy2;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   595
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   596
    val ((([induct], [rec_thms]), inject), thy3) =
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   597
      thy2
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   598
      |> Theory.add_path (space_implode "_" new_type_names)
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   599
      |> add_axiom "induct" (DatatypeProp.make_ind descr sorts) [case_names_induct]
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   600
      ||>> add_axioms "recs" rec_axs []
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   601
      ||> Theory.parent_path
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   602
      ||>> add_and_get_axiomss "inject" new_type_names
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   603
            (DatatypeProp.make_injs descr sorts);
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   604
    val (distinct, thy4) = add_and_get_axiomss "distinct" new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   605
      (DatatypeProp.make_distincts new_type_names descr sorts thy3) thy3;
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   606
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   607
    val exhaust_ts = DatatypeProp.make_casedists descr sorts;
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   608
    val (exhaustion, thy5) = add_and_get_axioms_atts "exhaust" new_type_names
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   609
      exhaust_ts (map single case_names_exhausts) thy4;
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   610
    val (case_thms, thy6) = add_and_get_axiomss "cases" new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   611
      (DatatypeProp.make_cases new_type_names descr sorts thy5) thy5;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   612
    val (split_ts, split_asm_ts) = ListPair.unzip
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   613
      (DatatypeProp.make_splits new_type_names descr sorts thy6);
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   614
    val (split, thy7) = add_and_get_axioms "split" new_type_names split_ts thy6;
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   615
    val (split_asm, thy8) = add_and_get_axioms "split_asm" new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   616
      split_asm_ts thy7;
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   617
    val (nchotomys, thy9) = add_and_get_axioms "nchotomy" new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   618
      (DatatypeProp.make_nchotomys descr sorts) thy8;
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   619
    val (case_congs, thy10) = add_and_get_axioms "case_cong" new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   620
      (DatatypeProp.make_case_congs new_type_names descr sorts thy9) thy9;
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   621
    val (weak_case_congs, thy11) = add_and_get_axioms "weak_case_cong" new_type_names
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   622
      (DatatypeProp.make_weak_case_congs new_type_names descr sorts thy10) thy10;
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   623
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   624
    val dt_infos = map (make_dt_info (length (hd descr)) descr' sorts induct
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   625
        reccomb_names' rec_thms)
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   626
      ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names' ~~ case_thms ~~
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   627
        exhaustion ~~ replicate (length (hd descr)) QuickAndDirty ~~ inject ~~
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   628
          nchotomys ~~ case_congs ~~ weak_case_congs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   629
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   630
    val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
9386
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   631
    val split_thms = split ~~ split_asm;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   632
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   633
    val thy12 =
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   634
      thy11
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   635
      |> add_case_tr' case_names'
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   636
      |> Theory.add_path (space_implode "_" new_type_names)
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   637
      |> add_rules simps case_thms rec_thms inject distinct
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   638
          weak_case_congs Simplifier.cong_add
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   639
      |> put_dt_infos dt_infos
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   640
      |> add_cases_induct dt_infos induct
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   641
      |> Theory.parent_path
19599
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19539
diff changeset
   642
      |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19539
diff changeset
   643
      |> snd
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   644
      |> DatatypeInterpretator.add_those [map fst dt_infos];
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   645
  in
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   646
    ({distinct = distinct,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   647
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   648
      exhaustion = exhaustion,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   649
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   650
      case_thms = case_thms,
9386
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   651
      split_thms = split_thms,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   652
      induction = induct,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   653
      simps = simps}, thy12)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   654
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   655
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   656
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   657
(******************* definitional introduction of datatypes *******************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   658
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   659
fun add_datatype_def flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   660
    case_names_induct case_names_exhausts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   661
  let
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   662
    val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   663
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   664
    val ((inject, distinct, dist_rewrites, simproc_dists, induct), thy2) = thy |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   665
      DatatypeRepProofs.representation_proofs flat_names dt_info new_type_names descr sorts
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   666
        types_syntax constr_syntax case_names_induct;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   667
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   668
    val (casedist_thms, thy3) = DatatypeAbsProofs.prove_casedist_thms new_type_names descr
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   669
      sorts induct case_names_exhausts thy2;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   670
    val ((reccomb_names, rec_thms), thy4) = DatatypeAbsProofs.prove_primrec_thms
20054
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   671
      flat_names new_type_names descr sorts dt_info inject dist_rewrites
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   672
      (Simplifier.theory_context thy3 dist_ss) induct thy3;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   673
    val ((case_thms, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   674
      flat_names new_type_names descr sorts reccomb_names rec_thms thy4;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   675
    val (split_thms, thy7) = DatatypeAbsProofs.prove_split_thms new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   676
      descr sorts inject dist_rewrites casedist_thms case_thms thy6;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   677
    val (nchotomys, thy8) = DatatypeAbsProofs.prove_nchotomys new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   678
      descr sorts casedist_thms thy7;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   679
    val (case_congs, thy9) = DatatypeAbsProofs.prove_case_congs new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   680
      descr sorts nchotomys case_thms thy8;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   681
    val (weak_case_congs, thy10) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   682
      descr sorts thy9;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   683
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   684
    val dt_infos = map (make_dt_info (length (hd descr)) (flat descr) sorts induct
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   685
        reccomb_names rec_thms)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   686
      ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   687
        casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   688
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   689
    val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   690
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   691
    val thy12 =
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   692
      thy10
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   693
      |> add_case_tr' case_names
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   694
      |> Theory.add_path (space_implode "_" new_type_names)
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   695
      |> add_rules simps case_thms rec_thms inject distinct
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   696
          weak_case_congs (Simplifier.attrib (op addcongs))
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   697
      |> put_dt_infos dt_infos
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   698
      |> add_cases_induct dt_infos induct
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   699
      |> Theory.parent_path
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   700
      |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms) |> snd
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   701
      |> DatatypeInterpretator.add_those [map fst dt_infos];
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   702
  in
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   703
    ({distinct = distinct,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   704
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   705
      exhaustion = casedist_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   706
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   707
      case_thms = case_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   708
      split_thms = split_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   709
      induction = induct,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   710
      simps = simps}, thy12)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   711
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   712
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   713
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   714
(*********************** declare existing type as datatype *********************)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   715
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   716
fun gen_rep_datatype apply_theorems alt_names raw_distinct raw_inject raw_induction thy0 =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   717
  let
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   718
    val (((distinct, inject), [induction]), thy1) =
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   719
      thy0
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   720
      |> fold_map apply_theorems raw_distinct
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   721
      ||>> fold_map apply_theorems raw_inject
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   722
      ||>> apply_theorems [raw_induction];
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   723
22598
f31a869077f0 rep_thm/cterm/ctyp: removed obsolete sign field;
wenzelm
parents: 22578
diff changeset
   724
    val ((_, [induction']), _) =
f31a869077f0 rep_thm/cterm/ctyp: removed obsolete sign field;
wenzelm
parents: 22578
diff changeset
   725
      Variable.importT_thms [induction] (Variable.thm_context induction);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   726
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   727
    fun err t = error ("Ill-formed predicate in induction rule: " ^
20715
c1f16b427d86 handling of \<^const> syntax for case; explicit case names for induction rules for rep_datatype
haftmann
parents: 20597
diff changeset
   728
      Sign.string_of_term thy1 t);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   729
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   730
    fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   731
          ((tname, map dest_TFree Ts) handle TERM _ => err t)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   732
      | get_typ t = err t;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   733
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   734
    val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induction')));
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   735
    val new_type_names = getOpt (alt_names, map fst dtnames);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   736
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   737
    fun get_constr t = (case Logic.strip_assums_concl t of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   738
        _ $ (_ $ t') => (case head_of t' of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   739
            Const (cname, cT) => (case strip_type cT of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   740
                (Ts, Type (tname, _)) => (tname, (cname, map (dtyp_of_typ dtnames) Ts))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   741
              | _ => err t)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   742
          | _ => err t)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   743
      | _ => err t);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   744
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   745
    fun make_dt_spec [] _ _ = []
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   746
      | make_dt_spec ((tname, tvs)::dtnames') i constrs =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   747
          let val (constrs', constrs'') = take_prefix (equal tname o fst) constrs
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   748
          in (i, (tname, map DtTFree tvs, map snd constrs'))::
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   749
            (make_dt_spec dtnames' (i + 1) constrs'')
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   750
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   751
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   752
    val descr = make_dt_spec dtnames 0 (map get_constr (prems_of induction'));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   753
    val sorts = add_term_tfrees (concl_of induction', []);
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   754
    val dt_info = get_datatypes thy1;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   755
21419
809e7520234a added instance for class size
haftmann
parents: 21350
diff changeset
   756
    val (case_names_induct, case_names_exhausts) =
809e7520234a added instance for class size
haftmann
parents: 21350
diff changeset
   757
      (mk_case_names_induct descr, mk_case_names_exhausts descr (map #1 dtnames));
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   758
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6423
diff changeset
   759
    val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   760
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   761
    val (casedist_thms, thy2) = thy1 |>
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   762
      DatatypeAbsProofs.prove_casedist_thms new_type_names [descr] sorts induction
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   763
        case_names_exhausts;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   764
    val ((reccomb_names, rec_thms), thy3) = DatatypeAbsProofs.prove_primrec_thms
20054
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   765
      false new_type_names [descr] sorts dt_info inject distinct
24d176b8f240 distinct simproc/simpset: proper context;
wenzelm
parents: 19925
diff changeset
   766
      (Simplifier.theory_context thy2 dist_ss) induction thy2;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   767
    val ((case_thms, case_names), thy4) = DatatypeAbsProofs.prove_case_thms false
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   768
      new_type_names [descr] sorts reccomb_names rec_thms thy3;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   769
    val (split_thms, thy5) = DatatypeAbsProofs.prove_split_thms
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   770
      new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   771
    val (nchotomys, thy6) = DatatypeAbsProofs.prove_nchotomys new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   772
      [descr] sorts casedist_thms thy5;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   773
    val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   774
      [descr] sorts nchotomys case_thms thy6;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   775
    val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   776
      [descr] sorts thy7;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   777
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   778
    val ((_, [induction']), thy10) =
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   779
      thy8
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   780
      |> store_thmss "inject" new_type_names inject
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   781
      ||>> store_thmss "distinct" new_type_names distinct
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   782
      ||> Theory.add_path (space_implode "_" new_type_names)
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   783
      ||>> PureThy.add_thms [(("induct", induction), [case_names_induct])];
9149
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   784
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   785
    val dt_infos = map (make_dt_info (length descr) descr sorts induction'
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   786
        reccomb_names rec_thms)
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   787
      ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   788
        map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   789
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   790
    val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   791
19599
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19539
diff changeset
   792
    val thy11 =
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19539
diff changeset
   793
      thy10
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   794
      |> add_case_tr' case_names
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   795
      |> add_rules simps case_thms rec_thms inject distinct
19599
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19539
diff changeset
   796
           weak_case_congs (Simplifier.attrib (op addcongs))
22777
2fc921376a86 - Moved parse / print translations for case to datatype_case.ML
berghofe
parents: 22709
diff changeset
   797
      |> put_dt_infos dt_infos
19599
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19539
diff changeset
   798
      |> add_cases_induct dt_infos induction'
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19539
diff changeset
   799
      |> Theory.parent_path
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19539
diff changeset
   800
      |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)
a5c7eb37d14f added DatatypeHooks
haftmann
parents: 19539
diff changeset
   801
      |> snd
24699
c6674504103f datatype interpretators for size and datatype_realizer
haftmann
parents: 24626
diff changeset
   802
      |> DatatypeInterpretator.add_those [map fst dt_infos];
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   803
  in
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   804
    ({distinct = distinct,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   805
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   806
      exhaustion = casedist_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   807
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   808
      case_thms = case_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   809
      split_thms = split_thms,
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   810
      induction = induction',
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   811
      simps = simps}, thy11)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   812
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   813
21350
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21253
diff changeset
   814
val rep_datatype = gen_rep_datatype IsarCmd.apply_theorems;
6e58289b6685 incorporated IsarThy into IsarCmd;
wenzelm
parents: 21253
diff changeset
   815
val rep_datatype_i = gen_rep_datatype IsarCmd.apply_theorems_i;
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   816
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   817
11958
2ece34b9fd8e Isar: fixed rep_datatype args;
wenzelm
parents: 11805
diff changeset
   818
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   819
(******************************** add datatype ********************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   820
14887
4938ce4ef295 Added exception Datatype_Empty.
berghofe
parents: 14799
diff changeset
   821
fun gen_add_datatype prep_typ err flat_names new_type_names dts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   822
  let
20820
58693343905f removed obsolete Datatype_Universe.thy (cf. Datatype.thy);
wenzelm
parents: 20715
diff changeset
   823
    val _ = Theory.requires thy "Datatype" "datatype definitions";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   824
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   825
    (* this theory is used just for parsing *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   826
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   827
    val tmp_thy = thy |>
5892
e5e44cc54eb2 Attribute.tthms_of;
wenzelm
parents: 5720
diff changeset
   828
      Theory.copy |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   829
      Theory.add_types (map (fn (tvs, tname, mx, _) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   830
        (tname, length tvs, mx)) dts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   831
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   832
    val (tyvars, _, _, _)::_ = dts;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   833
    val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   834
      let val full_tname = Sign.full_name tmp_thy (Syntax.type_name tname mx)
18964
67f572e03236 renamed gen_duplicates to duplicates;
wenzelm
parents: 18963
diff changeset
   835
      in (case duplicates (op =) tvs of
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   836
            [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   837
                  else error ("Mutually recursive datatypes must have same type parameters")
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   838
          | dups => error ("Duplicate parameter(s) for datatype " ^ full_tname ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   839
              " : " ^ commas dups))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   840
      end) dts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   841
18964
67f572e03236 renamed gen_duplicates to duplicates;
wenzelm
parents: 18963
diff changeset
   842
    val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   843
      [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   844
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   845
    fun prep_dt_spec (tvs, tname, mx, constrs) (dts', constr_syntax, sorts, i) =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   846
      let
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   847
        fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
5279
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   848
          let
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   849
            val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   850
            val _ = (case fold (curry add_typ_tfree_names) cargs' [] \\ tvs of
5279
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   851
                [] => ()
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   852
              | vs => error ("Extra type variables on rhs: " ^ commas vs))
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   853
          in (constrs @ [((if flat_names then Sign.full_name tmp_thy else
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   854
                Sign.full_name_path tmp_thy tname) (Syntax.const_name cname mx'),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   855
                   map (dtyp_of_typ new_dts) cargs')],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   856
              constr_syntax' @ [(cname, mx')], sorts'')
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18643
diff changeset
   857
          end handle ERROR msg =>
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18643
diff changeset
   858
            cat_error msg ("The error above occured in constructor " ^ cname ^
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   859
              " of datatype " ^ tname);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   860
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   861
        val (constrs', constr_syntax', sorts') =
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   862
          fold prep_constr constrs ([], [], sorts)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   863
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   864
      in
18964
67f572e03236 renamed gen_duplicates to duplicates;
wenzelm
parents: 18963
diff changeset
   865
        case duplicates (op =) (map fst constrs') of
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   866
           [] =>
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   867
             (dts' @ [(i, (Sign.full_name tmp_thy (Syntax.type_name tname mx),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   868
                map DtTFree tvs, constrs'))],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   869
              constr_syntax @ [constr_syntax'], sorts', i + 1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   870
         | dups => error ("Duplicate constructors " ^ commas dups ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   871
             " in datatype " ^ tname)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   872
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   873
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   874
    val (dts', constr_syntax, sorts', i) = fold prep_dt_spec dts ([], [], [], 0);
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   875
    val sorts = sorts' @ (map (rpair (Sign.defaultS tmp_thy)) (tyvars \\ map fst sorts'));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   876
    val dt_info = get_datatypes thy;
20597
65fe827aa595 code generation 2 adjustments
haftmann
parents: 20328
diff changeset
   877
    val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
14887
4938ce4ef295 Added exception Datatype_Empty.
berghofe
parents: 14799
diff changeset
   878
    val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
15661
9ef583b08647 reverted renaming of Some/None in comments and strings;
wenzelm
parents: 15574
diff changeset
   879
      if err then error ("Nonemptiness check failed for datatype " ^ s)
14887
4938ce4ef295 Added exception Datatype_Empty.
berghofe
parents: 14799
diff changeset
   880
      else raise exn;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   881
21045
66d6d1b0ddfa slight cleanup
haftmann
parents: 20820
diff changeset
   882
    val descr' = flat descr;
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   883
    val case_names_induct = mk_case_names_induct descr';
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   884
    val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   885
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   886
    (if (!quick_and_dirty) then add_datatype_axm else add_datatype_def)
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   887
      flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   888
      case_names_induct case_names_exhausts thy
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   889
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   890
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   891
val add_datatype_i = gen_add_datatype cert_typ;
14887
4938ce4ef295 Added exception Datatype_Empty.
berghofe
parents: 14799
diff changeset
   892
val add_datatype = gen_add_datatype read_typ true;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   893
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   894
18451
5ff0244e25e8 slight clean ups
haftmann
parents: 18418
diff changeset
   895
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   896
(** package setup **)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   897
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   898
(* setup theory *)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   899
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18702
diff changeset
   900
val setup =
24098
f1eb34ae33af replaced dtK ref by datatype_distinctness_limit configuration option;
wenzelm
parents: 22994
diff changeset
   901
  DatatypeProp.distinctness_limit_setup #>
f1eb34ae33af replaced dtK ref by datatype_distinctness_limit configuration option;
wenzelm
parents: 22994
diff changeset
   902
  Method.add_methods tactic_emulations #>
f1eb34ae33af replaced dtK ref by datatype_distinctness_limit configuration option;
wenzelm
parents: 22994
diff changeset
   903
  simproc_setup #>
24626
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
   904
  trfun_setup #>
85eceef2edc7 introduced generic concepts for theory interpretators
haftmann
parents: 24624
diff changeset
   905
  DatatypeInterpretator.init;
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   906
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   907
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   908
(* outer syntax *)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   909
17057
0934ac31985f OuterKeyword;
wenzelm
parents: 16973
diff changeset
   910
local structure P = OuterParse and K = OuterKeyword in
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   911
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   912
val datatype_decl =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   913
  Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12708
diff changeset
   914
    (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   915
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   916
fun mk_datatype args =
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   917
  let
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   918
    val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args;
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12708
diff changeset
   919
    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
   920
      (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   921
  in snd o add_datatype false names specs end;
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   922
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   923
val datatypeP =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   924
  OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   925
    (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   926
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   927
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   928
val rep_datatype_decl =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   929
  Scan.option (Scan.repeat1 P.name) --
22101
6d13239d5f52 moved parts of OuterParse to SpecParse;
wenzelm
parents: 21772
diff changeset
   930
    Scan.optional (P.$$$ "distinct" |-- P.!!! (P.and_list1 SpecParse.xthms1)) [[]] --
6d13239d5f52 moved parts of OuterParse to SpecParse;
wenzelm
parents: 21772
diff changeset
   931
    Scan.optional (P.$$$ "inject" |-- P.!!! (P.and_list1 SpecParse.xthms1)) [[]] --
6d13239d5f52 moved parts of OuterParse to SpecParse;
wenzelm
parents: 21772
diff changeset
   932
    (P.$$$ "induction" |-- P.!!! SpecParse.xthm);
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   933
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   934
fun mk_rep_datatype (((opt_ts, dss), iss), ind) = #2 o rep_datatype opt_ts dss iss ind;
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   935
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   936
val rep_datatypeP =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   937
  OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_decl
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   938
    (rep_datatype_decl >> (Toplevel.theory o mk_rep_datatype));
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   939
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   940
6479
b0448cab1b1e rep_datatype syntax: 'induction' instead of 'induct';
wenzelm
parents: 6441
diff changeset
   941
val _ = OuterSyntax.add_keywords ["distinct", "inject", "induction"];
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   942
val _ = OuterSyntax.add_parsers [datatypeP, rep_datatypeP];
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   943
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   944
end;
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   945
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   946
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   947
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   948
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   949
structure BasicDatatypePackage: BASIC_DATATYPE_PACKAGE = DatatypePackage;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   950
open BasicDatatypePackage;
15704
93163972dbdc *** MESSAGE REFERS TO PREVIOUS VERSION ***
wenzelm
parents: 15703
diff changeset
   951