src/HOL/Tools/datatype_package.ML
author wenzelm
Sun, 15 Jan 2006 19:58:53 +0100
changeset 18690 16a64bdc5485
parent 18688 abf0f018b5ec
child 18702 7dc7dcd63224
permissions -rw-r--r--
classical attributes: optional weight;
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
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    20
  val add_datatype : bool -> string list -> (string list * bstring * mixfix *
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    21
    (bstring * string list * mixfix) list) list -> theory ->
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    22
      {distinct : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    23
       inject : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    24
       exhaustion : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    25
       rec_thms : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    26
       case_thms : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    27
       split_thms : (thm * thm) list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    28
       induction : thm,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    29
       size : thm list,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    30
       simps : thm list} * theory
14887
4938ce4ef295 Added exception Datatype_Empty.
berghofe
parents: 14799
diff changeset
    31
  val add_datatype_i : bool -> bool -> string list -> (string list * bstring * mixfix *
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    32
    (bstring * typ list * mixfix) list) list -> theory ->
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    33
      {distinct : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    34
       inject : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    35
       exhaustion : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    36
       rec_thms : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    37
       case_thms : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    38
       split_thms : (thm * thm) list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    39
       induction : thm,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    40
       size : thm list,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    41
       simps : thm list} * theory
12708
31672377dadc clarified IsarThy.apply_theorems_i;
wenzelm
parents: 12448
diff changeset
    42
  val rep_datatype_i : string list option -> (thm list * theory attribute list) list list ->
31672377dadc clarified IsarThy.apply_theorems_i;
wenzelm
parents: 12448
diff changeset
    43
    (thm list * theory attribute list) list list -> (thm list * theory attribute list) ->
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    44
    theory ->
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    45
      {distinct : thm list list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    46
       inject : thm list list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    47
       exhaustion : thm list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    48
       rec_thms : thm list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    49
       case_thms : thm list list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    50
       split_thms : (thm * thm) list,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    51
       induction : thm,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    52
       size : thm list,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    53
       simps : thm list} * theory
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15661
diff changeset
    54
  val rep_datatype : string list option -> (thmref * Attrib.src list) list list ->
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    55
    (thmref * Attrib.src list) list list -> thmref * Attrib.src list -> theory ->
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    56
      {distinct : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    57
       inject : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    58
       exhaustion : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    59
       rec_thms : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    60
       case_thms : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    61
       split_thms : (thm * thm) list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    62
       induction : thm,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    63
       size : thm list,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
    64
       simps : thm list} * theory
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    65
  val get_datatypes : theory -> DatatypeAux.datatype_info Symtab.table
6441
268aa3c4a059 print_datatypes;
wenzelm
parents: 6427
diff changeset
    66
  val print_datatypes : theory -> unit
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
    67
  val datatype_info : theory -> string -> DatatypeAux.datatype_info option
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
    68
  val datatype_info_err : theory -> string -> DatatypeAux.datatype_info
18451
5ff0244e25e8 slight clean ups
haftmann
parents: 18418
diff changeset
    69
  val is_datatype : theory -> string -> bool
5ff0244e25e8 slight clean ups
haftmann
parents: 18418
diff changeset
    70
  val get_datatype : theory -> string -> ((string * sort) list * string list) option
5ff0244e25e8 slight clean ups
haftmann
parents: 18418
diff changeset
    71
  val get_datatype_cons : theory -> string * string -> typ list option
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
    72
  val get_datatype_case_consts : theory -> string list
18451
5ff0244e25e8 slight clean ups
haftmann
parents: 18418
diff changeset
    73
  val get_case_const_data : theory -> string -> (string * int) list option
18455
b293c1087f1d slight improvements
haftmann
parents: 18451
diff changeset
    74
  val get_all_datatype_cons : theory -> (string * string) list
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    75
  val constrs_of : theory -> string -> term list option
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    76
  val case_const_of : theory -> string -> term option
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
    77
  val weak_case_congs_of : theory -> thm list
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    78
  val setup: (theory -> theory) list
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    79
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    80
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    81
structure DatatypePackage : DATATYPE_PACKAGE =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    82
struct
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    83
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    84
open DatatypeAux;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    85
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    86
val quiet_mode = quiet_mode;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    87
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    88
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    89
(* data kind 'HOL/datatypes' *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    90
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    91
structure DatatypesData = TheoryDataFun
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    92
(struct
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    93
  val name = "HOL/datatypes";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    94
  type T = datatype_info Symtab.table;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    95
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    96
  val empty = Symtab.empty;
6556
daa00919502b theory data: copy;
wenzelm
parents: 6479
diff changeset
    97
  val copy = I;
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    98
  val extend = I;
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
    99
  fun merge _ tabs : T = Symtab.merge (K true) tabs;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   100
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   101
  fun print sg tab =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   102
    Pretty.writeln (Pretty.strs ("datatypes:" ::
16364
dc9f7066d80a refer to name spaces values instead of names;
wenzelm
parents: 16122
diff changeset
   103
      map #1 (NameSpace.extern_table (Sign.type_space sg, tab))));
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   104
end);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   105
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   106
val get_datatypes = DatatypesData.get;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   107
val put_datatypes = DatatypesData.put;
6441
268aa3c4a059 print_datatypes;
wenzelm
parents: 6427
diff changeset
   108
val print_datatypes = DatatypesData.print;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   109
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   110
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   111
(** theory information about datatypes **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   112
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17325
diff changeset
   113
val datatype_info = Symtab.lookup o get_datatypes;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   114
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   115
fun datatype_info_err thy name = (case datatype_info thy name of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   116
      SOME info => info
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   117
    | NONE => error ("Unknown datatype " ^ quote name));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   118
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   119
fun constrs_of thy tname = (case datatype_info thy tname of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   120
   SOME {index, descr, ...} =>
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   121
     let val (_, _, constrs) = valOf (AList.lookup (op =) descr index)
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   122
     in SOME (map (fn (cname, _) => Const (cname, Sign.the_const_type thy cname)) constrs)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   123
     end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   124
 | _ => NONE);
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   125
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   126
fun case_const_of thy tname = (case datatype_info thy tname of
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   127
   SOME {case_name, ...} => SOME (Const (case_name, Sign.the_const_type thy case_name))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   128
 | _ => NONE);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   129
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   130
val weak_case_congs_of = map (#weak_case_cong o #2) o Symtab.dest o get_datatypes;
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   131
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   132
fun is_datatype thy dtco =
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   133
  Symtab.lookup (get_datatypes thy) dtco
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   134
  |> is_some;
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   135
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   136
fun get_datatype thy dtco =
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   137
  dtco
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   138
  |> Symtab.lookup (get_datatypes thy)
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   139
  |> Option.map (fn info => (#sorts info,
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   140
      (get_first (fn (_, (dtco', _, cs)) =>
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   141
        if dtco = dtco'
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   142
        then SOME (map fst cs)
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   143
        else NONE) (#descr info) |> the)));
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   144
  
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   145
fun get_datatype_cons thy (co, dtco) =
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   146
  let
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   147
    val descr =
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   148
      dtco
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   149
      |> Symtab.lookup (get_datatypes thy)
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   150
      |> Option.map #descr
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   151
      |> these
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   152
    val one_descr =
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   153
      descr
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   154
      |> get_first (fn (_, (dtco', vs, cs)) =>
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   155
          if dtco = dtco'
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   156
          then SOME (vs, cs)
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   157
          else NONE);
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   158
  in
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   159
    if is_some one_descr
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   160
    then
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   161
      the one_descr
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   162
      |> (fn (vs, cs) =>
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   163
          co
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   164
          |> AList.lookup (op =) cs
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   165
          |> Option.map (map (DatatypeAux.typ_of_dtyp descr (map (rpair [])
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   166
               (map DatatypeAux.dest_DtTFree vs)))))
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   167
    else NONE
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   168
  end;
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   169
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   170
fun get_datatype_case_consts thy =
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   171
  Symtab.fold (fn (_, {case_name, ...}) => cons case_name) (get_datatypes thy) [];
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   172
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   173
fun get_case_const_data thy c =
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   174
  case find_first (fn (_, {index, descr, case_name, ...}) =>
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   175
      case_name = c
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   176
    ) ((Symtab.dest o get_datatypes) thy)
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   177
   of NONE => NONE
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   178
    | SOME (_, {index, descr, ...}) =>
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   179
        (SOME o map (apsnd length) o #3 o the o AList.lookup (op =) descr) index;
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   180
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   181
fun get_all_datatype_cons thy =
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   182
  Symtab.fold (fn (dtco, _) => fold
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   183
    (fn co => cons (co, dtco))
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   184
      ((snd o the oo get_datatype) thy dtco)) (get_datatypes thy) [];
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   185
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   186
fun find_tname var Bi =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   187
  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
   188
      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
   189
  in case AList.lookup (op =) (frees @ params) var of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   190
       NONE => error ("No such variable in subgoal: " ^ quote var)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   191
     | SOME(Type (tn, _)) => tn
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   192
     | _ => error ("Cannot determine type of " ^ quote var)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   193
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   194
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   195
fun infer_tname state i aterm =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   196
  let
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   197
    val sign = Thm.sign_of_thm state;
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   198
    val (_, _, Bi, _) = Thm.dest_state (state, i)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   199
    val params = Logic.strip_params Bi;   (*params of subgoal i*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   200
    val params = rev (rename_wrt_term Bi params);   (*as they are printed*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   201
    val (types, sorts) = types_sorts state;
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   202
    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
   203
      | types' ixn = types ixn;
8100
6186ee807f2e replaced HOLogic.termTVar by HOLogic.termT;
wenzelm
parents: 7904
diff changeset
   204
    val (ct, _) = read_def_cterm (sign, types', sorts) [] false (aterm, TypeInfer.logicT);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   205
  in case #T (rep_cterm ct) of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   206
       Type (tn, _) => tn
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   207
     | _ => error ("Cannot determine type of " ^ quote aterm)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   208
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   209
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   210
(*Warn if the (induction) variable occurs Free among the premises, which
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   211
  usually signals a mistake.  But calls the tactic either way!*)
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   212
fun occs_in_prems tacf vars =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   213
  SUBGOAL (fn (Bi, i) =>
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   214
           (if  exists (fn Free (a, _) => a mem vars)
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   215
                      (foldr add_term_frees [] (#2 (strip_context Bi)))
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   216
             then warning "Induction variable occurs also among premises!"
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   217
             else ();
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   218
            tacf i));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   219
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   220
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   221
(* generic induction tactic for datatypes *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   222
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   223
local
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   224
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   225
fun prep_var (Var (ixn, _), SOME x) = SOME (ixn, x)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   226
  | prep_var _ = NONE;
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   227
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   228
fun prep_inst (concl, xs) =	(*exception UnequalLengths *)
11725
d0c37d04906b HOLogic.dest_concls, InductAttrib.vars_of;
wenzelm
parents: 11539
diff changeset
   229
  let val vs = InductAttrib.vars_of concl
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   230
  in List.mapPartial prep_var (Library.drop (length vs - length xs, vs) ~~ xs) end;
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   231
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   232
in
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   233
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   234
fun gen_induct_tac inst_tac (varss, opt_rule) i state =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   235
  let
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   236
    val (_, _, Bi, _) = Thm.dest_state (state, i);
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   237
    val {sign, ...} = Thm.rep_thm state;
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   238
    val (rule, rule_name) =
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   239
      (case opt_rule of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   240
        SOME r => (r, "Induction rule")
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   241
      | NONE =>
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   242
          let val tn = find_tname (hd (List.mapPartial I (List.concat varss))) Bi
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   243
          in (#induction (datatype_info_err sign tn), "Induction rule for type " ^ tn) end);
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   244
11725
d0c37d04906b HOLogic.dest_concls, InductAttrib.vars_of;
wenzelm
parents: 11539
diff changeset
   245
    val concls = HOLogic.dest_concls (Thm.concl_of rule);
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   246
    val insts = List.concat (map prep_inst (concls ~~ varss)) handle UnequalLengths =>
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   247
      error (rule_name ^ " has different numbers of variables");
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   248
  in occs_in_prems (inst_tac insts rule) (map #2 insts) i state end;
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   249
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   250
fun induct_tac s =
15444
4f14c151d9f1 induct_tac and case_tac no longer depend on Syntax.string_of_vname.
berghofe
parents: 14981
diff changeset
   251
  gen_induct_tac Tactic.res_inst_tac'
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   252
    (map (Library.single o SOME) (Syntax.read_idents s), NONE);
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   253
9747
043098ba5098 introduced induct_thm_tac
nipkow
parents: 9739
diff changeset
   254
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
   255
  gen_induct_tac Tactic.res_inst_tac'
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   256
    ([map SOME (Syntax.read_idents s)], SOME th);
9747
043098ba5098 introduced induct_thm_tac
nipkow
parents: 9739
diff changeset
   257
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   258
end;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   259
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   260
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   261
(* generic case tactic for datatypes *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   262
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   263
fun case_inst_tac inst_tac t rule i state =
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   264
  let
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   265
    val _ $ Var (ixn, _) $ _ = HOLogic.dest_Trueprop
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   266
      (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
   267
  in inst_tac [(ixn, t)] rule i state end;
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   268
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   269
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
   270
      case_inst_tac inst_tac t rule i state
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   271
  | gen_case_tac inst_tac (t, NONE) i state =
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   272
      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
   273
        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
   274
        else case_inst_tac inst_tac t
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   275
               (#exhaustion (datatype_info_err (Thm.sign_of_thm state) tn))
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   276
               i state
14471
5688b05b2575 case_tac no longer raises THM exception if goal number is out of range.
berghofe
parents: 14412
diff changeset
   277
      end handle THM _ => Seq.empty;
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   278
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   279
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
   280
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   281
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   282
8541
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   283
(** Isar tactic emulations **)
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   284
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   285
local
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   286
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   287
val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.$$$ ":");
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   288
val opt_rule = Scan.option (rule_spec |-- Attrib.local_thm);
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   289
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15661
diff changeset
   290
val varss =
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15661
diff changeset
   291
  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
   292
15444
4f14c151d9f1 induct_tac and case_tac no longer depend on Syntax.string_of_vname.
berghofe
parents: 14981
diff changeset
   293
val inst_tac = Method.bires_inst_tac false;
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   294
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   295
fun induct_meth ctxt (varss, opt_rule) =
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   296
  gen_induct_tac (inst_tac ctxt) (varss, opt_rule);
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   297
fun case_meth ctxt (varss, opt_rule) =
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   298
  gen_case_tac (inst_tac ctxt) (varss, opt_rule);
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   299
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   300
in
8541
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   301
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   302
val tactic_emulations =
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   303
 [("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
   304
    "induct_tac emulation (dynamic instantiation)"),
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13641
diff changeset
   305
  ("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
   306
    "case_tac emulation (dynamic instantiation)")];
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   307
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   308
end;
8541
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   309
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   310
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   311
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   312
(** induct method setup **)
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   313
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   314
(* case names *)
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   315
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   316
local
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   317
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   318
fun dt_recs (DtTFree _) = []
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   319
  | dt_recs (DtType (_, dts)) = List.concat (map dt_recs dts)
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   320
  | dt_recs (DtRec i) = [i];
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   321
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   322
fun dt_cases (descr: descr) (_, args, constrs) =
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   323
  let
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   324
    fun the_bname i = Sign.base_name (#1 (valOf (AList.lookup (op =) descr i)));
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   325
    val bnames = map the_bname (distinct (List.concat (map dt_recs args)));
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   326
  in map (fn (c, _) => space_implode "_" (Sign.base_name c :: bnames)) constrs end;
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   327
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   328
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   329
fun induct_cases descr =
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   330
  DatatypeProp.indexify_names (List.concat (map (dt_cases descr) (map #2 descr)));
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   331
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   332
fun exhaust_cases descr i = dt_cases descr (valOf (AList.lookup (op =) descr i));
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   333
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   334
in
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   335
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   336
fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   337
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   338
fun mk_case_names_exhausts descr new =
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   339
  map (RuleCases.case_names o exhaust_cases descr o #1)
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   340
    (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr);
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   341
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   342
end;
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   343
18625
2db0982523fe _E suffix for compatibility with AddIffs
paulson
parents: 18518
diff changeset
   344
(*Name management for ATP linkup. The suffix here must agree with the one given
2db0982523fe _E suffix for compatibility with AddIffs
paulson
parents: 18518
diff changeset
   345
  for notE in Clasimp.addIff*)
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 17057
diff changeset
   346
fun name_notE th =
18625
2db0982523fe _E suffix for compatibility with AddIffs
paulson
parents: 18518
diff changeset
   347
    Thm.name_thm (Thm.name_of_thm th ^ "_iff1", th RS notE);
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 17057
diff changeset
   348
      
11345
cd605c85e421 improved iff_add_global, new function add_rules factoring out common behaviour
oheimb
parents: 10930
diff changeset
   349
fun add_rules simps case_thms size_thms rec_thms inject distinct
cd605c85e421 improved iff_add_global, new function add_rules factoring out common behaviour
oheimb
parents: 10930
diff changeset
   350
                  weak_case_congs cong_att =
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   351
  (snd o PureThy.add_thmss [(("simps", simps), []),
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   352
    (("", List.concat case_thms @ size_thms @ 
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18678
diff changeset
   353
          List.concat distinct  @ rec_thms), [Attrib.theory Simplifier.simp_add]),
16672
f83f3aef274d (intermediate commit)
haftmann
parents: 16646
diff changeset
   354
    (("", size_thms @ rec_thms),             [RecfunCodegen.add NONE]),
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18678
diff changeset
   355
    (("", List.concat inject),               [Attrib.theory iff_add]),
18690
16a64bdc5485 classical attributes: optional weight;
wenzelm
parents: 18688
diff changeset
   356
    (("", map name_notE (List.concat distinct)),  [Attrib.theory (Classical.safe_elim NONE)]),
16672
f83f3aef274d (intermediate commit)
haftmann
parents: 16646
diff changeset
   357
    (("", weak_case_congs),                  [cong_att])]);
11345
cd605c85e421 improved iff_add_global, new function add_rules factoring out common behaviour
oheimb
parents: 10930
diff changeset
   358
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   359
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   360
(* add_cases_induct *)
8306
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   361
11805
b110a1ea90da declare projected induction rules stemming from nested recursion;
wenzelm
parents: 11725
diff changeset
   362
fun add_cases_induct infos induction =
8306
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   363
  let
11805
b110a1ea90da declare projected induction rules stemming from nested recursion;
wenzelm
parents: 11725
diff changeset
   364
    val n = length (HOLogic.dest_concls (Thm.concl_of induction));
18462
b67d423b5234 *.inducts holds all projected rules;
wenzelm
parents: 18455
diff changeset
   365
    fun proj i = ProjectRule.project induction (i + 1);
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   366
11805
b110a1ea90da declare projected induction rules stemming from nested recursion;
wenzelm
parents: 11725
diff changeset
   367
    fun named_rules (name, {index, exhaustion, ...}: datatype_info) =
18643
89a7978f90e1 generic attributes;
wenzelm
parents: 18625
diff changeset
   368
      [(("", proj index), [Attrib.theory (InductAttrib.induct_type name)]),
89a7978f90e1 generic attributes;
wenzelm
parents: 18625
diff changeset
   369
       (("", exhaustion), [Attrib.theory (InductAttrib.cases_type name)])];
11805
b110a1ea90da declare projected induction rules stemming from nested recursion;
wenzelm
parents: 11725
diff changeset
   370
    fun unnamed_rule i =
18643
89a7978f90e1 generic attributes;
wenzelm
parents: 18625
diff changeset
   371
      (("", proj i), [Drule.kind_internal, Attrib.theory (InductAttrib.induct_type "")]);
18462
b67d423b5234 *.inducts holds all projected rules;
wenzelm
parents: 18455
diff changeset
   372
  in
b67d423b5234 *.inducts holds all projected rules;
wenzelm
parents: 18455
diff changeset
   373
    PureThy.add_thms
b67d423b5234 *.inducts holds all projected rules;
wenzelm
parents: 18455
diff changeset
   374
      (List.concat (map named_rules infos) @
b67d423b5234 *.inducts holds all projected rules;
wenzelm
parents: 18455
diff changeset
   375
        map unnamed_rule (length infos upto n - 1)) #> snd #>
b67d423b5234 *.inducts holds all projected rules;
wenzelm
parents: 18455
diff changeset
   376
    PureThy.add_thmss [(("inducts",
b67d423b5234 *.inducts holds all projected rules;
wenzelm
parents: 18455
diff changeset
   377
      map (proj #> standard #> RuleCases.save induction) (0 upto n - 1)), [])] #> snd
b67d423b5234 *.inducts holds all projected rules;
wenzelm
parents: 18455
diff changeset
   378
  end;
8306
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   379
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   380
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   381
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   382
(**** simplification procedure for showing distinctness of constructors ****)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   383
7060
berghofe
parents: 7049
diff changeset
   384
fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
berghofe
parents: 7049
diff changeset
   385
  | stripT p = p;
berghofe
parents: 7049
diff changeset
   386
berghofe
parents: 7049
diff changeset
   387
fun stripC (i, f $ x) = stripC (i + 1, f)
berghofe
parents: 7049
diff changeset
   388
  | stripC p = p;
berghofe
parents: 7049
diff changeset
   389
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   390
val distinctN = "constr_distinct";
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   391
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   392
exception ConstrDistinct of term;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   393
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16672
diff changeset
   394
fun distinct_proc sg ss (t as Const ("op =", _) $ t1 $ t2) =
7060
berghofe
parents: 7049
diff changeset
   395
  (case (stripC (0, t1), stripC (0, t2)) of
berghofe
parents: 7049
diff changeset
   396
     ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
berghofe
parents: 7049
diff changeset
   397
         (case (stripT (0, T1), stripT (0, T2)) of
berghofe
parents: 7049
diff changeset
   398
            ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
berghofe
parents: 7049
diff changeset
   399
                if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   400
                   (case (constrs_of sg tname1) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   401
                      SOME constrs => let val cnames = map (fst o dest_Const) constrs
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   402
                        in if cname1 mem cnames andalso cname2 mem cnames then
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   403
                             let val eq_t = Logic.mk_equals (t, Const ("False", HOLogic.boolT));
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   404
                                 val eq_ct = cterm_of sg eq_t;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   405
                                 val Datatype_thy = theory "Datatype";
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   406
                                 val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
16486
1a12cdb6ee6b get_thm(s): Name;
wenzelm
parents: 16430
diff changeset
   407
                                   map (get_thm Datatype_thy o Name)
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   408
                                     ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"]
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   409
                             in (case (#distinct (datatype_info_err sg tname1)) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   410
                                 QuickAndDirty => SOME (Thm.invoke_oracle
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   411
                                   Datatype_thy distinctN (sg, ConstrDistinct eq_t))
17956
369e2af8ee45 Goal.prove;
wenzelm
parents: 17875
diff changeset
   412
                               | FewConstrs thms => SOME (Goal.prove sg [] [] eq_t (K
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13466
diff changeset
   413
                                   (EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13466
diff changeset
   414
                                    atac 2, resolve_tac thms 1, etac FalseE 1])))
17956
369e2af8ee45 Goal.prove;
wenzelm
parents: 17875
diff changeset
   415
                               | ManyConstrs (thm, simpset) => SOME (Goal.prove sg [] [] eq_t (K
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13466
diff changeset
   416
                                   (EVERY [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
17875
d81094515061 change_claset/simpset;
wenzelm
parents: 17521
diff changeset
   417
                                    full_simp_tac (Simplifier.inherit_context ss simpset) 1,
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   418
                                    REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   419
                                    eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13466
diff changeset
   420
                                    etac FalseE 1]))))
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   421
                             end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   422
                           else NONE
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   423
                        end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   424
                    | NONE => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   425
                else NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   426
          | _ => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   427
   | _ => NONE)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   428
  | distinct_proc sg _ _ = NONE;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   429
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13340
diff changeset
   430
val distinct_simproc =
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   431
  Simplifier.simproc HOL.thy distinctN ["s = t"] distinct_proc;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   432
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   433
val dist_ss = HOL_ss addsimprocs [distinct_simproc];
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   434
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   435
val simproc_setup =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   436
  [Theory.add_oracle (distinctN, fn (_, ConstrDistinct t) => t),
17875
d81094515061 change_claset/simpset;
wenzelm
parents: 17521
diff changeset
   437
   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
   438
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   439
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   440
(**** translation rules for case ****)
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   441
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   442
fun find_first f = Library.find_first f;
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   443
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   444
fun case_tr sg [t, u] =
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   445
    let
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   446
      fun case_error s name ts = raise TERM ("Error in case expression" ^
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   447
        getOpt (Option.map (curry op ^ " for datatype ") name, "") ^ ":\n" ^ s, ts);
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   448
      fun dest_case1 (Const ("_case1", _) $ t $ u) = (case strip_comb t of
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   449
            (Const (s, _), ts) => (Sign.intern_const sg s, ts)
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   450
          | (Free (s, _), ts) => (Sign.intern_const sg s, ts)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   451
          | _ => case_error "Head is not a constructor" NONE [t, u], u)
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   452
        | dest_case1 t = raise TERM ("dest_case1", [t]);
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   453
      fun dest_case2 (Const ("_case2", _) $ t $ u) = t :: dest_case2 u
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   454
        | dest_case2 t = [t];
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   455
      val cases as ((cname, _), _) :: _ = map dest_case1 (dest_case2 u);
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   456
      val tab = Symtab.dest (get_datatypes sg);
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   457
      val (cases', default) = (case split_last cases of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   458
          (cases', (("dummy_pattern", []), t)) => (cases', SOME t)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   459
        | _ => (cases, NONE))
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   460
      fun abstr (Free (x, T), body) = Term.absfree (x, T, body)
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   461
        | abstr (Const ("_constrain", _) $ Free (x, T) $ tT, body) =
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   462
            Syntax.const Syntax.constrainAbsC $ Term.absfree (x, T, body) $ tT
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   463
        | abstr (Const ("Pair", _) $ x $ y, body) =
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   464
            Syntax.const "split" $ abstr (x, abstr (y, body))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   465
        | abstr (t, _) = case_error "Illegal pattern" NONE [t];
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   466
    in case find_first (fn (_, {descr, index, ...}) =>
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   467
      exists (equal cname o fst) (#3 (snd (List.nth (descr, index))))) tab of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   468
        NONE => case_error ("Not a datatype constructor: " ^ cname) NONE [u]
18319
c52b139ebde0 Added new component "sorts" to record datatype_info.
berghofe
parents: 18314
diff changeset
   469
      | SOME (tname, {descr, sorts, case_name, index, ...}) =>
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   470
        let
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   471
          val _ = if exists (equal "dummy_pattern" o fst o fst) cases' then
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   472
            case_error "Illegal occurrence of '_' dummy pattern" (SOME tname) [u] else ();
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   473
          val (_, (_, dts, constrs)) = List.nth (descr, index);
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   474
          fun find_case (cases, (s, dt)) =
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   475
            (case find_first (equal s o fst o fst) cases' of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   476
               NONE => (case default of
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   477
                   NONE => case_error ("No clause for constructor " ^ s) (SOME tname) [u]
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   478
                 | SOME t => (cases, list_abs (map (rpair dummyT) (DatatypeProp.make_tnames
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   479
                     (map (typ_of_dtyp descr sorts) dt)), t)))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   480
             | SOME (c as ((_, vs), t)) =>
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   481
                 if length dt <> length vs then
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   482
                    case_error ("Wrong number of arguments for constructor " ^ s)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   483
                      (SOME tname) vs
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   484
                 else (cases \ c, foldr abstr t vs))
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   485
          val (cases'', fs) = foldl_map find_case (cases', constrs)
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   486
        in case (cases'', length constrs = length cases', default) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   487
            ([], true, SOME _) =>
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   488
              case_error "Extra '_' dummy pattern" (SOME tname) [u]
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   489
          | (_ :: _, _, _) =>
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   490
              let val extra = distinct (map (fst o fst) cases'')
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   491
              in case extra \\ map fst constrs of
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   492
                  [] => case_error ("More than one clause for constructor(s) " ^
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   493
                    commas extra) (SOME tname) [u]
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   494
                | extra' => case_error ("Illegal constructor(s): " ^ commas extra')
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   495
                    (SOME tname) [u]
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   496
              end
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   497
          | _ => list_comb (Syntax.const case_name, fs) $ t
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   498
        end
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   499
    end
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   500
  | case_tr sg ts = raise TERM ("case_tr", ts);
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   501
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   502
fun case_tr' constrs sg ts =
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   503
  if length ts <> length constrs + 1 then raise Match else
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   504
  let
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   505
    val (fs, x) = split_last ts;
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   506
    fun strip_abs 0 t = ([], t)
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   507
      | strip_abs i (Abs p) =
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   508
        let val (x, u) = Syntax.atomic_abs_tr' p
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   509
        in apfst (cons x) (strip_abs (i-1) u) end
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   510
      | strip_abs i (Const ("split", _) $ t) = (case strip_abs (i+1) t of
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   511
          (v :: v' :: vs, u) => (Syntax.const "Pair" $ v $ v' :: vs, u));
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   512
    fun is_dependent i t =
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   513
      let val k = length (strip_abs_vars t) - i
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   514
      in k < 0 orelse exists (fn j => j >= k)
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   515
        (loose_bnos (strip_abs_body t))
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   516
      end;
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   517
    val cases = map (fn ((cname, dts), t) =>
16364
dc9f7066d80a refer to name spaces values instead of names;
wenzelm
parents: 16122
diff changeset
   518
      (Sign.extern_const sg cname,
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   519
       strip_abs (length dts) t, is_dependent (length dts) t))
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   520
      (constrs ~~ fs);
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   521
    fun count_cases (cs, (_, _, true)) = cs
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   522
      | count_cases (cs, (cname, (_, body), false)) = (case AList.lookup (op =) cs body of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   523
          NONE => (body, [cname]) :: cs
17325
d9d50222808e introduced new-style AList operations
haftmann
parents: 17261
diff changeset
   524
        | SOME cnames => AList.update (op =) (body, cnames @ [cname]) cs);
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   525
    val cases' = sort (int_ord o Library.swap o pairself (length o snd))
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   526
      (Library.foldl count_cases ([], cases));
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   527
    fun mk_case1 (cname, (vs, body), _) = Syntax.const "_case1" $
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   528
      list_comb (Syntax.const cname, vs) $ body;
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   529
  in
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   530
    Syntax.const "_case_syntax" $ x $
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   531
      foldr1 (fn (t, u) => Syntax.const "_case2" $ t $ u) (map mk_case1
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   532
        (case cases' of
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   533
           [] => cases
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   534
         | (default, cnames) :: _ =>
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   535
           if length cnames = 1 then cases
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   536
           else if length cnames = length constrs then
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   537
             [hd cases, ("dummy_pattern", ([], default), false)]
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   538
           else
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   539
             filter_out (fn (cname, _, _) => cname mem cnames) cases @
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   540
             [("dummy_pattern", ([], default), false)]))
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   541
  end;
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   542
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   543
fun make_case_tr' case_names descr = List.concat (map
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   544
  (fn ((_, (_, _, constrs)), case_name) => map (rpair (case_tr' constrs))
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   545
    (NameSpace.accesses' case_name)) (descr ~~ case_names));
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   546
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   547
val trfun_setup =
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   548
  [Theory.add_advanced_trfuns ([], [("_case_syntax", case_tr)], [], [])];
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   549
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   550
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   551
(* prepare types *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   552
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   553
fun read_typ sign ((Ts, sorts), str) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   554
  let
17521
0f1c48de39f5 introduced AList module in favor of assoc etc.
haftmann
parents: 17412
diff changeset
   555
    val T = Type.no_tvars (Sign.read_typ (sign, AList.lookup (op =)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   556
      (map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   557
  in (Ts @ [T], add_typ_tfrees (T, sorts)) end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   558
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   559
fun cert_typ sign ((Ts, sorts), raw_T) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   560
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   561
    val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   562
      TYPE (msg, _, _) => error msg;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   563
    val sorts' = add_typ_tfrees (T, sorts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   564
  in (Ts @ [T],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   565
      case duplicates (map fst sorts') of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   566
         [] => sorts'
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   567
       | dups => error ("Inconsistent sort constraints for " ^ commas dups))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   568
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   569
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   570
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   571
(**** make datatype info ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   572
18319
c52b139ebde0 Added new component "sorts" to record datatype_info.
berghofe
parents: 18314
diff changeset
   573
fun make_dt_info descr sorts induct reccomb_names rec_thms
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   574
    (((((((((i, (_, (tname, _, _))), case_name), case_thms),
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   575
      exhaustion_thm), distinct_thm), inject), nchotomy), case_cong), weak_case_cong) =
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   576
  (tname,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   577
   {index = i,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   578
    descr = descr,
18319
c52b139ebde0 Added new component "sorts" to record datatype_info.
berghofe
parents: 18314
diff changeset
   579
    sorts = sorts,
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   580
    rec_names = reccomb_names,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   581
    rec_rewrites = rec_thms,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   582
    case_name = case_name,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   583
    case_rewrites = case_thms,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   584
    induction = induct,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   585
    exhaustion = exhaustion_thm,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   586
    distinct = distinct_thm,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   587
    inject = inject,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   588
    nchotomy = nchotomy,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   589
    case_cong = case_cong,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   590
    weak_case_cong = weak_case_cong});
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   591
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   592
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   593
(********************* axiomatic introduction of datatypes ********************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   594
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   595
fun add_and_get_axioms_atts label tnames attss ts thy =
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   596
  foldr (fn (((tname, atts), t), (thy', axs)) =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   597
    let
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   598
      val ([ax], thy'') =
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   599
        thy'
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   600
        |> Theory.add_path tname
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   601
        |> PureThy.add_axioms_i [((label, t), atts)];
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   602
    in (Theory.parent_path thy'', ax::axs)
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   603
    end) (thy, []) (tnames ~~ attss ~~ ts) |> swap;
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   604
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   605
fun add_and_get_axioms label tnames =
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   606
  add_and_get_axioms_atts label tnames (replicate (length tnames) []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   607
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   608
fun add_and_get_axiomss label tnames tss thy =
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   609
  foldr (fn ((tname, ts), (thy', axss)) =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   610
    let
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   611
      val ([axs], thy'') =
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   612
        thy'
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   613
        |> Theory.add_path tname
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   614
        |> PureThy.add_axiomss_i [((label, ts), [])];
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   615
    in (Theory.parent_path thy'', axs::axss)
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   616
    end) (thy, []) (tnames ~~ tss) |> swap;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   617
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   618
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
   619
    case_names_induct case_names_exhausts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   620
  let
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   621
    val descr' = List.concat descr;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   622
    val recTs = get_rec_types descr' sorts;
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   623
    val used = foldr add_typ_tfree_names [] recTs;
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   624
    val newTs = Library.take (length (hd descr), recTs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   625
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   626
    val no_size = exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists
13641
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13480
diff changeset
   627
      (fn dt => is_rec_type dt andalso not (null (fst (strip_dtyp dt))))
63d1790a43ed Reimplemented parts of datatype package dealing with datatypes involving
berghofe
parents: 13480
diff changeset
   628
        cargs) constrs) descr';
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   629
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   630
    (**** declare new types and constants ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   631
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   632
    val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   633
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   634
    val constr_decls = map (fn (((_, (_, _, constrs)), T), constr_syntax') =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   635
      map (fn ((_, cargs), (cname, mx)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   636
        (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   637
          (constrs ~~ constr_syntax')) ((hd descr) ~~ newTs ~~ constr_syntax);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   638
15457
1fbd4aba46e3 Adapted to modified interface of PureThy.get_thm(s).
berghofe
parents: 15444
diff changeset
   639
    val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   640
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   641
    val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   642
    val reccomb_names = if length descr' = 1 then [big_reccomb_name] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   643
      (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   644
        (1 upto (length descr')));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   645
9739
8470c4662685 Improved names for size function.
berghofe
parents: 9714
diff changeset
   646
    val size_names = DatatypeProp.indexify_names
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   647
      (map (fn T => name_of_typ T ^ "_size") (Library.drop (length (hd descr), recTs)));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   648
12338
de0f4a63baa5 renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents: 12311
diff changeset
   649
    val freeT = TFree (variant used "'t", HOLogic.typeS);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   650
    val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   651
      map (fn (_, cargs) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   652
        let val Ts = map (typ_of_dtyp descr' sorts) cargs
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   653
        in Ts ---> freeT end) constrs) (hd descr);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   654
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   655
    val case_names = map (fn s => (s ^ "_case")) new_type_names;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   656
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   657
    val thy2' = thy |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   658
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   659
      (** new types **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   660
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   661
      curry (Library.foldr (fn (((name, mx), tvs), thy') => thy' |>
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   662
          TypedefPackage.add_typedecls [(name, tvs, mx)]))
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   663
        (types_syntax ~~ tyvars) |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   664
      add_path flat_names (space_implode "_" new_type_names) |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   665
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   666
      (** primrec combinators **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   667
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   668
      Theory.add_consts_i (map (fn ((name, T), T') =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   669
        (name, reccomb_fn_Ts @ [T] ---> T', NoSyn))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   670
          (reccomb_names ~~ recTs ~~ rec_result_Ts)) |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   671
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   672
      (** case combinators **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   673
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   674
      Theory.add_consts_i (map (fn ((name, T), Ts) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   675
        (name, Ts @ [T] ---> freeT, NoSyn))
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   676
          (case_names ~~ newTs ~~ case_fn_Ts));
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   677
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   678
    val reccomb_names' = map (Sign.intern_const thy2') reccomb_names;
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   679
    val case_names' = map (Sign.intern_const thy2') case_names;
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   680
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   681
    val thy2 = thy2' |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   682
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   683
      (** size functions **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   684
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   685
      (if no_size then I else Theory.add_consts_i (map (fn (s, T) =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   686
        (Sign.base_name s, T --> HOLogic.natT, NoSyn))
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   687
          (size_names ~~ Library.drop (length (hd descr), recTs)))) |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   688
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   689
      (** constructors **)
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   690
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   691
      parent_path flat_names |>
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   692
      curry (Library.foldr (fn (((((_, (_, _, constrs)), T), tname),
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   693
        constr_syntax'), thy') => thy' |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   694
          add_path flat_names tname |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   695
            Theory.add_consts_i (map (fn ((_, cargs), (cname, mx)) =>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   696
              (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   697
                (constrs ~~ constr_syntax')) |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   698
          parent_path flat_names))
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   699
            (hd descr ~~ newTs ~~ new_type_names ~~ constr_syntax);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   700
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   701
    (**** introduction of axioms ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   702
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   703
    val rec_axs = DatatypeProp.make_primrecs new_type_names descr sorts thy2;
9739
8470c4662685 Improved names for size function.
berghofe
parents: 9714
diff changeset
   704
    val size_axs = if no_size then [] else DatatypeProp.make_size descr sorts thy2;
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   705
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   706
    val ((([induct], [rec_thms]), inject), thy3) =
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   707
      thy2
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   708
      |> Theory.add_path (space_implode "_" new_type_names)
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   709
      |> PureThy.add_axioms_i [(("induct", DatatypeProp.make_ind descr sorts),
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   710
          [case_names_induct])]
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   711
      ||>> PureThy.add_axiomss_i [(("recs", rec_axs), [])]
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   712
      ||> (if no_size then I else snd o PureThy.add_axiomss_i [(("size", size_axs), [])])
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   713
      ||> Theory.parent_path
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   714
      ||>> add_and_get_axiomss "inject" new_type_names
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   715
            (DatatypeProp.make_injs descr sorts);
16486
1a12cdb6ee6b get_thm(s): Name;
wenzelm
parents: 16430
diff changeset
   716
    val size_thms = if no_size then [] else get_thms thy3 (Name "size");
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   717
    val (distinct, thy4) = add_and_get_axiomss "distinct" new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   718
      (DatatypeProp.make_distincts new_type_names descr sorts thy3) thy3;
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   719
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   720
    val exhaust_ts = DatatypeProp.make_casedists descr sorts;
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   721
    val (exhaustion, thy5) = add_and_get_axioms_atts "exhaust" new_type_names
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   722
      (map Library.single case_names_exhausts) exhaust_ts thy4;
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   723
    val (case_thms, thy6) = add_and_get_axiomss "cases" new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   724
      (DatatypeProp.make_cases new_type_names descr sorts thy5) thy5;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   725
    val (split_ts, split_asm_ts) = ListPair.unzip
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   726
      (DatatypeProp.make_splits new_type_names descr sorts thy6);
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   727
    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
   728
    val (split_asm, thy8) = add_and_get_axioms "split_asm" new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   729
      split_asm_ts thy7;
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   730
    val (nchotomys, thy9) = add_and_get_axioms "nchotomy" new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   731
      (DatatypeProp.make_nchotomys descr sorts) thy8;
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   732
    val (case_congs, thy10) = add_and_get_axioms "case_cong" new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   733
      (DatatypeProp.make_case_congs new_type_names descr sorts thy9) thy9;
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   734
    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
   735
      (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
   736
18319
c52b139ebde0 Added new component "sorts" to record datatype_info.
berghofe
parents: 18314
diff changeset
   737
    val dt_infos = map (make_dt_info descr' sorts induct reccomb_names' rec_thms)
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   738
      ((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
   739
        exhaustion ~~ replicate (length (hd descr)) QuickAndDirty ~~ inject ~~
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   740
          nchotomys ~~ case_congs ~~ weak_case_congs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   741
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   742
    val simps = List.concat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
9386
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   743
    val split_thms = split ~~ split_asm;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   744
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   745
    val thy12 =
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   746
      thy11
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   747
      |> Theory.add_advanced_trfuns ([], [], make_case_tr' case_names' (hd descr), [])
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   748
      |> Theory.add_path (space_implode "_" new_type_names)
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   749
      |> add_rules simps case_thms size_thms rec_thms inject distinct
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18678
diff changeset
   750
           weak_case_congs (Attrib.theory Simplifier.cong_add)
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   751
      |> put_datatypes (fold Symtab.update dt_infos dt_info)
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   752
      |> add_cases_induct dt_infos induct
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   753
      |> Theory.parent_path
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   754
      |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms) |> snd
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   755
      |> DatatypeRealizer.add_dt_realizers sorts (map snd dt_infos)
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   756
      |> fold (CodegenPackage.add_cg_case_const_i get_case_const_data) case_names';
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   757
  in
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   758
    ({distinct = distinct,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   759
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   760
      exhaustion = exhaustion,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   761
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   762
      case_thms = case_thms,
9386
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   763
      split_thms = split_thms,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   764
      induction = induct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   765
      size = size_thms,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   766
      simps = simps}, thy12)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   767
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   768
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   769
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   770
(******************* definitional introduction of datatypes *******************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   771
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   772
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
   773
    case_names_induct case_names_exhausts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   774
  let
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   775
    val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   776
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   777
    val ((inject, distinct, dist_rewrites, simproc_dists, induct), thy2) = thy |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   778
      DatatypeRepProofs.representation_proofs flat_names dt_info new_type_names descr sorts
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   779
        types_syntax constr_syntax case_names_induct;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   780
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   781
    val (casedist_thms, thy3) = DatatypeAbsProofs.prove_casedist_thms new_type_names descr
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   782
      sorts induct case_names_exhausts thy2;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   783
    val ((reccomb_names, rec_thms), thy4) = DatatypeAbsProofs.prove_primrec_thms
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   784
      flat_names new_type_names descr sorts dt_info inject dist_rewrites dist_ss induct thy3;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   785
    val ((case_thms, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   786
      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
   787
    val (split_thms, thy7) = DatatypeAbsProofs.prove_split_thms new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   788
      descr sorts inject dist_rewrites casedist_thms case_thms thy6;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   789
    val (nchotomys, thy8) = DatatypeAbsProofs.prove_nchotomys new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   790
      descr sorts casedist_thms thy7;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   791
    val (case_congs, thy9) = DatatypeAbsProofs.prove_case_congs new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   792
      descr sorts nchotomys case_thms thy8;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   793
    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
   794
      descr sorts thy9;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   795
    val (size_thms, thy11) = DatatypeAbsProofs.prove_size_thms flat_names new_type_names
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   796
      descr sorts reccomb_names rec_thms thy10;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   797
18319
c52b139ebde0 Added new component "sorts" to record datatype_info.
berghofe
parents: 18314
diff changeset
   798
    val dt_infos = map (make_dt_info (List.concat descr) sorts induct reccomb_names rec_thms)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   799
      ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   800
        casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   801
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   802
    val simps = List.concat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   803
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   804
    val thy12 =
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   805
      thy11
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   806
      |> Theory.add_advanced_trfuns ([], [], make_case_tr' case_names (hd descr), [])
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   807
      |> Theory.add_path (space_implode "_" new_type_names)
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   808
      |> add_rules simps case_thms size_thms rec_thms inject distinct
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18678
diff changeset
   809
            weak_case_congs (Attrib.theory (Simplifier.attrib (op addcongs)))
18518
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   810
      |> put_datatypes (fold Symtab.update dt_infos dt_info)
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   811
      |> add_cases_induct dt_infos induct
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   812
      |> Theory.parent_path
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   813
      |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms) |> snd
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   814
      |> DatatypeRealizer.add_dt_realizers sorts (map snd dt_infos)
3b1dfa53e64f adaptions to changes in code generator
haftmann
parents: 18462
diff changeset
   815
      |> fold (CodegenPackage.add_cg_case_const_i get_case_const_data) case_names;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   816
  in
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   817
    ({distinct = distinct,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   818
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   819
      exhaustion = casedist_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   820
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   821
      case_thms = case_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   822
      split_thms = split_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   823
      induction = induct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   824
      size = size_thms,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   825
      simps = simps}, thy12)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   826
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   827
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   828
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   829
(*********************** declare existing type as datatype *********************)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   830
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   831
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
   832
  let
12922
ed70a600f0ea clarified internal theory dependencies;
wenzelm
parents: 12876
diff changeset
   833
    val _ = Theory.requires thy0 "Inductive" "datatype representations";
ed70a600f0ea clarified internal theory dependencies;
wenzelm
parents: 12876
diff changeset
   834
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   835
    val (((distinct, inject), [induction]), thy1) =
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   836
      thy0
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   837
      |> fold_map apply_theorems raw_distinct
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   838
      ||>> fold_map apply_theorems raw_inject
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   839
      ||>> apply_theorems [raw_induction];
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6385
diff changeset
   840
    val sign = Theory.sign_of thy1;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   841
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   842
    val induction' = freezeT induction;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   843
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   844
    fun err t = error ("Ill-formed predicate in induction rule: " ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   845
      Sign.string_of_term sign t);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   846
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   847
    fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   848
          ((tname, map dest_TFree Ts) handle TERM _ => err t)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   849
      | get_typ t = err t;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   850
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   851
    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
   852
    val new_type_names = getOpt (alt_names, map fst dtnames);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   853
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   854
    fun get_constr t = (case Logic.strip_assums_concl t of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   855
        _ $ (_ $ t') => (case head_of t' of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   856
            Const (cname, cT) => (case strip_type cT of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   857
                (Ts, Type (tname, _)) => (tname, (cname, map (dtyp_of_typ dtnames) Ts))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   858
              | _ => err t)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   859
          | _ => err t)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   860
      | _ => err t);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   861
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   862
    fun make_dt_spec [] _ _ = []
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   863
      | make_dt_spec ((tname, tvs)::dtnames') i constrs =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   864
          let val (constrs', constrs'') = take_prefix (equal tname o fst) constrs
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   865
          in (i, (tname, map DtTFree tvs, map snd constrs'))::
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   866
            (make_dt_spec dtnames' (i + 1) constrs'')
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   867
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   868
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   869
    val descr = make_dt_spec dtnames 0 (map get_constr (prems_of induction'));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   870
    val sorts = add_term_tfrees (concl_of induction', []);
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   871
    val dt_info = get_datatypes thy1;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   872
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   873
    val case_names_induct = mk_case_names_induct descr;
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   874
    val case_names_exhausts = mk_case_names_exhausts descr (map #1 dtnames);
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   875
    
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   876
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6423
diff changeset
   877
    val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   878
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   879
    val (casedist_thms, thy2) = thy1 |>
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   880
      DatatypeAbsProofs.prove_casedist_thms new_type_names [descr] sorts induction
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   881
        case_names_exhausts;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   882
    val ((reccomb_names, rec_thms), thy3) = DatatypeAbsProofs.prove_primrec_thms
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   883
      false new_type_names [descr] sorts dt_info inject distinct dist_ss induction thy2;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   884
    val ((case_thms, case_names), thy4) = DatatypeAbsProofs.prove_case_thms false
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   885
      new_type_names [descr] sorts reccomb_names rec_thms thy3;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   886
    val (split_thms, thy5) = DatatypeAbsProofs.prove_split_thms
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   887
      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
   888
    val (nchotomys, thy6) = DatatypeAbsProofs.prove_nchotomys new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   889
      [descr] sorts casedist_thms thy5;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   890
    val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   891
      [descr] sorts nchotomys case_thms thy6;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   892
    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
   893
      [descr] sorts thy7;
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   894
    val (size_thms, thy9) =
16430
bc07926e4f03 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   895
      if Context.exists_name "NatArith" thy8 then
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   896
        DatatypeAbsProofs.prove_size_thms false new_type_names
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   897
          [descr] sorts reccomb_names rec_thms thy8
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   898
      else ([], thy8);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   899
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   900
    val ((_, [induction']), thy10) =
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   901
      thy9
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   902
      |> store_thmss "inject" new_type_names inject
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   903
      ||>> store_thmss "distinct" new_type_names distinct
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   904
      ||> Theory.add_path (space_implode "_" new_type_names)
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18319
diff changeset
   905
      ||>> PureThy.add_thms [(("induct", induction), [case_names_induct])];
9149
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   906
18319
c52b139ebde0 Added new component "sorts" to record datatype_info.
berghofe
parents: 18314
diff changeset
   907
    val dt_infos = map (make_dt_info descr sorts induction' reccomb_names rec_thms)
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   908
      ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   909
        map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   910
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   911
    val simps = List.concat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   912
9149
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   913
    val thy11 = thy10 |>
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
   914
      Theory.add_advanced_trfuns ([], [], make_case_tr' case_names descr, []) |>
11345
cd605c85e421 improved iff_add_global, new function add_rules factoring out common behaviour
oheimb
parents: 10930
diff changeset
   915
      add_rules simps case_thms size_thms rec_thms inject distinct
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18678
diff changeset
   916
                weak_case_congs (Attrib.theory (Simplifier.attrib (op addcongs))) |> 
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17325
diff changeset
   917
      put_datatypes (fold Symtab.update dt_infos dt_info) |>
11805
b110a1ea90da declare projected induction rules stemming from nested recursion;
wenzelm
parents: 11725
diff changeset
   918
      add_cases_induct dt_infos induction' |>
9386
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   919
      Theory.parent_path |>
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   920
      (snd o store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)) |>
13466
42766aa25460 Added calls to add_dt_realizers.
berghofe
parents: 13462
diff changeset
   921
      DatatypeRealizer.add_dt_realizers sorts (map snd dt_infos);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   922
  in
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   923
    ({distinct = distinct,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   924
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   925
      exhaustion = casedist_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   926
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   927
      case_thms = case_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   928
      split_thms = split_thms,
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   929
      induction = induction',
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   930
      size = size_thms,
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
   931
      simps = simps}, thy11)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   932
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   933
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   934
val rep_datatype = gen_rep_datatype IsarThy.apply_theorems;
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   935
val rep_datatype_i = gen_rep_datatype IsarThy.apply_theorems_i;
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   936
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   937
11958
2ece34b9fd8e Isar: fixed rep_datatype args;
wenzelm
parents: 11805
diff changeset
   938
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   939
(******************************** add datatype ********************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   940
14887
4938ce4ef295 Added exception Datatype_Empty.
berghofe
parents: 14799
diff changeset
   941
fun gen_add_datatype prep_typ err flat_names new_type_names dts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   942
  let
12922
ed70a600f0ea clarified internal theory dependencies;
wenzelm
parents: 12876
diff changeset
   943
    val _ = Theory.requires thy "Datatype_Universe" "datatype definitions";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   944
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   945
    (* this theory is used just for parsing *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   946
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   947
    val tmp_thy = thy |>
5892
e5e44cc54eb2 Attribute.tthms_of;
wenzelm
parents: 5720
diff changeset
   948
      Theory.copy |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   949
      Theory.add_types (map (fn (tvs, tname, mx, _) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   950
        (tname, length tvs, mx)) dts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   951
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6385
diff changeset
   952
    val sign = Theory.sign_of tmp_thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   953
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   954
    val (tyvars, _, _, _)::_ = dts;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   955
    val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   956
      let val full_tname = Sign.full_name sign (Syntax.type_name tname mx)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   957
      in (case duplicates tvs of
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   958
            [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   959
                  else error ("Mutually recursive datatypes must have same type parameters")
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   960
          | dups => error ("Duplicate parameter(s) for datatype " ^ full_tname ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   961
              " : " ^ commas dups))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   962
      end) dts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   963
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   964
    val _ = (case duplicates (map fst new_dts) @ duplicates new_type_names of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   965
      [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   966
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   967
    fun prep_dt_spec ((dts', constr_syntax, sorts, i), (tvs, tname, mx, constrs)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   968
      let
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   969
        fun prep_constr ((constrs, constr_syntax', sorts'), (cname, cargs, mx')) =
5279
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   970
          let
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   971
            val (cargs', sorts'') = Library.foldl (prep_typ sign) (([], sorts'), cargs);
15574
b1d1b5bfc464 Removed practically all references to Library.foldr.
skalberg
parents: 15570
diff changeset
   972
            val _ = (case foldr add_typ_tfree_names [] cargs' \\ tvs of
5279
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   973
                [] => ()
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   974
              | vs => error ("Extra type variables on rhs: " ^ commas vs))
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   975
          in (constrs @ [((if flat_names then Sign.full_name sign else
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   976
                Sign.full_name_path sign tname) (Syntax.const_name cname mx'),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   977
                   map (dtyp_of_typ new_dts) cargs')],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   978
              constr_syntax' @ [(cname, mx')], sorts'')
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18643
diff changeset
   979
          end handle ERROR msg =>
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18643
diff changeset
   980
            cat_error msg ("The error above occured in constructor " ^ cname ^
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   981
              " of datatype " ^ tname);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   982
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   983
        val (constrs', constr_syntax', sorts') =
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   984
          Library.foldl prep_constr (([], [], sorts), constrs)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   985
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   986
      in
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   987
        case duplicates (map fst constrs') of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   988
           [] =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   989
             (dts' @ [(i, (Sign.full_name sign (Syntax.type_name tname mx),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   990
                map DtTFree tvs, constrs'))],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   991
              constr_syntax @ [constr_syntax'], sorts', i + 1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   992
         | dups => error ("Duplicate constructors " ^ commas dups ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   993
             " in datatype " ^ tname)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   994
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   995
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   996
    val (dts', constr_syntax, sorts', i) = Library.foldl prep_dt_spec (([], [], [], 0), dts);
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   997
    val sorts = sorts' @ (map (rpair (Sign.defaultS sign)) (tyvars \\ map fst sorts'));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   998
    val dt_info = get_datatypes thy;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   999
    val (descr, _) = unfold_datatypes sign dts' sorts dt_info dts' i;
14887
4938ce4ef295 Added exception Datatype_Empty.
berghofe
parents: 14799
diff changeset
  1000
    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
  1001
      if err then error ("Nonemptiness check failed for datatype " ^ s)
14887
4938ce4ef295 Added exception Datatype_Empty.
berghofe
parents: 14799
diff changeset
  1002
      else raise exn;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
  1003
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
  1004
    val descr' = List.concat descr;
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
  1005
    val case_names_induct = mk_case_names_induct descr';
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
  1006
    val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
  1007
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
  1008
    (if (!quick_and_dirty) then add_datatype_axm else add_datatype_def)
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
  1009
      flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
  1010
      case_names_induct case_names_exhausts thy
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
  1011
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
  1012
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
  1013
val add_datatype_i = gen_add_datatype cert_typ;
14887
4938ce4ef295 Added exception Datatype_Empty.
berghofe
parents: 14799
diff changeset
  1014
val add_datatype = gen_add_datatype read_typ true;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
  1015
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1016
18451
5ff0244e25e8 slight clean ups
haftmann
parents: 18418
diff changeset
  1017
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1018
(** package setup **)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1019
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1020
(* setup theory *)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1021
14799
a405aadff16c Added more flexible parse / print translations for case expressions.
berghofe
parents: 14471
diff changeset
  1022
val setup = [DatatypesData.init, Method.add_methods tactic_emulations] @ simproc_setup @ trfun_setup;
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1023
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1024
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1025
(* outer syntax *)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1026
17057
0934ac31985f OuterKeyword;
wenzelm
parents: 16973
diff changeset
  1027
local structure P = OuterParse and K = OuterKeyword in
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1028
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1029
val datatype_decl =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
  1030
  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
  1031
    (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1032
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1033
fun mk_datatype args =
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1034
  let
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
  1035
    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
  1036
    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
  1037
      (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
  1038
  in snd o add_datatype false names specs end;
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1039
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1040
val datatypeP =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
  1041
  OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
  1042
    (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1043
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
  1044
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
  1045
val rep_datatype_decl =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
  1046
  Scan.option (Scan.repeat1 P.name) --
11958
2ece34b9fd8e Isar: fixed rep_datatype args;
wenzelm
parents: 11805
diff changeset
  1047
    Scan.optional (P.$$$ "distinct" |-- P.!!! (P.and_list1 P.xthms1)) [[]] --
2ece34b9fd8e Isar: fixed rep_datatype args;
wenzelm
parents: 11805
diff changeset
  1048
    Scan.optional (P.$$$ "inject" |-- P.!!! (P.and_list1 P.xthms1)) [[]] --
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
  1049
    (P.$$$ "induction" |-- P.!!! P.xthm);
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
  1050
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17956
diff changeset
  1051
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
  1052
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
  1053
val rep_datatypeP =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
  1054
  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
  1055
    (rep_datatype_decl >> (Toplevel.theory o mk_rep_datatype));
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
  1056
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
  1057
6479
b0448cab1b1e rep_datatype syntax: 'induction' instead of 'induct';
wenzelm
parents: 6441
diff changeset
  1058
val _ = OuterSyntax.add_keywords ["distinct", "inject", "induction"];
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
  1059
val _ = OuterSyntax.add_parsers [datatypeP, rep_datatypeP];
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
  1060
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
  1061
end;
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
  1062
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1063
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
  1064
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
  1065
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1066
structure BasicDatatypePackage: BASIC_DATATYPE_PACKAGE = DatatypePackage;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
  1067
open BasicDatatypePackage;
15704
93163972dbdc *** MESSAGE REFERS TO PREVIOUS VERSION ***
wenzelm
parents: 15703
diff changeset
  1068