src/HOL/Tools/datatype_package.ML
author wenzelm
Thu, 19 Oct 2000 21:23:47 +0200
changeset 10279 b223a9a3350e
parent 10214 77349ed89f45
child 10525 3e21ab3e5114
permissions -rw-r--r--
InductAttrib;
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$
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     3
    Author:     Stefan Berghofer
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     4
    Copyright   1998  TU Muenchen
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     5
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
     6
Datatype package for Isabelle/HOL.
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     7
*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
     8
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
     9
signature BASIC_DATATYPE_PACKAGE =
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    10
sig
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    11
  val induct_tac : string -> int -> tactic
9747
043098ba5098 introduced induct_thm_tac
nipkow
parents: 9739
diff changeset
    12
  val induct_thm_tac : thm -> string -> int -> tactic
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
    13
  val case_tac : string -> int -> tactic
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
    14
  val distinct_simproc : simproc
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    15
end;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    16
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    17
signature DATATYPE_PACKAGE =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    18
sig
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    19
  include BASIC_DATATYPE_PACKAGE
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    20
  val quiet_mode : bool ref
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    21
  val add_datatype : bool -> string list -> (string list * bstring * mixfix *
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    22
    (bstring * string list * mixfix) list) list -> theory -> theory *
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    23
      {distinct : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    24
       inject : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    25
       exhaustion : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    26
       rec_thms : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    27
       case_thms : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    28
       split_thms : (thm * thm) list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    29
       induction : thm,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    30
       size : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    31
       simps : thm list}
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    32
  val add_datatype_i : bool -> string list -> (string list * bstring * mixfix *
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    33
    (bstring * typ list * mixfix) list) list -> theory -> theory *
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    34
      {distinct : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    35
       inject : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    36
       exhaustion : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    37
       rec_thms : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    38
       case_thms : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    39
       split_thms : (thm * thm) list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    40
       induction : thm,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    41
       size : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    42
       simps : thm list}
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    43
  val rep_datatype_i : string list option -> (thm * theory attribute list) list list ->
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    44
    (thm * theory attribute list) list list -> (thm * theory attribute list) -> theory -> theory *
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,
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    53
       simps : thm list}
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    54
  val rep_datatype : string list option -> (xstring * Args.src list) list list ->
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
    55
    (xstring * Args.src list) list list -> xstring * Args.src list -> theory -> 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,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    64
       simps : thm list}
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    65
  val get_datatypes : theory -> DatatypeAux.datatype_info Symtab.table
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
    66
  val get_datatypes_sg : Sign.sg -> DatatypeAux.datatype_info Symtab.table
6441
268aa3c4a059 print_datatypes;
wenzelm
parents: 6427
diff changeset
    67
  val print_datatypes : theory -> unit
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
    68
  val datatype_info_sg : Sign.sg -> string -> DatatypeAux.datatype_info option
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
    69
  val datatype_info : theory -> string -> DatatypeAux.datatype_info option
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
    70
  val datatype_info_sg_err : Sign.sg -> string -> DatatypeAux.datatype_info
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
    71
  val datatype_info_err : theory -> string -> DatatypeAux.datatype_info
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    72
  val constrs_of : theory -> string -> term list option
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
    73
  val constrs_of_sg : Sign.sg -> string -> term list option
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    74
  val case_const_of : theory -> string -> term option
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
    75
  val weak_case_congs_of : theory -> thm list
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
    76
  val weak_case_congs_of_sg : Sign.sg -> thm list
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    77
  val setup: (theory -> theory) list
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    78
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    79
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    80
structure DatatypePackage : DATATYPE_PACKAGE =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    81
struct
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    82
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    83
open DatatypeAux;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    84
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    85
val quiet_mode = quiet_mode;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    86
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    87
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    88
(* data kind 'HOL/datatypes' *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    89
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    90
structure DatatypesArgs =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    91
struct
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    92
  val name = "HOL/datatypes";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    93
  type T = datatype_info Symtab.table;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    94
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    95
  val empty = Symtab.empty;
6556
daa00919502b theory data: copy;
wenzelm
parents: 6479
diff changeset
    96
  val copy = I;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    97
  val prep_ext = I;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    98
  val merge: T * T -> T = Symtab.merge (K true);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    99
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   100
  fun print sg tab =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   101
    Pretty.writeln (Pretty.strs ("datatypes:" ::
6851
526c0b90bcef cond_extern_table;
wenzelm
parents: 6766
diff changeset
   102
      map #1 (Sign.cond_extern_table sg Sign.typeK tab)));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   103
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   104
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   105
structure DatatypesData = TheoryDataFun(DatatypesArgs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   106
val get_datatypes_sg = DatatypesData.get_sg;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   107
val get_datatypes = DatatypesData.get;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   108
val put_datatypes = DatatypesData.put;
6441
268aa3c4a059 print_datatypes;
wenzelm
parents: 6427
diff changeset
   109
val print_datatypes = DatatypesData.print;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   110
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   111
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   112
(** theory information about datatypes **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   113
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   114
fun datatype_info_sg sg name = Symtab.lookup (get_datatypes_sg sg, name);
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   115
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   116
fun datatype_info_sg_err sg name = (case datatype_info_sg sg name of
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   117
      Some info => info
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   118
    | None => error ("Unknown datatype " ^ quote name));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   119
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6385
diff changeset
   120
val datatype_info = datatype_info_sg o Theory.sign_of;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   121
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   122
fun datatype_info_err thy name = (case datatype_info thy name of
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   123
      Some info => info
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   124
    | None => error ("Unknown datatype " ^ quote name));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   125
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   126
fun constrs_of_sg sg tname = (case datatype_info_sg sg tname of
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   127
   Some {index, descr, ...} =>
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   128
     let val (_, _, constrs) = the (assoc (descr, index))
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   129
     in Some (map (fn (cname, _) => Const (cname, the (Sign.const_type sg cname))) constrs)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   130
     end
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   131
 | _ => None);
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   132
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   133
val constrs_of = constrs_of_sg o Theory.sign_of;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   134
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   135
fun case_const_of thy tname = (case datatype_info thy tname of
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   136
   Some {case_name, ...} => Some (Const (case_name, the (Sign.const_type
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   137
     (Theory.sign_of thy) case_name)))
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   138
 | _ => None);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   139
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   140
val weak_case_congs_of_sg = map (#weak_case_cong o #2) o Symtab.dest o get_datatypes_sg;
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   141
val weak_case_congs_of = weak_case_congs_of_sg o Theory.sign_of;
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   142
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   143
fun find_tname var Bi =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   144
  let val frees = map dest_Free (term_frees Bi)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   145
      val params = Logic.strip_params Bi;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   146
  in case assoc (frees @ params, var) of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   147
       None => error ("No such variable in subgoal: " ^ quote var)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   148
     | Some(Type (tn, _)) => tn
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   149
     | _ => error ("Cannot determine type of " ^ quote var)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   150
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   151
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   152
fun infer_tname state i aterm =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   153
  let
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   154
    val sign = Thm.sign_of_thm state;
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   155
    val (_, _, Bi, _) = Thm.dest_state (state, i)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   156
    val params = Logic.strip_params Bi;   (*params of subgoal i*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   157
    val params = rev (rename_wrt_term Bi params);   (*as they are printed*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   158
    val (types, sorts) = types_sorts state;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   159
    fun types' (a, ~1) = (case assoc (params, a) of None => types(a, ~1) | sm => sm)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   160
      | types' ixn = types ixn;
8100
6186ee807f2e replaced HOLogic.termTVar by HOLogic.termT;
wenzelm
parents: 7904
diff changeset
   161
    val (ct, _) = read_def_cterm (sign, types', sorts) [] false (aterm, TypeInfer.logicT);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   162
  in case #T (rep_cterm ct) of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   163
       Type (tn, _) => tn
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   164
     | _ => error ("Cannot determine type of " ^ quote aterm)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   165
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   166
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   167
(*Warn if the (induction) variable occurs Free among the premises, which
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   168
  usually signals a mistake.  But calls the tactic either way!*)
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   169
fun occs_in_prems tacf vars =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   170
  SUBGOAL (fn (Bi, i) =>
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   171
           (if  exists (fn Free (a, _) => a mem vars)
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   172
                      (foldr add_term_frees (#2 (strip_context Bi), []))
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   173
             then warning "Induction variable occurs also among premises!"
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   174
             else ();
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   175
            tacf i));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   176
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   177
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   178
(* generic induction tactic for datatypes *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   179
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   180
local
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   181
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   182
fun prep_var (Var (ixn, _), Some x) = Some (implode (tl (explode (Syntax.string_of_vname ixn))), x)
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   183
  | prep_var _ = None;
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   184
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   185
fun prep_inst (concl, xs) =	(*exception LIST*)
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   186
  let val vs = InductMethod.vars_of concl
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   187
  in mapfilter prep_var (Library.drop (length vs - length xs, vs) ~~ xs) end;
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   188
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   189
in
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   190
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   191
fun gen_induct_tac (varss, opt_rule) i state =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   192
  let
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   193
    val (_, _, Bi, _) = Thm.dest_state (state, i);
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   194
    val {sign, ...} = Thm.rep_thm state;
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   195
    val (rule, rule_name) =
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   196
      (case opt_rule of
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   197
        Some r => (r, "Induction rule")
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   198
      | None =>
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   199
          let val tn = find_tname (hd (mapfilter I (flat varss))) Bi
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   200
          in (#induction (datatype_info_sg_err sign tn), "Induction rule for type " ^ tn) end);
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   201
8697
88dafea5955c InductMethod.concls_of;
wenzelm
parents: 8689
diff changeset
   202
    val concls = InductMethod.concls_of rule;
88dafea5955c InductMethod.concls_of;
wenzelm
parents: 8689
diff changeset
   203
    val insts = flat (map prep_inst (concls ~~ varss)) handle LIST _ =>
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   204
      error (rule_name ^ " has different numbers of variables");
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   205
  in occs_in_prems (Tactic.res_inst_tac insts rule) (map #2 insts) i state end;
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   206
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   207
fun induct_tac s = gen_induct_tac (map (Library.single o Some) (Syntax.read_idents s), None);
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   208
9747
043098ba5098 introduced induct_thm_tac
nipkow
parents: 9739
diff changeset
   209
fun induct_thm_tac th s =
043098ba5098 introduced induct_thm_tac
nipkow
parents: 9739
diff changeset
   210
  gen_induct_tac ([map Some (Syntax.read_idents s)], Some th);
043098ba5098 introduced induct_thm_tac
nipkow
parents: 9739
diff changeset
   211
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   212
end;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   213
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   214
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   215
(* generic case tactic for datatypes *)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   216
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   217
fun case_inst_tac t rule i state =
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   218
  let
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   219
    val _ $ Var (ixn, _) $ _ = HOLogic.dest_Trueprop
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   220
      (hd (Logic.strip_assums_hyp (hd (Thm.prems_of rule))));
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   221
    val exh_vname = implode (tl (explode (Syntax.string_of_vname ixn)));
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   222
  in Tactic.res_inst_tac [(exh_vname, t)] rule i state end;
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   223
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   224
fun gen_case_tac (t, Some rule) i state = case_inst_tac t rule i state
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   225
  | gen_case_tac (t, None) i state =
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   226
      let val tn = infer_tname state i t in
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   227
        if tn = HOLogic.boolN then Tactic.res_inst_tac [("P", t)] case_split_thm i state
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   228
        else case_inst_tac t (#exhaustion (datatype_info_sg_err (Thm.sign_of_thm state) tn)) i state
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   229
      end;
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   230
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   231
fun case_tac t = gen_case_tac (t, None);
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   232
8279
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   233
3453f73fad71 added cases_tac;
wenzelm
parents: 8100
diff changeset
   234
8541
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   235
(** Isar tactic emulations **)
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   236
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   237
local
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   238
8689
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   239
val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.$$$ ":");
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   240
val opt_rule = Scan.option (rule_spec |-- Attrib.local_thm);
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   241
a2e82eed6454 improved induct_tac;
wenzelm
parents: 8686
diff changeset
   242
val varss = Args.and_list (Scan.repeat (Scan.unless rule_spec (Scan.lift Args.name_dummy)));
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   243
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   244
in
8541
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   245
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   246
val tactic_emulations =
9704
c2f2f70bbb60 'induct_tac' / 'case_tac': Method.goal_args';
wenzelm
parents: 9386
diff changeset
   247
 [("induct_tac", Method.goal_args' (varss -- opt_rule) gen_induct_tac,
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   248
    "induct_tac emulation (dynamic instantiation!)"),
9704
c2f2f70bbb60 'induct_tac' / 'case_tac': Method.goal_args';
wenzelm
parents: 9386
diff changeset
   249
  ("case_tac", Method.goal_args' (Scan.lift Args.name -- opt_rule) gen_case_tac,
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   250
    "case_tac emulation (dynamic instantiation!)")];
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   251
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   252
end;
8541
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   253
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   254
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   255
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   256
(** induct method setup **)
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   257
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   258
(* case names *)
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   259
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   260
local
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   261
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   262
fun dt_recs (DtTFree _) = []
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   263
  | dt_recs (DtType (_, dts)) = flat (map dt_recs dts)
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   264
  | dt_recs (DtRec i) = [i];
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   265
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   266
fun dt_cases (descr: descr) (_, args, constrs) =
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   267
  let
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   268
    fun the_bname i = Sign.base_name (#1 (the (assoc (descr, i))));
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   269
    val bnames = map the_bname (distinct (flat (map dt_recs args)));
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   270
  in map (fn (c, _) => space_implode "_" (Sign.base_name c :: bnames)) constrs end;
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   271
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   272
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   273
fun induct_cases descr =
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   274
  DatatypeProp.indexify_names (flat (map (dt_cases descr) (map #2 descr)));
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   275
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   276
fun exhaust_cases descr i = dt_cases descr (the (assoc (descr, i)));
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   277
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   278
in
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   279
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   280
fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   281
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   282
fun mk_case_names_exhausts descr new =
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   283
  map (RuleCases.case_names o exhaust_cases descr o #1)
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   284
    (filter (fn ((_, (name, _, _))) => name mem_string new) descr);
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   285
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   286
end;
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   287
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   288
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   289
(* add_cases_induct *)
8306
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   290
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   291
fun add_cases_induct infos =
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   292
  let
8325
80855ae484ce project induct rule;
wenzelm
parents: 8306
diff changeset
   293
    fun proj _ 1 thm = thm
80855ae484ce project induct rule;
wenzelm
parents: 8306
diff changeset
   294
      | proj i n thm =
80855ae484ce project induct rule;
wenzelm
parents: 8306
diff changeset
   295
          (if i + 1 < n then (fn th => th RS conjunct1) else I)
80855ae484ce project induct rule;
wenzelm
parents: 8306
diff changeset
   296
            (Library.funpow i (fn th => th RS conjunct2) thm)
8672
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   297
          |> Drule.standard
1f51c411da5a induct/case_tac emulation: optional rule;
wenzelm
parents: 8645
diff changeset
   298
          |> RuleCases.name (RuleCases.get thm);
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   299
8325
80855ae484ce project induct rule;
wenzelm
parents: 8306
diff changeset
   300
    fun add (ths, (name, {index, descr, induction, exhaustion, ...}: datatype_info)) =
10279
b223a9a3350e InductAttrib;
wenzelm
parents: 10214
diff changeset
   301
      (("", proj index (length descr) induction), [InductAttrib.induct_type_global name]) ::
b223a9a3350e InductAttrib;
wenzelm
parents: 10214
diff changeset
   302
      (("", exhaustion), [InductAttrib.cases_type_global name]) :: ths;
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   303
  in #1 o PureThy.add_thms (foldl add ([], infos)) end;
8306
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   304
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   305
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   306
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   307
(**** simplification procedure for showing distinctness of constructors ****)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   308
7060
berghofe
parents: 7049
diff changeset
   309
fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
berghofe
parents: 7049
diff changeset
   310
  | stripT p = p;
berghofe
parents: 7049
diff changeset
   311
berghofe
parents: 7049
diff changeset
   312
fun stripC (i, f $ x) = stripC (i + 1, f)
berghofe
parents: 7049
diff changeset
   313
  | stripC p = p;
berghofe
parents: 7049
diff changeset
   314
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   315
val distinctN = "constr_distinct";
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   316
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   317
exception ConstrDistinct of term;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   318
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   319
fun distinct_proc sg _ (t as Const ("op =", _) $ t1 $ t2) =
7060
berghofe
parents: 7049
diff changeset
   320
  (case (stripC (0, t1), stripC (0, t2)) of
berghofe
parents: 7049
diff changeset
   321
     ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
berghofe
parents: 7049
diff changeset
   322
         (case (stripT (0, T1), stripT (0, T2)) of
berghofe
parents: 7049
diff changeset
   323
            ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
berghofe
parents: 7049
diff changeset
   324
                if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   325
                   (case (constrs_of_sg sg tname1) of
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   326
                      Some constrs => let val cnames = map (fst o dest_Const) constrs
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   327
                        in if cname1 mem cnames andalso cname2 mem cnames then
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   328
                             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
   329
                                 val eq_ct = cterm_of sg eq_t;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   330
                                 val Datatype_thy = theory "Datatype";
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   331
                                 val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   332
                                   map (get_thm Datatype_thy)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   333
                                     ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"]
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   334
                             in (case (#distinct (datatype_info_sg_err sg tname1)) of
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   335
                                 QuickAndDirty => Some (Thm.invoke_oracle
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   336
                                   Datatype_thy distinctN (sg, ConstrDistinct eq_t))
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   337
                               | FewConstrs thms => Some (prove_goalw_cterm [] eq_ct (K
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   338
                                   [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   339
                                    atac 2, resolve_tac thms 1, etac FalseE 1]))
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   340
                               | ManyConstrs (thm, ss) => Some (prove_goalw_cterm [] eq_ct (K
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   341
                                   [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   342
                                    full_simp_tac ss 1,
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   343
                                    REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   344
                                    eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   345
                                    etac FalseE 1])))
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   346
                             end
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   347
                           else None
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   348
                        end
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   349
                    | None => None)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   350
                else None
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   351
          | _ => None)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   352
   | _ => None)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   353
  | distinct_proc sg _ _ = None;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   354
8100
6186ee807f2e replaced HOLogic.termTVar by HOLogic.termT;
wenzelm
parents: 7904
diff changeset
   355
val distinct_pats = [Thm.read_cterm (Theory.sign_of HOL.thy) ("s = t", HOLogic.termT)];
7049
f59d33c6e1c7 Eliminated addDistinct.
berghofe
parents: 7015
diff changeset
   356
val distinct_simproc = mk_simproc distinctN distinct_pats distinct_proc;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   357
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   358
val dist_ss = HOL_ss addsimprocs [distinct_simproc];
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   359
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   360
val simproc_setup =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   361
  [Theory.add_oracle (distinctN, fn (_, ConstrDistinct t) => t),
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   362
   fn thy => (simpset_ref_of thy := simpset_of thy addsimprocs [distinct_simproc]; thy)];
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   363
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   364
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   365
(* prepare types *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   366
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   367
fun read_typ sign ((Ts, sorts), str) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   368
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   369
    val T = Type.no_tvars (Sign.read_typ (sign, (curry assoc)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   370
      (map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   371
  in (Ts @ [T], add_typ_tfrees (T, sorts)) end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   372
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   373
fun cert_typ sign ((Ts, sorts), raw_T) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   374
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   375
    val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   376
      TYPE (msg, _, _) => error msg;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   377
    val sorts' = add_typ_tfrees (T, sorts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   378
  in (Ts @ [T],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   379
      case duplicates (map fst sorts') of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   380
         [] => sorts'
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   381
       | dups => error ("Inconsistent sort constraints for " ^ commas dups))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   382
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   383
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   384
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   385
(**** make datatype info ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   386
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   387
fun make_dt_info descr induct reccomb_names rec_thms
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   388
    (((((((((i, (_, (tname, _, _))), case_name), case_thms),
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   389
      exhaustion_thm), distinct_thm), inject), nchotomy), case_cong), weak_case_cong) =
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   390
  (tname,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   391
   {index = i,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   392
    descr = descr,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   393
    rec_names = reccomb_names,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   394
    rec_rewrites = rec_thms,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   395
    case_name = case_name,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   396
    case_rewrites = case_thms,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   397
    induction = induct,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   398
    exhaustion = exhaustion_thm,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   399
    distinct = distinct_thm,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   400
    inject = inject,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   401
    nchotomy = nchotomy,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   402
    case_cong = case_cong,
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   403
    weak_case_cong = weak_case_cong});
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   404
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   405
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   406
(********************* axiomatic introduction of datatypes ********************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   407
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   408
fun add_and_get_axioms_atts label tnames attss ts thy =
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   409
  foldr (fn (((tname, atts), t), (thy', axs)) =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   410
    let
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   411
      val (thy'', [ax]) = thy' |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   412
        Theory.add_path tname |>
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   413
        PureThy.add_axioms_i [((label, t), atts)];
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   414
    in (Theory.parent_path thy'', ax::axs)
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   415
    end) (tnames ~~ attss ~~ ts, (thy, []));
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   416
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   417
fun add_and_get_axioms label tnames =
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   418
  add_and_get_axioms_atts label tnames (replicate (length tnames) []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   419
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   420
fun add_and_get_axiomss label tnames tss thy =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   421
  foldr (fn ((tname, ts), (thy', axss)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   422
    let
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   423
      val (thy'', [axs]) = thy' |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   424
        Theory.add_path tname |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   425
        PureThy.add_axiomss_i [((label, ts), [])];
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   426
    in (Theory.parent_path thy'', axs::axss)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   427
    end) (tnames ~~ tss, (thy, []));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   428
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   429
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
   430
    case_names_induct case_names_exhausts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   431
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   432
    val descr' = flat descr;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   433
    val recTs = get_rec_types descr' sorts;
5578
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5279
diff changeset
   434
    val used = foldr add_typ_tfree_names (recTs, []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   435
    val newTs = take (length (hd descr), recTs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   436
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   437
    val no_size = exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   438
      (fn (DtType ("fun", [_, DtRec _])) => true | _ => false) cargs) constrs) descr';
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   439
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   440
    (**** declare new types and constants ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   441
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   442
    val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   443
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   444
    val constr_decls = map (fn (((_, (_, _, constrs)), T), constr_syntax') =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   445
      map (fn ((_, cargs), (cname, mx)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   446
        (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   447
          (constrs ~~ constr_syntax')) ((hd descr) ~~ newTs ~~ constr_syntax);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   448
5578
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5279
diff changeset
   449
    val rec_result_Ts = map TFree (variantlist (replicate (length descr') "'t", used) ~~
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5279
diff changeset
   450
      replicate (length descr') HOLogic.termS);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   451
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   452
    val reccomb_fn_Ts = flat (map (fn (i, (_, _, constrs)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   453
      map (fn (_, cargs) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   454
        let
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   455
          val Ts = map (typ_of_dtyp descr' sorts) cargs;
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   456
          val recs = filter (is_rec_type o fst) (cargs ~~ Ts);
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   457
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   458
          fun mk_argT (DtRec k, _) = nth_elem (k, rec_result_Ts)
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   459
            | mk_argT (DtType ("fun", [_, DtRec k]), Type ("fun", [T, _])) =
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   460
               T --> nth_elem (k, rec_result_Ts);
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   461
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   462
          val argTs = Ts @ map mk_argT recs
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   463
        in argTs ---> nth_elem (i, rec_result_Ts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   464
        end) constrs) descr');
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   465
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   466
    val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   467
    val reccomb_names = if length descr' = 1 then [big_reccomb_name] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   468
      (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   469
        (1 upto (length descr')));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   470
9739
8470c4662685 Improved names for size function.
berghofe
parents: 9714
diff changeset
   471
    val size_names = DatatypeProp.indexify_names
8470c4662685 Improved names for size function.
berghofe
parents: 9714
diff changeset
   472
      (map (fn T => name_of_typ T ^ "_size") (drop (length (hd descr), recTs)));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   473
5578
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5279
diff changeset
   474
    val freeT = TFree (variant used "'t", HOLogic.termS);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   475
    val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   476
      map (fn (_, cargs) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   477
        let val Ts = map (typ_of_dtyp descr' sorts) cargs
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   478
        in Ts ---> freeT end) constrs) (hd descr);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   479
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   480
    val case_names = map (fn s => (s ^ "_case")) new_type_names;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   481
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   482
    val thy2' = thy |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   483
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   484
      (** new types **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   485
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   486
      curry (foldr (fn (((name, mx), tvs), thy') => thy' |>
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   487
          TypedefPackage.add_typedecls [(name, tvs, mx)]))
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   488
        (types_syntax ~~ tyvars) |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   489
      add_path flat_names (space_implode "_" new_type_names) |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   490
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   491
      (** primrec combinators **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   492
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   493
      Theory.add_consts_i (map (fn ((name, T), T') =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   494
        (name, reccomb_fn_Ts @ [T] ---> T', NoSyn))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   495
          (reccomb_names ~~ recTs ~~ rec_result_Ts)) |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   496
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   497
      (** case combinators **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   498
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   499
      Theory.add_consts_i (map (fn ((name, T), Ts) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   500
        (name, Ts @ [T] ---> freeT, NoSyn))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   501
          (case_names ~~ newTs ~~ case_fn_Ts)) |>
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   502
      Theory.add_trrules_i (DatatypeProp.make_case_trrules new_type_names descr);
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   503
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6385
diff changeset
   504
    val reccomb_names' = map (Sign.intern_const (Theory.sign_of thy2')) reccomb_names;
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6385
diff changeset
   505
    val case_names' = map (Sign.intern_const (Theory.sign_of thy2')) case_names;
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   506
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   507
    val thy2 = thy2' |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   508
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   509
      (** size functions **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   510
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   511
      (if no_size then I else Theory.add_consts_i (map (fn (s, T) =>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   512
        (Sign.base_name s, T --> HOLogic.natT, NoSyn))
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   513
          (size_names ~~ drop (length (hd descr), recTs)))) |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   514
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   515
      (** constructors **)
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   516
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   517
      parent_path flat_names |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   518
      curry (foldr (fn (((((_, (_, _, constrs)), T), tname),
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   519
        constr_syntax'), thy') => thy' |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   520
          add_path flat_names tname |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   521
            Theory.add_consts_i (map (fn ((_, cargs), (cname, mx)) =>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   522
              (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   523
                (constrs ~~ constr_syntax')) |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   524
          parent_path flat_names))
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   525
            (hd descr ~~ newTs ~~ new_type_names ~~ constr_syntax);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   526
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   527
    (**** introduction of axioms ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   528
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   529
    val rec_axs = DatatypeProp.make_primrecs new_type_names descr sorts thy2;
9739
8470c4662685 Improved names for size function.
berghofe
parents: 9714
diff changeset
   530
    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
   531
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   532
    val (thy3, (([induct], [rec_thms]), inject)) =
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   533
      thy2 |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   534
      Theory.add_path (space_implode "_" new_type_names) |>
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   535
      PureThy.add_axioms_i [(("induct", DatatypeProp.make_ind descr sorts), [case_names_induct])] |>>>
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   536
      PureThy.add_axiomss_i [(("recs", rec_axs), [])] |>>
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   537
      (if no_size then I else #1 o PureThy.add_axiomss_i [(("size", size_axs), [])]) |>>
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   538
      Theory.parent_path |>>>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   539
      add_and_get_axiomss "inject" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   540
        (DatatypeProp.make_injs descr sorts);
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   541
    val size_thms = if no_size then [] else get_thms thy3 "size";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   542
    val (thy4, distinct) = add_and_get_axiomss "distinct" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   543
      (DatatypeProp.make_distincts new_type_names descr sorts thy3) thy3;
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   544
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   545
    val exhaust_ts = DatatypeProp.make_casedists descr sorts;
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   546
    val (thy5, exhaustion) = add_and_get_axioms_atts "exhaust" new_type_names
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   547
      (map Library.single case_names_exhausts) exhaust_ts thy4;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   548
    val (thy6, case_thms) = add_and_get_axiomss "cases" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   549
      (DatatypeProp.make_cases new_type_names descr sorts thy5) thy5;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   550
    val (split_ts, split_asm_ts) = ListPair.unzip
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   551
      (DatatypeProp.make_splits new_type_names descr sorts thy6);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   552
    val (thy7, split) = add_and_get_axioms "split" new_type_names split_ts thy6;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   553
    val (thy8, split_asm) = add_and_get_axioms "split_asm" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   554
      split_asm_ts thy7;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   555
    val (thy9, nchotomys) = add_and_get_axioms "nchotomy" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   556
      (DatatypeProp.make_nchotomys descr sorts) thy8;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   557
    val (thy10, case_congs) = add_and_get_axioms "case_cong" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   558
      (DatatypeProp.make_case_congs new_type_names descr sorts thy9) thy9;
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   559
    val (thy11, weak_case_congs) = add_and_get_axioms "weak_case_cong" new_type_names
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   560
      (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
   561
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   562
    val dt_infos = map (make_dt_info descr' induct reccomb_names' rec_thms)
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   563
      ((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
   564
        exhaustion ~~ replicate (length (hd descr)) QuickAndDirty ~~ inject ~~
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   565
          nchotomys ~~ case_congs ~~ weak_case_congs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   566
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   567
    val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
9386
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   568
    val split_thms = split ~~ split_asm;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   569
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   570
    val thy12 = thy11 |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   571
      Theory.add_path (space_implode "_" new_type_names) |>
8478
6053a5cd2881 Eliminated store_clasimp.
berghofe
parents: 8442
diff changeset
   572
      (#1 o PureThy.add_thmss [(("simps", simps), []),
6053a5cd2881 Eliminated store_clasimp.
berghofe
parents: 8442
diff changeset
   573
        (("", flat case_thms @ size_thms @ rec_thms), [Simplifier.simp_add_global]),
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   574
        (("", flat (inject @ distinct)), [iff_add_global]),
9714
79db0e5b7824 Simplifier.cong_add_global;
wenzelm
parents: 9704
diff changeset
   575
        (("", weak_case_congs), [Simplifier.cong_add_global])]) |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   576
      put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |>
8306
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   577
      add_cases_induct dt_infos |>
9386
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   578
      Theory.parent_path |>
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   579
      (#1 o store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   580
  in
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   581
    (thy12,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   582
     {distinct = distinct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   583
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   584
      exhaustion = exhaustion,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   585
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   586
      case_thms = case_thms,
9386
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   587
      split_thms = split_thms,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   588
      induction = induct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   589
      size = size_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   590
      simps = simps})
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   591
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   592
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   593
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   594
(******************* definitional introduction of datatypes *******************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   595
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   596
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
   597
    case_names_induct case_names_exhausts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   598
  let
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   599
    val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   600
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   601
    val (thy2, inject, distinct, dist_rewrites, simproc_dists, induct) = thy |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   602
      DatatypeRepProofs.representation_proofs flat_names dt_info new_type_names descr sorts
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   603
        types_syntax constr_syntax case_names_induct;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   604
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   605
    val (thy3, casedist_thms) = DatatypeAbsProofs.prove_casedist_thms new_type_names descr
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   606
      sorts induct case_names_exhausts thy2;
8478
6053a5cd2881 Eliminated store_clasimp.
berghofe
parents: 8442
diff changeset
   607
    val (thy4, (reccomb_names, rec_thms)) = DatatypeAbsProofs.prove_primrec_thms
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   608
      flat_names new_type_names descr sorts dt_info inject dist_rewrites dist_ss induct thy3;
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   609
    val (thy6, (case_thms, case_names)) = DatatypeAbsProofs.prove_case_thms
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   610
      flat_names new_type_names descr sorts reccomb_names rec_thms thy4;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   611
    val (thy7, split_thms) = DatatypeAbsProofs.prove_split_thms new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   612
      descr sorts inject dist_rewrites casedist_thms case_thms thy6;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   613
    val (thy8, nchotomys) = DatatypeAbsProofs.prove_nchotomys new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   614
      descr sorts casedist_thms thy7;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   615
    val (thy9, case_congs) = DatatypeAbsProofs.prove_case_congs new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   616
      descr sorts nchotomys case_thms thy8;
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   617
    val (thy10, weak_case_congs) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   618
      descr sorts thy9;
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   619
    val (thy11, size_thms) = DatatypeAbsProofs.prove_size_thms flat_names new_type_names
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   620
      descr sorts reccomb_names rec_thms thy10;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   621
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   622
    val dt_infos = map (make_dt_info (flat descr) induct reccomb_names rec_thms)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   623
      ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   624
        casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   625
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   626
    val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   627
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   628
    val thy12 = thy11 |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   629
      Theory.add_path (space_implode "_" new_type_names) |>
8478
6053a5cd2881 Eliminated store_clasimp.
berghofe
parents: 8442
diff changeset
   630
      (#1 o PureThy.add_thmss [(("simps", simps), []),
6053a5cd2881 Eliminated store_clasimp.
berghofe
parents: 8442
diff changeset
   631
        (("", flat case_thms @ size_thms @ rec_thms), [Simplifier.simp_add_global]),
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   632
        (("", flat (inject @ distinct)), [iff_add_global]),
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   633
        (("", weak_case_congs), [Simplifier.change_global_ss (op addcongs)])]) |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   634
      put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |>
8306
9855f1801d2b HOLogic.dest_conj;
wenzelm
parents: 8279
diff changeset
   635
      add_cases_induct dt_infos |>
9386
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   636
      Theory.parent_path |>
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   637
      (#1 o store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   638
  in
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   639
    (thy12,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   640
     {distinct = distinct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   641
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   642
      exhaustion = casedist_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   643
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   644
      case_thms = case_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   645
      split_thms = split_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   646
      induction = induct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   647
      size = size_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   648
      simps = simps})
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   649
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   650
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   651
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   652
(*********************** declare existing type as datatype *********************)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   653
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   654
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
   655
  let
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   656
    fun app_thmss srcs thy = foldl_map (fn (thy, x) => apply_theorems x thy) (thy, srcs);
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   657
    fun app_thm src thy = apsnd Library.hd (apply_theorems [src] thy);
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   658
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   659
    val (((thy1, induction), inject), distinct) = thy0
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   660
      |> app_thmss raw_distinct
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   661
      |> apfst (app_thmss raw_inject)
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   662
      |> apfst (apfst (app_thm raw_induction));
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6385
diff changeset
   663
    val sign = Theory.sign_of thy1;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   664
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   665
    val induction' = freezeT induction;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   666
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   667
    fun err t = error ("Ill-formed predicate in induction rule: " ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   668
      Sign.string_of_term sign t);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   669
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   670
    fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   671
          ((tname, map dest_TFree Ts) handle TERM _ => err t)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   672
      | get_typ t = err t;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   673
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   674
    val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induction')));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   675
    val new_type_names = if_none alt_names (map fst dtnames);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   676
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   677
    fun get_constr t = (case Logic.strip_assums_concl t of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   678
        _ $ (_ $ t') => (case head_of t' of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   679
            Const (cname, cT) => (case strip_type cT of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   680
                (Ts, Type (tname, _)) => (tname, (cname, map (dtyp_of_typ dtnames) Ts))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   681
              | _ => err t)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   682
          | _ => err t)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   683
      | _ => err t);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   684
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   685
    fun make_dt_spec [] _ _ = []
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   686
      | make_dt_spec ((tname, tvs)::dtnames') i constrs =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   687
          let val (constrs', constrs'') = take_prefix (equal tname o fst) constrs
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   688
          in (i, (tname, map DtTFree tvs, map snd constrs'))::
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   689
            (make_dt_spec dtnames' (i + 1) constrs'')
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   690
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   691
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   692
    val descr = make_dt_spec dtnames 0 (map get_constr (prems_of induction'));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   693
    val sorts = add_term_tfrees (concl_of induction', []);
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   694
    val dt_info = get_datatypes thy1;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   695
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   696
    val case_names_induct = mk_case_names_induct descr;
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   697
    val case_names_exhausts = mk_case_names_exhausts descr (map #1 dtnames);
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   698
    
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   699
6427
fd36b2e7d80e tuned messages;
wenzelm
parents: 6423
diff changeset
   700
    val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   701
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   702
    val (thy2, casedist_thms) = thy1 |>
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   703
      DatatypeAbsProofs.prove_casedist_thms new_type_names [descr] sorts induction
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   704
        case_names_exhausts;
8478
6053a5cd2881 Eliminated store_clasimp.
berghofe
parents: 8442
diff changeset
   705
    val (thy3, (reccomb_names, rec_thms)) = DatatypeAbsProofs.prove_primrec_thms
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   706
      false new_type_names [descr] sorts dt_info inject distinct dist_ss induction thy2;
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   707
    val (thy4, (case_thms, case_names)) = DatatypeAbsProofs.prove_case_thms false
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   708
      new_type_names [descr] sorts reccomb_names rec_thms thy3;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   709
    val (thy5, split_thms) = DatatypeAbsProofs.prove_split_thms
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   710
      new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   711
    val (thy6, nchotomys) = DatatypeAbsProofs.prove_nchotomys new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   712
      [descr] sorts casedist_thms thy5;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   713
    val (thy7, case_congs) = DatatypeAbsProofs.prove_case_congs new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   714
      [descr] sorts nchotomys case_thms thy6;
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   715
    val (thy8, weak_case_congs) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   716
      [descr] sorts thy7;
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   717
    val (thy9, size_thms) =
10214
77349ed89f45 *** empty log message ***
nipkow
parents: 10212
diff changeset
   718
      if exists (equal "NatArith") (Sign.stamp_names_of (Theory.sign_of thy8)) then
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   719
        DatatypeAbsProofs.prove_size_thms false new_type_names
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   720
          [descr] sorts reccomb_names rec_thms thy8
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   721
      else (thy8, []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   722
9149
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   723
    val (thy10, [induction']) = thy9 |>
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   724
      (#1 o store_thmss "inject" new_type_names inject) |>
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   725
      (#1 o store_thmss "distinct" new_type_names distinct) |>
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   726
      Theory.add_path (space_implode "_" new_type_names) |>
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   727
      PureThy.add_thms [(("induct", induction), [case_names_induct])];
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   728
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   729
    val dt_infos = map (make_dt_info descr induction' reccomb_names rec_thms)
10121
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   730
      ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
fb9be005cc44 export get_datatypes_sg;
wenzelm
parents: 9747
diff changeset
   731
        map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   732
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   733
    val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   734
9149
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   735
    val thy11 = thy10 |>
8478
6053a5cd2881 Eliminated store_clasimp.
berghofe
parents: 8442
diff changeset
   736
      (#1 o PureThy.add_thmss [(("simps", simps), []),
6053a5cd2881 Eliminated store_clasimp.
berghofe
parents: 8442
diff changeset
   737
        (("", flat case_thms @ size_thms @ rec_thms), [Simplifier.simp_add_global]),
8601
8fb3a81b4ccf added weak_case_cong feature
nipkow
parents: 8541
diff changeset
   738
        (("", flat (inject @ distinct)), [iff_add_global]),
9149
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   739
        (("", weak_case_congs), [Simplifier.change_global_ss (op addcongs)])]) |>
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   740
      put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |>
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   741
      add_cases_induct dt_infos |>
9386
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   742
      Theory.parent_path |>
8800603d99f6 theorems foo.splits = foo.split foo.split_asm;
wenzelm
parents: 9149
diff changeset
   743
      (#1 o store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   744
  in
9149
a126a40cba76 export proper induction rule;
wenzelm
parents: 8697
diff changeset
   745
    (thy11,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   746
     {distinct = distinct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   747
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   748
      exhaustion = casedist_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   749
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   750
      case_thms = case_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   751
      split_thms = split_thms,
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   752
      induction = induction',
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   753
      size = size_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   754
      simps = simps})
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   755
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   756
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   757
val rep_datatype = gen_rep_datatype IsarThy.apply_theorems;
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   758
val rep_datatype_i = gen_rep_datatype IsarThy.apply_theorems_i;
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   759
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   760
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   761
(******************************** add datatype ********************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   762
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   763
fun gen_add_datatype prep_typ flat_names new_type_names dts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   764
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   765
    val _ = Theory.requires thy "Datatype" "datatype definitions";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   766
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   767
    (* this theory is used just for parsing *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   768
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   769
    val tmp_thy = thy |>
5892
e5e44cc54eb2 Attribute.tthms_of;
wenzelm
parents: 5720
diff changeset
   770
      Theory.copy |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   771
      Theory.add_types (map (fn (tvs, tname, mx, _) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   772
        (tname, length tvs, mx)) dts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   773
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6385
diff changeset
   774
    val sign = Theory.sign_of tmp_thy;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   775
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   776
    val (tyvars, _, _, _)::_ = dts;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   777
    val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   778
      let val full_tname = Sign.full_name sign (Syntax.type_name tname mx)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   779
      in (case duplicates tvs of
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   780
            [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   781
                  else error ("Mutually recursive datatypes must have same type parameters")
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   782
          | dups => error ("Duplicate parameter(s) for datatype " ^ full_tname ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   783
              " : " ^ commas dups))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   784
      end) dts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   785
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   786
    val _ = (case duplicates (map fst new_dts) @ duplicates new_type_names of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   787
      [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   788
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   789
    fun prep_dt_spec ((dts', constr_syntax, sorts, i), (tvs, tname, mx, constrs)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   790
      let
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   791
        fun prep_constr ((constrs, constr_syntax', sorts'), (cname, cargs, mx')) =
5279
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   792
          let
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   793
            val (cargs', sorts'') = foldl (prep_typ sign) (([], sorts'), cargs);
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   794
            val _ = (case foldr add_typ_tfree_names (cargs', []) \\ tvs of
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   795
                [] => ()
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   796
              | vs => error ("Extra type variables on rhs: " ^ commas vs))
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   797
          in (constrs @ [((if flat_names then Sign.full_name sign else
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   798
                Sign.full_name_path sign tname) (Syntax.const_name cname mx'),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   799
                   map (dtyp_of_typ new_dts) cargs')],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   800
              constr_syntax' @ [(cname, mx')], sorts'')
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   801
          end handle ERROR =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   802
            error ("The error above occured in constructor " ^ cname ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   803
              " of datatype " ^ tname);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   804
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   805
        val (constrs', constr_syntax', sorts') =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   806
          foldl prep_constr (([], [], sorts), constrs)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   807
8405
0255394a05da add_cases_induct: produce proper case names;
wenzelm
parents: 8325
diff changeset
   808
      in
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   809
        case duplicates (map fst constrs') of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   810
           [] =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   811
             (dts' @ [(i, (Sign.full_name sign (Syntax.type_name tname mx),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   812
                map DtTFree tvs, constrs'))],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   813
              constr_syntax @ [constr_syntax'], sorts', i + 1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   814
         | dups => error ("Duplicate constructors " ^ commas dups ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   815
             " in datatype " ^ tname)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   816
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   817
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   818
    val (dts', constr_syntax, sorts', i) = foldl prep_dt_spec (([], [], [], 0), dts);
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   819
    val sorts = sorts' @ (map (rpair (Sign.defaultS sign)) (tyvars \\ map fst sorts'));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   820
    val dt_info = get_datatypes thy;
7015
85be09eb136c - Datatype package now also supports arbitrarily branching datatypes
berghofe
parents: 6851
diff changeset
   821
    val (descr, _) = unfold_datatypes sign dts' sorts dt_info dts' i;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   822
    val _ = check_nonempty descr;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   823
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   824
    val descr' = flat descr;
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   825
    val case_names_induct = mk_case_names_induct descr';
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   826
    val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   827
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   828
    (if (!quick_and_dirty) then add_datatype_axm else add_datatype_def)
8437
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   829
      flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
258281c43dea removed cases_of;
wenzelm
parents: 8405
diff changeset
   830
      case_names_induct case_names_exhausts thy
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   831
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   832
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   833
val add_datatype_i = gen_add_datatype cert_typ;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   834
val add_datatype = gen_add_datatype read_typ;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   835
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   836
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   837
(** package setup **)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   838
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   839
(* setup theory *)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   840
8541
b0d2002f9f04 proof methods: 'case_tac' / 'induct_tac';
wenzelm
parents: 8478
diff changeset
   841
val setup = [DatatypesData.init, Method.add_methods tactic_emulations] @ simproc_setup;
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   842
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   843
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   844
(* outer syntax *)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   845
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   846
local structure P = OuterParse and K = OuterSyntax.Keyword in
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   847
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   848
val datatype_decl =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   849
  Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
6729
b6e167580a32 formal comments (still dummy);
wenzelm
parents: 6723
diff changeset
   850
    (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix --| P.marg_comment));
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   851
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   852
fun mk_datatype args =
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   853
  let
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   854
    val names = map (fn ((((None, _), t), _), _) => t | ((((Some t, _), _), _), _) => t) args;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   855
    val specs = map (fn ((((_, vs), t), mx), cons) => (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   856
  in #1 o add_datatype false names specs end;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   857
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   858
val datatypeP =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   859
  OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   860
    (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   861
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   862
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   863
val rep_datatype_decl =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   864
  Scan.option (Scan.repeat1 P.name) --
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   865
    Scan.optional (P.$$$ "distinct" |-- P.!!! (P.and_list1 P.xthms1)) [] --
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   866
    Scan.optional (P.$$$ "inject" |-- P.!!! (P.and_list1 P.xthms1)) [] --
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   867
    (P.$$$ "induction" |-- P.!!! P.xthm);
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   868
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   869
fun mk_rep_datatype (((opt_ts, dss), iss), ind) = #1 o rep_datatype opt_ts dss iss ind;
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   870
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   871
val rep_datatypeP =
6723
f342449d73ca outer syntax keyword classification;
wenzelm
parents: 6556
diff changeset
   872
  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
   873
    (rep_datatype_decl >> (Toplevel.theory o mk_rep_datatype));
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   874
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   875
6479
b0448cab1b1e rep_datatype syntax: 'induction' instead of 'induct';
wenzelm
parents: 6441
diff changeset
   876
val _ = OuterSyntax.add_keywords ["distinct", "inject", "induction"];
6385
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   877
val _ = OuterSyntax.add_parsers [datatypeP, rep_datatypeP];
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   878
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   879
end;
5a6570458d9e rep_datatype: '_i' version, attributes, outer syntax;
wenzelm
parents: 6360
diff changeset
   880
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   881
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   882
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   883
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   884
structure BasicDatatypePackage: BASIC_DATATYPE_PACKAGE = DatatypePackage;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   885
open BasicDatatypePackage;