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