src/HOL/Tools/datatype_package.ML
author wenzelm
Thu, 11 Mar 1999 21:59:26 +0100
changeset 6360 83573ae0f22c
parent 6305 4cbdb974220c
child 6385 5a6570458d9e
permissions -rw-r--r--
outer syntax for 'datatype';
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.
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
     7
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
     8
TODO:
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
     9
  - streamline internal interfaces (!??);
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    10
  - rep_datatype outer syntax ('and' vs. ',' (!?));
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    11
  - methods: induct, exhaust;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    12
*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    13
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    14
signature BASIC_DATATYPE_PACKAGE =
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    15
sig
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    16
  val mutual_induct_tac : string list -> int -> tactic
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    17
  val induct_tac : string -> int -> tactic
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    18
  val exhaust_tac : string -> int -> tactic
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    19
end;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    20
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    21
signature DATATYPE_PACKAGE =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    22
sig
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    23
  include BASIC_DATATYPE_PACKAGE
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    24
  val quiet_mode : bool ref
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    25
  val add_datatype : bool -> string list -> (string list * bstring * mixfix *
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    26
    (bstring * string list * mixfix) list) list -> theory -> theory *
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    27
      {distinct : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    28
       inject : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    29
       exhaustion : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    30
       rec_thms : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    31
       case_thms : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    32
       split_thms : (thm * thm) list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    33
       induction : thm,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    34
       size : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    35
       simps : thm list}
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    36
  val add_datatype_i : bool -> string list -> (string list * bstring * mixfix *
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    37
    (bstring * typ list * mixfix) list) list -> theory -> theory *
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    38
      {distinct : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    39
       inject : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    40
       exhaustion : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    41
       rec_thms : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    42
       case_thms : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    43
       split_thms : (thm * thm) list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    44
       induction : thm,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    45
       size : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    46
       simps : thm list}
6103
36f272ea9413 get_tthms witness theorems;
wenzelm
parents: 6092
diff changeset
    47
  val rep_datatype : string list option -> thm list list ->
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    48
    thm list list -> thm -> theory -> theory *
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    49
      {distinct : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    50
       inject : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    51
       exhaustion : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    52
       rec_thms : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    53
       case_thms : thm list list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    54
       split_thms : (thm * thm) list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    55
       induction : thm,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    56
       size : thm list,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    57
       simps : thm list}
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    58
  val get_datatypes : theory -> DatatypeAux.datatype_info Symtab.table
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    59
  val datatype_info_sg : Sign.sg -> string -> DatatypeAux.datatype_info
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    60
  val datatype_info : theory -> string -> DatatypeAux.datatype_info
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    61
  val constrs_of : theory -> string -> term list option
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    62
  val case_const_of : theory -> string -> term option
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    63
  val setup: (theory -> theory) list
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    64
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    65
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    66
structure DatatypePackage : DATATYPE_PACKAGE =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    67
struct
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    68
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    69
open DatatypeAux;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    70
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    71
val quiet_mode = quiet_mode;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
    72
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
    73
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    74
(* data kind 'HOL/datatypes' *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    75
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    76
structure DatatypesArgs =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    77
struct
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    78
  val name = "HOL/datatypes";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    79
  type T = datatype_info Symtab.table;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    80
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    81
  val empty = Symtab.empty;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    82
  val prep_ext = I;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    83
  val merge: T * T -> T = Symtab.merge (K true);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    84
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    85
  fun print sg tab =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    86
    Pretty.writeln (Pretty.strs ("datatypes:" ::
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    87
      map (Sign.cond_extern sg Sign.typeK o fst) (Symtab.dest tab)));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    88
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    89
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    90
structure DatatypesData = TheoryDataFun(DatatypesArgs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    91
val get_datatypes_sg = DatatypesData.get_sg;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    92
val get_datatypes = DatatypesData.get;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    93
val put_datatypes = DatatypesData.put;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    94
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    95
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    96
(** theory information about datatypes **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    97
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    98
fun datatype_info_sg sg name =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
    99
  (case Symtab.lookup (get_datatypes_sg sg, name) of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   100
    Some info => info
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   101
  | None => error ("Unknown datatype " ^ quote name));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   102
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   103
val datatype_info = datatype_info_sg o sign_of;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   104
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   105
fun constrs_of thy tname =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   106
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   107
    val {index, descr, ...} = datatype_info thy tname;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   108
    val (_, _, constrs) = the (assoc (descr, index))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   109
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   110
    Some (map (fn (cname, _) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   111
      Const (cname, the (Sign.const_type (sign_of thy) cname))) constrs)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   112
  end handle _ => None;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   113
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   114
fun case_const_of thy tname =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   115
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   116
    val {case_name, ...} = datatype_info thy tname;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   117
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   118
    Some (Const (case_name, the (Sign.const_type (sign_of thy) case_name)))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   119
  end handle _ => None;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   120
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   121
fun find_tname var Bi =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   122
  let val frees = map dest_Free (term_frees Bi)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   123
      val params = Logic.strip_params Bi;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   124
  in case assoc (frees @ params, var) of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   125
       None => error ("No such variable in subgoal: " ^ quote var)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   126
     | Some(Type (tn, _)) => tn
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   127
     | _ => error ("Cannot determine type of " ^ quote var)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   128
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   129
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   130
fun infer_tname state sign i aterm =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   131
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   132
    val (_, _, Bi, _) = dest_state (state, i)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   133
    val params = Logic.strip_params Bi;   (*params of subgoal i*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   134
    val params = rev (rename_wrt_term Bi params);   (*as they are printed*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   135
    val (types, sorts) = types_sorts state;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   136
    fun types' (a, ~1) = (case assoc (params, a) of None => types(a, ~1) | sm => sm)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   137
      | types' ixn = types ixn;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   138
    val (ct, _) = read_def_cterm (sign, types', sorts) [] false
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   139
                                  (aterm, TVar (("", 0), []));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   140
  in case #T (rep_cterm ct) of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   141
       Type (tn, _) => tn
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   142
     | _ => error ("Cannot determine type of " ^ quote aterm)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   143
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   144
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   145
(*Warn if the (induction) variable occurs Free among the premises, which
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   146
  usually signals a mistake.  But calls the tactic either way!*)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   147
fun occs_in_prems tacf vars = 
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   148
  SUBGOAL (fn (Bi, i) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   149
	   (if  exists (fn Free (a, _) => a mem vars)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   150
	              (foldr add_term_frees (#2 (strip_context Bi), []))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   151
	     then warning "Induction variable occurs also among premises!"
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   152
	     else ();
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   153
	    tacf i));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   154
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   155
(* generic induction tactic for datatypes *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   156
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   157
fun mutual_induct_tac vars i state =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   158
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   159
    val (_, _, Bi, _) = dest_state (state, i);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   160
    val {sign, ...} = rep_thm state;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   161
    val tn = find_tname (hd vars) Bi;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   162
    val {induction, ...} = datatype_info_sg sign tn;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   163
    val ind_vnames = map (fn (_ $ Var (ixn, _)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   164
      implode (tl (explode (Syntax.string_of_vname ixn))))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   165
        (dest_conj (HOLogic.dest_Trueprop (concl_of induction)));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   166
    val insts = (ind_vnames ~~ vars) handle _ =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   167
      error ("Induction rule for type " ^ tn ^ " has different number of variables")
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   168
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   169
    occs_in_prems (res_inst_tac insts induction) vars i state
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   170
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   171
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   172
fun induct_tac var = mutual_induct_tac [var];
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   173
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   174
(* generic exhaustion tactic for datatypes *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   175
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   176
fun exhaust_tac aterm i state =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   177
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   178
    val {sign, ...} = rep_thm state;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   179
    val tn = infer_tname state sign i aterm;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   180
    val {exhaustion, ...} = datatype_info_sg sign tn;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   181
    val _ $ Var (ixn, _) $ _ = HOLogic.dest_Trueprop
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   182
      (hd (Logic.strip_assums_hyp (hd (prems_of exhaustion))));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   183
    val exh_vname = implode (tl (explode (Syntax.string_of_vname ixn)))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   184
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   185
    res_inst_tac [(exh_vname, aterm)] exhaustion i state
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   186
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   187
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   188
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   189
(* prepare types *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   190
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   191
fun read_typ sign ((Ts, sorts), str) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   192
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   193
    val T = Type.no_tvars (Sign.read_typ (sign, (curry assoc)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   194
      (map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   195
  in (Ts @ [T], add_typ_tfrees (T, sorts)) end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   196
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   197
fun cert_typ sign ((Ts, sorts), raw_T) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   198
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   199
    val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   200
      TYPE (msg, _, _) => error msg;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   201
    val sorts' = add_typ_tfrees (T, sorts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   202
  in (Ts @ [T],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   203
      case duplicates (map fst sorts') of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   204
         [] => sorts'
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   205
       | dups => error ("Inconsistent sort constraints for " ^ commas dups))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   206
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   207
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   208
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   209
(**** make datatype info ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   210
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   211
fun make_dt_info descr induct reccomb_names rec_thms
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   212
  ((((((((i, (_, (tname, _, _))), case_name), case_thms),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   213
    exhaustion_thm), distinct_thm), inject), nchotomy), case_cong) = (tname,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   214
      {index = i,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   215
       descr = descr,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   216
       rec_names = reccomb_names,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   217
       rec_rewrites = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   218
       case_name = case_name,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   219
       case_rewrites = case_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   220
       induction = induct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   221
       exhaustion = exhaustion_thm,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   222
       distinct = distinct_thm,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   223
       inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   224
       nchotomy = nchotomy,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   225
       case_cong = case_cong});
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   226
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   227
fun store_clasimp thy (cla, simp) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   228
  (claset_ref_of thy := cla; simpset_ref_of thy := simp);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   229
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   230
infix 4 addDistinct;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   231
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   232
fun clasimp addDistinct ([], _) = clasimp
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   233
  | clasimp addDistinct (thms::thmss, (_, (_, _, constrs))::descr) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   234
      if length constrs < DatatypeProp.dtK then
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   235
        clasimp addIffs thms addDistinct (thmss, descr)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   236
      else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   237
        clasimp addsimps2 thms addDistinct (thmss, descr);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   238
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   239
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   240
(********************* axiomatic introduction of datatypes ********************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   241
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   242
fun add_and_get_axioms label tnames ts thy =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   243
  foldr (fn ((tname, t), (thy', axs)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   244
    let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   245
      val thy'' = thy' |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   246
        Theory.add_path tname |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   247
        PureThy.add_axioms_i [((label, t), [])];
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   248
      val ax = get_axiom thy'' label
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   249
    in (Theory.parent_path thy'', ax::axs)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   250
    end) (tnames ~~ ts, (thy, []));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   251
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   252
fun add_and_get_axiomss label tnames tss thy =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   253
  foldr (fn ((tname, ts), (thy', axss)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   254
    let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   255
      val thy'' = thy' |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   256
        Theory.add_path tname |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   257
        PureThy.add_axiomss_i [((label, ts), [])];
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   258
      val axs = PureThy.get_thms thy'' label
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   259
    in (Theory.parent_path thy'', axs::axss)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   260
    end) (tnames ~~ tss, (thy, []));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   261
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   262
fun add_datatype_axm flat_names new_type_names descr sorts types_syntax constr_syntax dt_info thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   263
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   264
    val descr' = flat descr;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   265
    val recTs = get_rec_types descr' sorts;
5578
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5279
diff changeset
   266
    val used = foldr add_typ_tfree_names (recTs, []);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   267
    val newTs = take (length (hd descr), recTs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   268
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   269
    val _ = message ("Adding axioms for datatype(s) " ^ commas_quote new_type_names);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   270
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   271
    (**** declare new types and constants ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   272
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   273
    val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   274
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   275
    val constr_decls = map (fn (((_, (_, _, constrs)), T), constr_syntax') =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   276
      map (fn ((_, cargs), (cname, mx)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   277
        (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   278
          (constrs ~~ constr_syntax')) ((hd descr) ~~ newTs ~~ constr_syntax);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   279
5578
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5279
diff changeset
   280
    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
   281
      replicate (length descr') HOLogic.termS);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   282
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   283
    val reccomb_fn_Ts = flat (map (fn (i, (_, _, constrs)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   284
      map (fn (_, cargs) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   285
        let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   286
          val recs = filter is_rec_type cargs;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   287
          val argTs = (map (typ_of_dtyp descr' sorts) cargs) @
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   288
            (map (fn r => nth_elem (dest_DtRec r, rec_result_Ts)) recs)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   289
        in argTs ---> nth_elem (i, rec_result_Ts)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   290
        end) constrs) descr');
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   291
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   292
    val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   293
    val reccomb_names = if length descr' = 1 then [big_reccomb_name] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   294
      (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   295
        (1 upto (length descr')));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   296
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   297
    val big_size_name = space_implode "_" new_type_names ^ "_size";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   298
    val size_names = if length (flat (tl descr)) = 1 then [big_size_name] else
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   299
      map (fn i => big_size_name ^ "_" ^ string_of_int i)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   300
        (1 upto length (flat (tl descr)));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   301
5578
7de426cf179c Package now chooses type variable names more carefully to
berghofe
parents: 5279
diff changeset
   302
    val freeT = TFree (variant used "'t", HOLogic.termS);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   303
    val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   304
      map (fn (_, cargs) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   305
        let val Ts = map (typ_of_dtyp descr' sorts) cargs
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   306
        in Ts ---> freeT end) constrs) (hd descr);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   307
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   308
    val case_names = map (fn s => (s ^ "_case")) new_type_names;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   309
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   310
    val thy2' = thy |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   311
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   312
      (** new types **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   313
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   314
      curry (foldr (fn (((name, mx), tvs), thy') => thy' |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   315
        PureThy.add_typedecls [(name, tvs, mx)] |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   316
        Theory.add_arities_i
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   317
          [(Sign.full_name (sign_of thy') (Syntax.type_name name mx),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   318
            replicate (length tvs) HOLogic.termS, HOLogic.termS)]))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   319
              (types_syntax ~~ tyvars) |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   320
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   321
      add_path flat_names (space_implode "_" new_type_names) |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   322
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   323
      (** primrec combinators **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   324
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   325
      Theory.add_consts_i (map (fn ((name, T), T') =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   326
        (name, reccomb_fn_Ts @ [T] ---> T', NoSyn))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   327
          (reccomb_names ~~ recTs ~~ rec_result_Ts)) |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   328
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   329
      (** case combinators **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   330
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   331
      Theory.add_consts_i (map (fn ((name, T), Ts) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   332
        (name, Ts @ [T] ---> freeT, NoSyn))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   333
          (case_names ~~ newTs ~~ case_fn_Ts)) |>
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   334
      Theory.add_trrules_i (DatatypeProp.make_case_trrules new_type_names descr);
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   335
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   336
    val reccomb_names' = map (Sign.intern_const (sign_of thy2')) reccomb_names;
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   337
    val case_names' = map (Sign.intern_const (sign_of thy2')) case_names;
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   338
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   339
    val thy2 = thy2' |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   340
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   341
      (** t_ord functions **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   342
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   343
      Theory.add_consts_i
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   344
        (foldr (fn ((((_, (_, _, constrs)), tname), T), decls) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   345
          if length constrs < DatatypeProp.dtK then decls
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   346
          else (tname ^ "_ord", T --> HOLogic.natT, NoSyn)::decls)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   347
            ((hd descr) ~~ new_type_names ~~ newTs, [])) |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   348
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   349
      (** size functions **)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   350
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   351
      Theory.add_consts_i (map (fn (s, T) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   352
        (Sign.base_name s, T --> HOLogic.natT, NoSyn))
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   353
          (size_names ~~ drop (length (hd descr), recTs))) |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   354
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   355
      (** constructors **)
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   356
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   357
      parent_path flat_names |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   358
      curry (foldr (fn (((((_, (_, _, constrs)), T), tname),
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   359
        constr_syntax'), thy') => thy' |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   360
          add_path flat_names tname |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   361
            Theory.add_consts_i (map (fn ((_, cargs), (cname, mx)) =>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   362
              (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   363
                (constrs ~~ constr_syntax')) |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   364
          parent_path flat_names))
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   365
            (hd descr ~~ newTs ~~ new_type_names ~~ constr_syntax);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   366
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   367
    (**** introduction of axioms ****)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   368
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   369
    val rec_axs = DatatypeProp.make_primrecs new_type_names descr sorts thy2;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   370
    val size_axs = DatatypeProp.make_size new_type_names descr sorts thy2;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   371
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   372
    val (thy3, inject) = thy2 |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   373
      Theory.add_path (space_implode "_" new_type_names) |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   374
      PureThy.add_axioms_i [(("induct", DatatypeProp.make_ind descr sorts), [])] |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   375
      PureThy.add_axiomss_i [(("recs", rec_axs), [])] |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   376
      PureThy.add_axiomss_i [(("size", size_axs), [])] |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   377
      Theory.parent_path |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   378
      add_and_get_axiomss "inject" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   379
        (DatatypeProp.make_injs descr sorts);
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   380
    val induct = get_axiom thy3 "induct";
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   381
    val rec_thms = get_thms thy3 "recs";
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   382
    val size_thms = get_thms thy3 "size";
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   383
    val (thy4, distinct) = add_and_get_axiomss "distinct" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   384
      (DatatypeProp.make_distincts new_type_names descr sorts thy3) thy3;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   385
    val (thy5, exhaustion) = add_and_get_axioms "exhaust" new_type_names
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   386
      (DatatypeProp.make_casedists descr sorts) thy4;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   387
    val (thy6, case_thms) = add_and_get_axiomss "cases" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   388
      (DatatypeProp.make_cases new_type_names descr sorts thy5) thy5;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   389
    val (split_ts, split_asm_ts) = ListPair.unzip
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   390
      (DatatypeProp.make_splits new_type_names descr sorts thy6);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   391
    val (thy7, split) = add_and_get_axioms "split" new_type_names split_ts thy6;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   392
    val (thy8, split_asm) = add_and_get_axioms "split_asm" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   393
      split_asm_ts thy7;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   394
    val (thy9, nchotomys) = add_and_get_axioms "nchotomy" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   395
      (DatatypeProp.make_nchotomys descr sorts) thy8;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   396
    val (thy10, case_congs) = add_and_get_axioms "case_cong" new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   397
      (DatatypeProp.make_case_congs new_type_names descr sorts thy9) thy9;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   398
    
6305
4cbdb974220c Fixed bug in add_datatype_axm:
berghofe
parents: 6103
diff changeset
   399
    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
   400
      ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names' ~~ case_thms ~~
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   401
        exhaustion ~~ distinct ~~ inject ~~ nchotomys ~~ case_congs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   402
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   403
    val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   404
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   405
    val thy11 = thy10 |>
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   406
      Theory.add_path (space_implode "_" new_type_names) |>
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5892
diff changeset
   407
      PureThy.add_thmss [(("simps", simps), [])] |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   408
      put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   409
      Theory.parent_path;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   410
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   411
    val _ = store_clasimp thy11 ((claset_of thy11, simpset_of thy11)
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   412
      addsimps2 flat case_thms addsimps2 size_thms addsimps2 rec_thms
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   413
      addIffs flat inject addDistinct (distinct, hd descr));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   414
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   415
  in
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   416
    (thy11,
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   417
     {distinct = distinct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   418
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   419
      exhaustion = exhaustion,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   420
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   421
      case_thms = case_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   422
      split_thms = split ~~ split_asm,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   423
      induction = induct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   424
      size = size_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   425
      simps = simps})
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   426
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   427
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   428
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   429
(******************* definitional introduction of datatypes *******************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   430
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   431
fun add_datatype_def flat_names new_type_names descr sorts types_syntax constr_syntax dt_info thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   432
  let
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   433
    val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   434
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   435
    val (thy2, inject, dist_rewrites, induct) = thy |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   436
      DatatypeRepProofs.representation_proofs flat_names dt_info new_type_names descr sorts
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   437
        types_syntax constr_syntax;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   438
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   439
    val (thy3, casedist_thms) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   440
      DatatypeAbsProofs.prove_casedist_thms new_type_names descr sorts induct thy2;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   441
    val (thy4, reccomb_names, rec_thms) = DatatypeAbsProofs.prove_primrec_thms
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   442
      flat_names new_type_names descr sorts dt_info inject dist_rewrites induct thy3;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   443
    val (thy5, case_names, case_thms) = DatatypeAbsProofs.prove_case_thms
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   444
      flat_names new_type_names descr sorts reccomb_names rec_thms thy4;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   445
    val (thy6, distinct) = DatatypeAbsProofs.prove_distinctness_thms
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   446
      flat_names new_type_names descr sorts dist_rewrites case_thms thy5;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   447
    val (thy7, split_thms) = DatatypeAbsProofs.prove_split_thms new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   448
      descr sorts inject dist_rewrites casedist_thms case_thms thy6;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   449
    val (thy8, nchotomys) = DatatypeAbsProofs.prove_nchotomys new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   450
      descr sorts casedist_thms thy7;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   451
    val (thy9, case_congs) = DatatypeAbsProofs.prove_case_congs new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   452
      descr sorts nchotomys case_thms thy8;
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   453
    val (thy10, size_thms) = DatatypeAbsProofs.prove_size_thms flat_names new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   454
      descr sorts reccomb_names rec_thms thy9;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   455
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   456
    val dt_infos = map (make_dt_info (flat descr) induct reccomb_names rec_thms)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   457
      ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   458
        casedist_thms ~~ distinct ~~ inject ~~ nchotomys ~~ case_congs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   459
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   460
    val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   461
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   462
    val thy11 = thy10 |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   463
      Theory.add_path (space_implode "_" new_type_names) |>
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5892
diff changeset
   464
      PureThy.add_thmss [(("simps", simps), [])] |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   465
      put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |>
5663
aad79a127628 Fixed bug (improper handling of flag flat_names).
berghofe
parents: 5661
diff changeset
   466
      Theory.parent_path;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   467
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   468
    val _ = store_clasimp thy11 ((claset_of thy11, simpset_of thy11)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   469
      addsimps2 flat case_thms addsimps2 size_thms addsimps2 rec_thms
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   470
      addIffs flat inject addDistinct (distinct, hd descr));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   471
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   472
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   473
    (thy11,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   474
     {distinct = distinct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   475
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   476
      exhaustion = casedist_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   477
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   478
      case_thms = case_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   479
      split_thms = split_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   480
      induction = induct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   481
      size = size_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   482
      simps = simps})
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   483
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   484
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   485
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   486
(*********************** declare non-datatype as datatype *********************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   487
6103
36f272ea9413 get_tthms witness theorems;
wenzelm
parents: 6092
diff changeset
   488
fun rep_datatype alt_names distinct inject induction thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   489
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   490
    val sign = sign_of thy;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   491
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   492
    val induction' = freezeT induction;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   493
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   494
    fun err t = error ("Ill-formed predicate in induction rule: " ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   495
      Sign.string_of_term sign t);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   496
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   497
    fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   498
          ((tname, map dest_TFree Ts) handle _ => err t)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   499
      | get_typ t = err t;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   500
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   501
    val dtnames = map get_typ (dest_conj (HOLogic.dest_Trueprop (concl_of induction')));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   502
    val new_type_names = if_none alt_names (map fst dtnames);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   503
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   504
    fun get_constr t = (case Logic.strip_assums_concl t of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   505
        _ $ (_ $ t') => (case head_of t' of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   506
            Const (cname, cT) => (case strip_type cT of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   507
                (Ts, Type (tname, _)) => (tname, (cname, map (dtyp_of_typ dtnames) Ts))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   508
              | _ => err t)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   509
          | _ => err t)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   510
      | _ => err t);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   511
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   512
    fun make_dt_spec [] _ _ = []
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   513
      | make_dt_spec ((tname, tvs)::dtnames') i constrs =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   514
          let val (constrs', constrs'') = take_prefix (equal tname o fst) constrs
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   515
          in (i, (tname, map DtTFree tvs, map snd constrs'))::
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   516
            (make_dt_spec dtnames' (i + 1) constrs'')
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   517
          end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   518
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   519
    val descr = make_dt_spec dtnames 0 (map get_constr (prems_of induction'));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   520
    val sorts = add_term_tfrees (concl_of induction', []);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   521
    val dt_info = get_datatypes thy;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   522
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   523
    val _ = writeln ("Proofs for datatype(s) " ^ commas new_type_names);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   524
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   525
    val (thy2, casedist_thms) = thy |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   526
      DatatypeAbsProofs.prove_casedist_thms new_type_names [descr] sorts induction;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   527
    val (thy3, reccomb_names, rec_thms) = DatatypeAbsProofs.prove_primrec_thms
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   528
      false new_type_names [descr] sorts dt_info inject distinct induction thy2;
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   529
    val (thy4, case_names, case_thms) = DatatypeAbsProofs.prove_case_thms false
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   530
      new_type_names [descr] sorts reccomb_names rec_thms thy3;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   531
    val (thy5, split_thms) = DatatypeAbsProofs.prove_split_thms
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   532
      new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   533
    val (thy6, nchotomys) = DatatypeAbsProofs.prove_nchotomys new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   534
      [descr] sorts casedist_thms thy5;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   535
    val (thy7, case_congs) = DatatypeAbsProofs.prove_case_congs new_type_names
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   536
      [descr] sorts nchotomys case_thms thy6;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   537
    val (thy8, size_thms) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   538
      if exists (equal "Arith") (Sign.stamp_names_of (sign_of thy7)) then
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   539
        DatatypeAbsProofs.prove_size_thms false new_type_names
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   540
          [descr] sorts reccomb_names rec_thms thy7
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   541
      else (thy7, []);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   542
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   543
    val dt_infos = map (make_dt_info descr induction reccomb_names rec_thms)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   544
      ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   545
        casedist_thms ~~ distinct ~~ inject ~~ nchotomys ~~ case_congs);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   546
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   547
    val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   548
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   549
    val thy9 = thy8 |>
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   550
      Theory.add_path (space_implode "_" new_type_names) |>
6092
d9db67970c73 eliminated tthm type and Attribute structure;
wenzelm
parents: 5892
diff changeset
   551
      PureThy.add_thmss [(("simps", simps), [])] |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   552
      put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   553
      Theory.parent_path;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   554
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   555
    val _ = store_clasimp thy9 ((claset_of thy9, simpset_of thy9)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   556
      addsimps2 flat case_thms addsimps2 size_thms addsimps2 rec_thms
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   557
      addIffs flat inject addDistinct (distinct, descr));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   558
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   559
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   560
    (thy9,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   561
     {distinct = distinct,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   562
      inject = inject,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   563
      exhaustion = casedist_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   564
      rec_thms = rec_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   565
      case_thms = case_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   566
      split_thms = split_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   567
      induction = induction,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   568
      size = size_thms,
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   569
      simps = simps})
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   570
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   571
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   572
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   573
(******************************** add datatype ********************************)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   574
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   575
fun gen_add_datatype prep_typ flat_names new_type_names dts thy =
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   576
  let
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   577
    val _ = Theory.requires thy "Datatype" "datatype definitions";
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   578
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   579
    (* this theory is used just for parsing *)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   580
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   581
    val tmp_thy = thy |>
5892
e5e44cc54eb2 Attribute.tthms_of;
wenzelm
parents: 5720
diff changeset
   582
      Theory.copy |>
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   583
      Theory.add_types (map (fn (tvs, tname, mx, _) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   584
        (tname, length tvs, mx)) dts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   585
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   586
    val sign = sign_of tmp_thy;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   587
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   588
    val (tyvars, _, _, _)::_ = dts;
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   589
    val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   590
      let val full_tname = Sign.full_name sign (Syntax.type_name tname mx)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   591
      in (case duplicates tvs of
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   592
            [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   593
                  else error ("Mutually recursive datatypes must have same type parameters")
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   594
          | dups => error ("Duplicate parameter(s) for datatype " ^ full_tname ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   595
              " : " ^ commas dups))
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   596
      end) dts);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   597
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   598
    val _ = (case duplicates (map fst new_dts) @ duplicates new_type_names of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   599
      [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   600
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   601
    fun prep_dt_spec ((dts', constr_syntax, sorts, i), (tvs, tname, mx, constrs)) =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   602
      let
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   603
        fun prep_constr ((constrs, constr_syntax', sorts'), (cname, cargs, mx')) =
5279
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   604
          let
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   605
            val (cargs', sorts'') = foldl (prep_typ sign) (([], sorts'), cargs);
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   606
            val _ = (case foldr add_typ_tfree_names (cargs', []) \\ tvs of
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   607
                [] => ()
cba6a96f5812 Improved well-formedness check.
berghofe
parents: 5177
diff changeset
   608
              | vs => error ("Extra type variables on rhs: " ^ commas vs))
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   609
          in (constrs @ [((if flat_names then Sign.full_name sign else
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   610
                Sign.full_name_path sign tname) (Syntax.const_name cname mx'),
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   611
                   map (dtyp_of_typ new_dts) cargs')],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   612
              constr_syntax' @ [(cname, mx')], sorts'')
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   613
          end handle ERROR =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   614
            error ("The error above occured in constructor " ^ cname ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   615
              " of datatype " ^ tname);
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   616
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   617
        val (constrs', constr_syntax', sorts') =
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   618
          foldl prep_constr (([], [], sorts), constrs)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   619
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   620
      in 
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   621
        case duplicates (map fst constrs') of
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   622
           [] =>
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   623
             (dts' @ [(i, (Sign.full_name sign (Syntax.type_name tname mx),
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   624
                map DtTFree tvs, constrs'))],
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   625
              constr_syntax @ [constr_syntax'], sorts', i + 1)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   626
         | dups => error ("Duplicate constructors " ^ commas dups ^
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   627
             " in datatype " ^ tname)
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   628
      end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   629
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   630
    val (dts', constr_syntax, sorts', i) = foldl prep_dt_spec (([], [], [], 0), dts);
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   631
    val dt_info = get_datatypes thy;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   632
    val (descr, _) = unfold_datatypes dt_info dts' i;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   633
    val _ = check_nonempty descr;
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   634
    val sorts = sorts' @ (map (rpair (Sign.defaultS sign)) (tyvars \\ map fst sorts'));
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   635
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   636
  in
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   637
    (if (!quick_and_dirty) then add_datatype_axm else add_datatype_def)
5661
6ecb6ea25f19 - Changed structure of name spaces
berghofe
parents: 5578
diff changeset
   638
      flat_names new_type_names descr sorts types_syntax constr_syntax dt_info thy
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   639
  end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   640
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   641
val add_datatype_i = gen_add_datatype cert_typ;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   642
val add_datatype = gen_add_datatype read_typ;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   643
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   644
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   645
(** package setup **)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   646
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   647
(* setup theory *)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   648
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   649
val setup = [DatatypesData.init];
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   650
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   651
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   652
(* outer syntax *)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   653
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   654
open OuterParse;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   655
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   656
val datatype_decl =
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   657
  Scan.option ($$$ "(" |-- name --| $$$ ")") -- type_args -- name -- opt_infix --
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   658
    ($$$ "=" |-- enum1 "|" (name -- Scan.repeat typ -- opt_mixfix));
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   659
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   660
fun mk_datatype args =
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   661
  let
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   662
    val names = map (fn ((((None, _), t), _), _) => t | ((((Some t, _), _), _), _) => t) args;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   663
    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
   664
  in #1 o add_datatype false names specs end;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   665
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   666
val datatypeP =
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   667
  OuterSyntax.parser false "datatype" "define inductive datatypes"
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   668
    (enum1 "and" datatype_decl >> (Toplevel.theory o mk_datatype));
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   669
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   670
val _ = OuterSyntax.add_keywords ["distinct", "inject", "induct"];
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   671
(* FIXME val _ = OuterSyntax.add_parsers [datatypeP, rep_datatypeP]; *)
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   672
val _ = OuterSyntax.add_parsers [datatypeP];
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   673
5177
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   674
end;
0d3a168e4d44 New datatype definition package
berghofe
parents:
diff changeset
   675
6360
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   676
structure BasicDatatypePackage: BASIC_DATATYPE_PACKAGE = DatatypePackage;
83573ae0f22c outer syntax for 'datatype';
wenzelm
parents: 6305
diff changeset
   677
open BasicDatatypePackage;