src/HOL/Tools/datatype_package.ML
author wenzelm
Sat, 06 Oct 2007 16:50:04 +0200
changeset 24867 e5b55d7be9bb
parent 24830 a7b3ab44d993
child 24959 119793c84647
permissions -rw-r--r--
simplified interfaces for outer syntax;
     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_i : bool -> bool -> string list -> (string list * bstring * mixfix *
    21     (bstring * typ 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        simps : thm list} * theory
    30   val add_datatype : bool -> string list -> (string list * bstring * mixfix *
    31     (bstring * string list * mixfix) list) list -> theory ->
    32       {distinct : thm list list,
    33        inject : thm list list,
    34        exhaustion : thm list,
    35        rec_thms : thm list,
    36        case_thms : thm list list,
    37        split_thms : (thm * thm) list,
    38        induction : thm,
    39        simps : thm list} * theory
    40   val rep_datatype_i : string list option -> (thm list * attribute list) list list ->
    41     (thm list * attribute list) list list -> (thm list * attribute list) ->
    42     theory ->
    43       {distinct : thm list list,
    44        inject : thm list list,
    45        exhaustion : thm list,
    46        rec_thms : thm list,
    47        case_thms : thm list list,
    48        split_thms : (thm * thm) list,
    49        induction : thm,
    50        simps : thm list} * theory
    51   val rep_datatype : string list option -> (thmref * Attrib.src list) list list ->
    52     (thmref * Attrib.src list) list list -> thmref * Attrib.src list -> theory ->
    53       {distinct : thm list list,
    54        inject : thm list list,
    55        exhaustion : thm list,
    56        rec_thms : thm list,
    57        case_thms : thm list list,
    58        split_thms : (thm * thm) list,
    59        induction : thm,
    60        simps : thm list} * theory
    61   val get_datatypes : theory -> DatatypeAux.datatype_info Symtab.table
    62   val get_datatype : theory -> string -> DatatypeAux.datatype_info option
    63   val the_datatype : theory -> string -> DatatypeAux.datatype_info
    64   val datatype_of_constr : theory -> string -> DatatypeAux.datatype_info option
    65   val datatype_of_case : theory -> string -> DatatypeAux.datatype_info option
    66   val get_datatype_spec : theory -> string -> ((string * sort) list * (string * typ list) list) option
    67   val get_datatype_constrs : theory -> string -> (string * typ) list option
    68   val interpretation: (string list -> theory -> theory) -> theory -> theory
    69   val print_datatypes : theory -> unit
    70   val make_case :  Proof.context -> bool -> string list -> term ->
    71     (term * term) list -> term * (term * (int * bool)) list
    72   val strip_case: Proof.context -> bool ->
    73     term -> (term * (term * term) list) option
    74   val setup: theory -> theory
    75 end;
    76 
    77 structure DatatypePackage : DATATYPE_PACKAGE =
    78 struct
    79 
    80 open DatatypeAux;
    81 
    82 val quiet_mode = quiet_mode;
    83 
    84 
    85 (* theory data *)
    86 
    87 structure DatatypesData = TheoryDataFun
    88 (
    89   type T =
    90     {types: datatype_info Symtab.table,
    91      constrs: datatype_info Symtab.table,
    92      cases: datatype_info Symtab.table};
    93 
    94   val empty =
    95     {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
    96   val copy = I;
    97   val extend = I;
    98   fun merge _
    99     ({types = types1, constrs = constrs1, cases = cases1},
   100      {types = types2, constrs = constrs2, cases = cases2}) =
   101     {types = Symtab.merge (K true) (types1, types2),
   102      constrs = Symtab.merge (K true) (constrs1, constrs2),
   103      cases = Symtab.merge (K true) (cases1, cases2)};
   104 );
   105 
   106 val get_datatypes = #types o DatatypesData.get;
   107 val map_datatypes = DatatypesData.map;
   108 
   109 fun print_datatypes thy =
   110   Pretty.writeln (Pretty.strs ("datatypes:" ::
   111     map #1 (NameSpace.extern_table (Sign.type_space thy, get_datatypes thy))));
   112 
   113 
   114 (** theory information about datatypes **)
   115 
   116 fun put_dt_infos (dt_infos : (string * datatype_info) list) =
   117   map_datatypes (fn {types, constrs, cases} =>
   118     {types = fold Symtab.update dt_infos types,
   119      constrs = fold Symtab.update
   120        (maps (fn (_, info as {descr, index, ...}) => map (rpair info o fst)
   121           (#3 (the (AList.lookup op = descr index)))) dt_infos) constrs,
   122      cases = fold Symtab.update
   123        (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)
   124        cases});
   125 
   126 val get_datatype = Symtab.lookup o get_datatypes;
   127 
   128 fun the_datatype thy name = (case get_datatype thy name of
   129       SOME info => info
   130     | NONE => error ("Unknown datatype " ^ quote name));
   131 
   132 val datatype_of_constr = Symtab.lookup o #constrs o DatatypesData.get;
   133 val datatype_of_case = Symtab.lookup o #cases o DatatypesData.get;
   134 
   135 fun get_datatype_descr thy dtco =
   136   get_datatype thy dtco
   137   |> Option.map (fn info as { descr, index, ... } =>
   138        (info, (((fn SOME (_, dtys, cos) => (dtys, cos)) o AList.lookup (op =) descr) index)));
   139 
   140 fun get_datatype_spec thy dtco =
   141   let
   142     fun mk_cons typ_of_dtyp (co, tys) =
   143       (co, map typ_of_dtyp tys);
   144     fun mk_dtyp ({ sorts = raw_sorts, descr, ... } : DatatypeAux.datatype_info, (dtys, cos)) =
   145       let
   146         val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
   147           o DatatypeAux.dest_DtTFree) dtys;
   148         val typ_of_dtyp = DatatypeAux.typ_of_dtyp descr sorts;
   149         val tys = map typ_of_dtyp dtys;
   150       in (sorts, map (mk_cons typ_of_dtyp) cos) end;
   151   in Option.map mk_dtyp (get_datatype_descr thy dtco) end;
   152 
   153 fun get_datatype_constrs thy dtco =
   154   case get_datatype_spec thy dtco
   155    of SOME (sorts, cos) =>
   156         let
   157           fun subst (v, sort) = TVar ((v, 0), sort);
   158           fun subst_ty (TFree v) = subst v
   159             | subst_ty ty = ty;
   160           val dty = Type (dtco, map subst sorts);
   161           fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
   162         in SOME (map mk_co cos) end
   163     | NONE => NONE;
   164 
   165 fun find_tname var Bi =
   166   let val frees = map dest_Free (term_frees Bi)
   167       val params = rename_wrt_term Bi (Logic.strip_params Bi);
   168   in case AList.lookup (op =) (frees @ params) var of
   169        NONE => error ("No such variable in subgoal: " ^ quote var)
   170      | SOME(Type (tn, _)) => tn
   171      | _ => error ("Cannot determine type of " ^ quote var)
   172   end;
   173 
   174 fun infer_tname state i aterm =
   175   let
   176     val sign = Thm.theory_of_thm state;
   177     val (_, _, Bi, _) = Thm.dest_state (state, i)
   178     val params = Logic.strip_params Bi;   (*params of subgoal i*)
   179     val params = rev (rename_wrt_term Bi params);   (*as they are printed*)
   180     val (types, sorts) = types_sorts state;
   181     fun types' (a, ~1) = (case AList.lookup (op =) params a of NONE => types(a, ~1) | sm => sm)
   182       | types' ixn = types ixn;
   183     val ([ct], _) = Thm.read_def_cterms (sign, types', sorts) [] false [(aterm, dummyT)];
   184   in case #T (rep_cterm ct) of
   185        Type (tn, _) => tn
   186      | _ => error ("Cannot determine type of " ^ quote aterm)
   187   end;
   188 
   189 (*Warn if the (induction) variable occurs Free among the premises, which
   190   usually signals a mistake.  But calls the tactic either way!*)
   191 fun occs_in_prems tacf vars =
   192   SUBGOAL (fn (Bi, i) =>
   193            (if exists (fn (a, _) => member (op =) vars a)
   194                       (fold Term.add_frees (#2 (strip_context Bi)) [])
   195              then warning "Induction variable occurs also among premises!"
   196              else ();
   197             tacf i));
   198 
   199 
   200 (* generic induction tactic for datatypes *)
   201 
   202 local
   203 
   204 fun prep_var (Var (ixn, _), SOME x) = SOME (ixn, x)
   205   | prep_var _ = NONE;
   206 
   207 fun prep_inst (concl, xs) = (*exception Library.UnequalLengths*)
   208   let val vs = Induct.vars_of concl
   209   in map_filter prep_var (Library.drop (length vs - length xs, vs) ~~ xs) end;
   210 
   211 in
   212 
   213 fun gen_induct_tac inst_tac (varss, opt_rule) i state =
   214   SUBGOAL (fn (Bi,_) =>
   215   let
   216     val (rule, rule_name) =
   217       case opt_rule of
   218           SOME r => (r, "Induction rule")
   219         | NONE =>
   220             let val tn = find_tname (hd (map_filter I (flat varss))) Bi
   221                 val thy = Thm.theory_of_thm state
   222             in (#induction (the_datatype thy tn), "Induction rule for type " ^ tn)
   223             end
   224     val concls = HOLogic.dest_concls (Thm.concl_of rule);
   225     val insts = maps prep_inst (concls ~~ varss) handle Library.UnequalLengths =>
   226       error (rule_name ^ " has different numbers of variables");
   227   in occs_in_prems (inst_tac insts rule) (map #2 insts) i end)
   228   i state;
   229 
   230 fun induct_tac s =
   231   gen_induct_tac Tactic.res_inst_tac'
   232     (map (single o SOME) (Syntax.read_idents s), NONE);
   233 
   234 fun induct_thm_tac th s =
   235   gen_induct_tac Tactic.res_inst_tac'
   236     ([map SOME (Syntax.read_idents s)], SOME th);
   237 
   238 end;
   239 
   240 
   241 (* generic case tactic for datatypes *)
   242 
   243 fun case_inst_tac inst_tac t rule i state =
   244   let
   245     val _ $ Var (ixn, _) $ _ = HOLogic.dest_Trueprop
   246       (hd (Logic.strip_assums_hyp (hd (Thm.prems_of rule))));
   247   in inst_tac [(ixn, t)] rule i state end;
   248 
   249 fun gen_case_tac inst_tac (t, SOME rule) i state =
   250       case_inst_tac inst_tac t rule i state
   251   | gen_case_tac inst_tac (t, NONE) i state =
   252       let val tn = infer_tname state i t in
   253         if tn = HOLogic.boolN then inst_tac [(("P", 0), t)] case_split_thm i state
   254         else case_inst_tac inst_tac t
   255                (#exhaustion (the_datatype (Thm.theory_of_thm state) tn))
   256                i state
   257       end handle THM _ => Seq.empty;
   258 
   259 fun case_tac t = gen_case_tac Tactic.res_inst_tac' (t, NONE);
   260 
   261 
   262 
   263 (** Isar tactic emulations **)
   264 
   265 local
   266 
   267 val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.$$$ ":");
   268 val opt_rule = Scan.option (rule_spec |-- Attrib.thm);
   269 
   270 val varss =
   271   Args.and_list (Scan.repeat (Scan.unless rule_spec (Scan.lift (Args.maybe Args.name))));
   272 
   273 val inst_tac = RuleInsts.bires_inst_tac false;
   274 
   275 fun induct_meth ctxt (varss, opt_rule) =
   276   gen_induct_tac (inst_tac ctxt) (varss, opt_rule);
   277 fun case_meth ctxt (varss, opt_rule) =
   278   gen_case_tac (inst_tac ctxt) (varss, opt_rule);
   279 
   280 in
   281 
   282 val tactic_emulations =
   283  [("induct_tac", Method.goal_args_ctxt' (varss -- opt_rule) induct_meth,
   284     "induct_tac emulation (dynamic instantiation)"),
   285   ("case_tac", Method.goal_args_ctxt' (Scan.lift Args.name -- opt_rule) case_meth,
   286     "case_tac emulation (dynamic instantiation)")];
   287 
   288 end;
   289 
   290 
   291 
   292 (** induct method setup **)
   293 
   294 (* case names *)
   295 
   296 local
   297 
   298 fun dt_recs (DtTFree _) = []
   299   | dt_recs (DtType (_, dts)) = maps dt_recs dts
   300   | dt_recs (DtRec i) = [i];
   301 
   302 fun dt_cases (descr: descr) (_, args, constrs) =
   303   let
   304     fun the_bname i = Sign.base_name (#1 (the (AList.lookup (op =) descr i)));
   305     val bnames = map the_bname (distinct (op =) (maps dt_recs args));
   306   in map (fn (c, _) => space_implode "_" (Sign.base_name c :: bnames)) constrs end;
   307 
   308 
   309 fun induct_cases descr =
   310   DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
   311 
   312 fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
   313 
   314 in
   315 
   316 fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
   317 
   318 fun mk_case_names_exhausts descr new =
   319   map (RuleCases.case_names o exhaust_cases descr o #1)
   320     (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
   321 
   322 end;
   323 
   324 fun add_rules simps case_thms rec_thms inject distinct
   325                   weak_case_congs cong_att =
   326   PureThy.add_thmss [(("simps", simps), []),
   327     (("", flat case_thms @
   328           flat distinct @ rec_thms), [Simplifier.simp_add]),
   329     (("", rec_thms), [RecfunCodegen.add_default]),
   330     (("", flat inject), [iff_add]),
   331     (("", map (fn th => th RS notE) (flat distinct)), [Classical.safe_elim NONE]),
   332     (("", weak_case_congs), [cong_att])]
   333   #> snd;
   334 
   335 
   336 (* add_cases_induct *)
   337 
   338 fun add_cases_induct infos induction thy =
   339   let
   340     val inducts = ProjectRule.projections (ProofContext.init thy) induction;
   341 
   342     fun named_rules (name, {index, exhaustion, ...}: datatype_info) =
   343       [(("", nth inducts index), [Induct.induct_type name]),
   344        (("", exhaustion), [Induct.cases_type name])];
   345     fun unnamed_rule i =
   346       (("", nth inducts i), [PureThy.kind_internal, Induct.induct_type ""]);
   347   in
   348     thy |> PureThy.add_thms
   349       (maps named_rules infos @
   350         map unnamed_rule (length infos upto length inducts - 1)) |> snd
   351     |> PureThy.add_thmss [(("inducts", inducts), [])] |> snd
   352   end;
   353 
   354 
   355 
   356 (**** simplification procedure for showing distinctness of constructors ****)
   357 
   358 fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
   359   | stripT p = p;
   360 
   361 fun stripC (i, f $ x) = stripC (i + 1, f)
   362   | stripC p = p;
   363 
   364 val distinctN = "constr_distinct";
   365 
   366 exception ConstrDistinct of term;
   367 
   368 fun distinct_proc thy ss (t as Const ("op =", _) $ t1 $ t2) =
   369   (case (stripC (0, t1), stripC (0, t2)) of
   370      ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
   371          (case (stripT (0, T1), stripT (0, T2)) of
   372             ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
   373                 if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
   374                    (case (get_datatype_descr thy) tname1 of
   375                       SOME (_, (_, constrs)) => let val cnames = map fst constrs
   376                         in if cname1 mem cnames andalso cname2 mem cnames then
   377                              let val eq_t = Logic.mk_equals (t, Const ("False", HOLogic.boolT));
   378                                  val eq_ct = cterm_of thy eq_t;
   379                                  val Datatype_thy = ThyInfo.the_theory "Datatype" thy;
   380                                  val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
   381                                    map (get_thm Datatype_thy o Name)
   382                                      ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"]
   383                              in (case (#distinct (the_datatype thy tname1)) of
   384                                  QuickAndDirty => SOME (Thm.invoke_oracle
   385                                    Datatype_thy distinctN (thy, ConstrDistinct eq_t))
   386                                | FewConstrs thms =>
   387                                    SOME (Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
   388                                      (EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
   389                                        atac 2, resolve_tac thms 1, etac FalseE 1])))
   390                                | ManyConstrs (thm, simpset) =>
   391                                    SOME (Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
   392                                      (EVERY [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
   393                                       full_simp_tac (Simplifier.inherit_context ss simpset) 1,
   394                                       REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
   395                                       eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
   396                                       etac FalseE 1]))))
   397                              end
   398                            else NONE
   399                         end
   400                     | NONE => NONE)
   401                 else NONE
   402           | _ => NONE)
   403    | _ => NONE)
   404   | distinct_proc _ _ _ = NONE;
   405 
   406 val distinct_simproc =
   407   Simplifier.simproc HOL.thy distinctN ["s = t"] distinct_proc;
   408 
   409 val dist_ss = HOL_ss addsimprocs [distinct_simproc];
   410 
   411 val simproc_setup =
   412   Theory.add_oracle (distinctN, fn (_, ConstrDistinct t) => t) #>
   413   (fn thy => ((change_simpset_of thy) (fn ss => ss addsimprocs [distinct_simproc]); thy));
   414 
   415 
   416 (**** translation rules for case ****)
   417 
   418 fun make_case ctxt = DatatypeCase.make_case
   419   (datatype_of_constr (ProofContext.theory_of ctxt)) ctxt;
   420 
   421 fun strip_case ctxt = DatatypeCase.strip_case
   422   (datatype_of_case (ProofContext.theory_of ctxt));
   423 
   424 fun add_case_tr' case_names thy =
   425   Sign.add_advanced_trfuns ([], [],
   426     map (fn case_name =>
   427       let val case_name' = Sign.const_syntax_name thy case_name
   428       in (case_name', DatatypeCase.case_tr' datatype_of_case case_name')
   429       end) case_names, []) thy;
   430 
   431 val trfun_setup =
   432   Sign.add_advanced_trfuns ([],
   433     [("_case_syntax", DatatypeCase.case_tr true datatype_of_constr)],
   434     [], []);
   435 
   436 
   437 (* prepare types *)
   438 
   439 fun read_typ sign ((Ts, sorts), str) =
   440   let
   441     val T = Type.no_tvars (Sign.read_def_typ (sign, AList.lookup (op =)
   442       (map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg
   443   in (Ts @ [T], add_typ_tfrees (T, sorts)) end;
   444 
   445 fun cert_typ sign ((Ts, sorts), raw_T) =
   446   let
   447     val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
   448       TYPE (msg, _, _) => error msg;
   449     val sorts' = add_typ_tfrees (T, sorts)
   450   in (Ts @ [T],
   451       case duplicates (op =) (map fst sorts') of
   452          [] => sorts'
   453        | dups => error ("Inconsistent sort constraints for " ^ commas dups))
   454   end;
   455 
   456 
   457 (**** make datatype info ****)
   458 
   459 fun make_dt_info head_len descr sorts induct reccomb_names rec_thms
   460     (((((((((i, (_, (tname, _, _))), case_name), case_thms),
   461       exhaustion_thm), distinct_thm), inject), nchotomy), case_cong), weak_case_cong) =
   462   (tname,
   463    {index = i,
   464     head_len = head_len,
   465     descr = descr,
   466     sorts = sorts,
   467     rec_names = reccomb_names,
   468     rec_rewrites = rec_thms,
   469     case_name = case_name,
   470     case_rewrites = case_thms,
   471     induction = induct,
   472     exhaustion = exhaustion_thm,
   473     distinct = distinct_thm,
   474     inject = inject,
   475     nchotomy = nchotomy,
   476     case_cong = case_cong,
   477     weak_case_cong = weak_case_cong});
   478 
   479 
   480 (********************* axiomatic introduction of datatypes ********************)
   481 
   482 fun add_axiom label t atts thy =
   483   thy
   484   |> PureThy.add_axioms_i [((label, t), atts)];
   485 
   486 fun add_axioms label ts atts thy =
   487   thy
   488   |> PureThy.add_axiomss_i [((label, ts), atts)];
   489 
   490 fun add_and_get_axioms_atts label tnames ts attss =
   491   fold_map (fn (tname, (atts, t)) => fn thy =>
   492     thy
   493     |> Sign.add_path tname
   494     |> add_axiom label t atts
   495     ||> Sign.parent_path
   496     |-> (fn [ax] => pair ax)) (tnames ~~ (attss ~~ ts));
   497 
   498 fun add_and_get_axioms label tnames ts =
   499   add_and_get_axioms_atts label tnames ts (replicate (length tnames) []);
   500 
   501 fun add_and_get_axiomss label tnames tss =
   502   fold_map (fn (tname, ts) => fn thy =>
   503     thy
   504     |> Sign.add_path tname
   505     |> add_axioms label ts []
   506     ||> Sign.parent_path
   507     |-> (fn [ax] => pair ax)) (tnames ~~ tss);
   508 
   509 fun gen_specify_consts add args thy =
   510   let
   511     val specs = map (fn (c, T, mx) =>
   512       Const (Sign.full_name thy (Syntax.const_name c mx), T)) args;
   513   in
   514     thy
   515     |> add args
   516     |> Theory.add_finals_i false specs
   517   end;
   518 
   519 val specify_consts = gen_specify_consts Sign.add_consts_i;
   520 val specify_consts_authentic = gen_specify_consts (Sign.add_consts_authentic []);
   521 
   522 structure DatatypeInterpretation = InterpretationFun(type T = string list val eq = op =);
   523 val interpretation = DatatypeInterpretation.interpretation;
   524 
   525 fun add_datatype_axm flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
   526     case_names_induct case_names_exhausts thy =
   527   let
   528     val descr' = flat descr;
   529     val recTs = get_rec_types descr' sorts;
   530     val used = map fst (fold Term.add_tfreesT recTs []);
   531     val newTs = Library.take (length (hd descr), recTs);
   532 
   533     (**** declare new types and constants ****)
   534 
   535     val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
   536 
   537     val constr_decls = map (fn (((_, (_, _, constrs)), T), constr_syntax') =>
   538       map (fn ((_, cargs), (cname, mx)) =>
   539         (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
   540           (constrs ~~ constr_syntax')) ((hd descr) ~~ newTs ~~ constr_syntax);
   541 
   542     val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used;
   543 
   544     val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
   545     val reccomb_names = if length descr' = 1 then [big_reccomb_name] else
   546       (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
   547         (1 upto (length descr')));
   548 
   549     val freeT = TFree (Name.variant used "'t", HOLogic.typeS);
   550     val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
   551       map (fn (_, cargs) =>
   552         let val Ts = map (typ_of_dtyp descr' sorts) cargs
   553         in Ts ---> freeT end) constrs) (hd descr);
   554 
   555     val case_names = map (fn s => (s ^ "_case")) new_type_names;
   556 
   557     val thy2' = thy
   558 
   559       (** new types **)
   560       |> fold2 (fn (name, mx) => fn tvs => TypedefPackage.add_typedecls [(name, tvs, mx)])
   561            types_syntax tyvars
   562       |> add_path flat_names (space_implode "_" new_type_names)
   563 
   564       (** primrec combinators **)
   565 
   566       |> specify_consts (map (fn ((name, T), T') =>
   567            (name, reccomb_fn_Ts @ [T] ---> T', NoSyn)) (reccomb_names ~~ recTs ~~ rec_result_Ts))
   568 
   569       (** case combinators **)
   570 
   571       |> specify_consts_authentic (map (fn ((name, T), Ts) =>
   572            (name, Ts @ [T] ---> freeT, NoSyn)) (case_names ~~ newTs ~~ case_fn_Ts));
   573 
   574     val reccomb_names' = map (Sign.full_name thy2') reccomb_names;
   575     val case_names' = map (Sign.full_name thy2') case_names;
   576 
   577     val thy2 = thy2'
   578 
   579       (** constructors **)
   580 
   581       |> parent_path flat_names
   582       |> fold (fn ((((_, (_, _, constrs)), T), tname),
   583         constr_syntax') =>
   584           add_path flat_names tname #>
   585             specify_consts (map (fn ((_, cargs), (cname, mx)) =>
   586               (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
   587                 (constrs ~~ constr_syntax')) #>
   588           parent_path flat_names)
   589             (hd descr ~~ newTs ~~ new_type_names ~~ constr_syntax);
   590 
   591     (**** introduction of axioms ****)
   592 
   593     val rec_axs = DatatypeProp.make_primrecs new_type_names descr sorts thy2;
   594 
   595     val ((([induct], [rec_thms]), inject), thy3) =
   596       thy2
   597       |> Sign.add_path (space_implode "_" new_type_names)
   598       |> add_axiom "induct" (DatatypeProp.make_ind descr sorts) [case_names_induct]
   599       ||>> add_axioms "recs" rec_axs []
   600       ||> Sign.parent_path
   601       ||>> add_and_get_axiomss "inject" new_type_names
   602             (DatatypeProp.make_injs descr sorts);
   603     val (distinct, thy4) = add_and_get_axiomss "distinct" new_type_names
   604       (DatatypeProp.make_distincts new_type_names descr sorts thy3) thy3;
   605 
   606     val exhaust_ts = DatatypeProp.make_casedists descr sorts;
   607     val (exhaustion, thy5) = add_and_get_axioms_atts "exhaust" new_type_names
   608       exhaust_ts (map single case_names_exhausts) thy4;
   609     val (case_thms, thy6) = add_and_get_axiomss "cases" new_type_names
   610       (DatatypeProp.make_cases new_type_names descr sorts thy5) thy5;
   611     val (split_ts, split_asm_ts) = ListPair.unzip
   612       (DatatypeProp.make_splits new_type_names descr sorts thy6);
   613     val (split, thy7) = add_and_get_axioms "split" new_type_names split_ts thy6;
   614     val (split_asm, thy8) = add_and_get_axioms "split_asm" new_type_names
   615       split_asm_ts thy7;
   616     val (nchotomys, thy9) = add_and_get_axioms "nchotomy" new_type_names
   617       (DatatypeProp.make_nchotomys descr sorts) thy8;
   618     val (case_congs, thy10) = add_and_get_axioms "case_cong" new_type_names
   619       (DatatypeProp.make_case_congs new_type_names descr sorts thy9) thy9;
   620     val (weak_case_congs, thy11) = add_and_get_axioms "weak_case_cong" new_type_names
   621       (DatatypeProp.make_weak_case_congs new_type_names descr sorts thy10) thy10;
   622 
   623     val dt_infos = map (make_dt_info (length (hd descr)) descr' sorts induct
   624         reccomb_names' rec_thms)
   625       ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names' ~~ case_thms ~~
   626         exhaustion ~~ replicate (length (hd descr)) QuickAndDirty ~~ inject ~~
   627           nchotomys ~~ case_congs ~~ weak_case_congs);
   628 
   629     val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
   630     val split_thms = split ~~ split_asm;
   631 
   632     val thy12 =
   633       thy11
   634       |> add_case_tr' case_names'
   635       |> Sign.add_path (space_implode "_" new_type_names)
   636       |> add_rules simps case_thms rec_thms inject distinct
   637           weak_case_congs Simplifier.cong_add
   638       |> put_dt_infos dt_infos
   639       |> add_cases_induct dt_infos induct
   640       |> Sign.parent_path
   641       |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)
   642       |> snd
   643       |> DatatypeInterpretation.data (map fst dt_infos);
   644   in
   645     ({distinct = distinct,
   646       inject = inject,
   647       exhaustion = exhaustion,
   648       rec_thms = rec_thms,
   649       case_thms = case_thms,
   650       split_thms = split_thms,
   651       induction = induct,
   652       simps = simps}, thy12)
   653   end;
   654 
   655 
   656 (******************* definitional introduction of datatypes *******************)
   657 
   658 fun add_datatype_def flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
   659     case_names_induct case_names_exhausts thy =
   660   let
   661     val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
   662 
   663     val ((inject, distinct, dist_rewrites, simproc_dists, induct), thy2) = thy |>
   664       DatatypeRepProofs.representation_proofs flat_names dt_info new_type_names descr sorts
   665         types_syntax constr_syntax case_names_induct;
   666 
   667     val (casedist_thms, thy3) = DatatypeAbsProofs.prove_casedist_thms new_type_names descr
   668       sorts induct case_names_exhausts thy2;
   669     val ((reccomb_names, rec_thms), thy4) = DatatypeAbsProofs.prove_primrec_thms
   670       flat_names new_type_names descr sorts dt_info inject dist_rewrites
   671       (Simplifier.theory_context thy3 dist_ss) induct thy3;
   672     val ((case_thms, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
   673       flat_names new_type_names descr sorts reccomb_names rec_thms thy4;
   674     val (split_thms, thy7) = DatatypeAbsProofs.prove_split_thms new_type_names
   675       descr sorts inject dist_rewrites casedist_thms case_thms thy6;
   676     val (nchotomys, thy8) = DatatypeAbsProofs.prove_nchotomys new_type_names
   677       descr sorts casedist_thms thy7;
   678     val (case_congs, thy9) = DatatypeAbsProofs.prove_case_congs new_type_names
   679       descr sorts nchotomys case_thms thy8;
   680     val (weak_case_congs, thy10) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
   681       descr sorts thy9;
   682 
   683     val dt_infos = map (make_dt_info (length (hd descr)) (flat descr) sorts induct
   684         reccomb_names rec_thms)
   685       ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
   686         casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
   687 
   688     val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
   689 
   690     val thy12 =
   691       thy10
   692       |> add_case_tr' case_names
   693       |> Sign.add_path (space_implode "_" new_type_names)
   694       |> add_rules simps case_thms rec_thms inject distinct
   695           weak_case_congs (Simplifier.attrib (op addcongs))
   696       |> put_dt_infos dt_infos
   697       |> add_cases_induct dt_infos induct
   698       |> Sign.parent_path
   699       |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms) |> snd
   700       |> DatatypeInterpretation.data (map fst dt_infos);
   701   in
   702     ({distinct = distinct,
   703       inject = inject,
   704       exhaustion = casedist_thms,
   705       rec_thms = rec_thms,
   706       case_thms = case_thms,
   707       split_thms = split_thms,
   708       induction = induct,
   709       simps = simps}, thy12)
   710   end;
   711 
   712 
   713 (*********************** declare existing type as datatype *********************)
   714 
   715 fun gen_rep_datatype apply_theorems alt_names raw_distinct raw_inject raw_induction thy0 =
   716   let
   717     val (((distinct, inject), [induction]), thy1) =
   718       thy0
   719       |> fold_map apply_theorems raw_distinct
   720       ||>> fold_map apply_theorems raw_inject
   721       ||>> apply_theorems [raw_induction];
   722 
   723     val ((_, [induction']), _) =
   724       Variable.importT_thms [induction] (Variable.thm_context induction);
   725 
   726     fun err t = error ("Ill-formed predicate in induction rule: " ^
   727       Sign.string_of_term thy1 t);
   728 
   729     fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
   730           ((tname, map dest_TFree Ts) handle TERM _ => err t)
   731       | get_typ t = err t;
   732 
   733     val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induction')));
   734     val new_type_names = getOpt (alt_names, map fst dtnames);
   735 
   736     fun get_constr t = (case Logic.strip_assums_concl t of
   737         _ $ (_ $ t') => (case head_of t' of
   738             Const (cname, cT) => (case strip_type cT of
   739                 (Ts, Type (tname, _)) => (tname, (cname, map (dtyp_of_typ dtnames) Ts))
   740               | _ => err t)
   741           | _ => err t)
   742       | _ => err t);
   743 
   744     fun make_dt_spec [] _ _ = []
   745       | make_dt_spec ((tname, tvs)::dtnames') i constrs =
   746           let val (constrs', constrs'') = take_prefix (equal tname o fst) constrs
   747           in (i, (tname, map DtTFree tvs, map snd constrs'))::
   748             (make_dt_spec dtnames' (i + 1) constrs'')
   749           end;
   750 
   751     val descr = make_dt_spec dtnames 0 (map get_constr (prems_of induction'));
   752     val sorts = add_term_tfrees (concl_of induction', []);
   753     val dt_info = get_datatypes thy1;
   754 
   755     val (case_names_induct, case_names_exhausts) =
   756       (mk_case_names_induct descr, mk_case_names_exhausts descr (map #1 dtnames));
   757 
   758     val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
   759 
   760     val (casedist_thms, thy2) = thy1 |>
   761       DatatypeAbsProofs.prove_casedist_thms new_type_names [descr] sorts induction
   762         case_names_exhausts;
   763     val ((reccomb_names, rec_thms), thy3) = DatatypeAbsProofs.prove_primrec_thms
   764       false new_type_names [descr] sorts dt_info inject distinct
   765       (Simplifier.theory_context thy2 dist_ss) induction thy2;
   766     val ((case_thms, case_names), thy4) = DatatypeAbsProofs.prove_case_thms false
   767       new_type_names [descr] sorts reccomb_names rec_thms thy3;
   768     val (split_thms, thy5) = DatatypeAbsProofs.prove_split_thms
   769       new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
   770     val (nchotomys, thy6) = DatatypeAbsProofs.prove_nchotomys new_type_names
   771       [descr] sorts casedist_thms thy5;
   772     val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
   773       [descr] sorts nchotomys case_thms thy6;
   774     val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
   775       [descr] sorts thy7;
   776 
   777     val ((_, [induction']), thy10) =
   778       thy8
   779       |> store_thmss "inject" new_type_names inject
   780       ||>> store_thmss "distinct" new_type_names distinct
   781       ||> Sign.add_path (space_implode "_" new_type_names)
   782       ||>> PureThy.add_thms [(("induct", induction), [case_names_induct])];
   783 
   784     val dt_infos = map (make_dt_info (length descr) descr sorts induction'
   785         reccomb_names rec_thms)
   786       ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
   787         map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
   788 
   789     val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
   790 
   791     val thy11 =
   792       thy10
   793       |> add_case_tr' case_names
   794       |> add_rules simps case_thms rec_thms inject distinct
   795            weak_case_congs (Simplifier.attrib (op addcongs))
   796       |> put_dt_infos dt_infos
   797       |> add_cases_induct dt_infos induction'
   798       |> Sign.parent_path
   799       |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)
   800       |> snd
   801       |> DatatypeInterpretation.data (map fst dt_infos);
   802   in
   803     ({distinct = distinct,
   804       inject = inject,
   805       exhaustion = casedist_thms,
   806       rec_thms = rec_thms,
   807       case_thms = case_thms,
   808       split_thms = split_thms,
   809       induction = induction',
   810       simps = simps}, thy11)
   811   end;
   812 
   813 val rep_datatype = gen_rep_datatype IsarCmd.apply_theorems;
   814 val rep_datatype_i = gen_rep_datatype IsarCmd.apply_theorems_i;
   815 
   816 
   817 
   818 (******************************** add datatype ********************************)
   819 
   820 fun gen_add_datatype prep_typ err flat_names new_type_names dts thy =
   821   let
   822     val _ = Theory.requires thy "Datatype" "datatype definitions";
   823 
   824     (* this theory is used just for parsing *)
   825 
   826     val tmp_thy = thy |>
   827       Theory.copy |>
   828       Sign.add_types (map (fn (tvs, tname, mx, _) =>
   829         (tname, length tvs, mx)) dts);
   830 
   831     val (tyvars, _, _, _)::_ = dts;
   832     val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
   833       let val full_tname = Sign.full_name tmp_thy (Syntax.type_name tname mx)
   834       in (case duplicates (op =) tvs of
   835             [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
   836                   else error ("Mutually recursive datatypes must have same type parameters")
   837           | dups => error ("Duplicate parameter(s) for datatype " ^ full_tname ^
   838               " : " ^ commas dups))
   839       end) dts);
   840 
   841     val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
   842       [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
   843 
   844     fun prep_dt_spec (tvs, tname, mx, constrs) (dts', constr_syntax, sorts, i) =
   845       let
   846         fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
   847           let
   848             val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
   849             val _ = (case fold (curry add_typ_tfree_names) cargs' [] \\ tvs of
   850                 [] => ()
   851               | vs => error ("Extra type variables on rhs: " ^ commas vs))
   852           in (constrs @ [((if flat_names then Sign.full_name tmp_thy else
   853                 Sign.full_name_path tmp_thy tname) (Syntax.const_name cname mx'),
   854                    map (dtyp_of_typ new_dts) cargs')],
   855               constr_syntax' @ [(cname, mx')], sorts'')
   856           end handle ERROR msg =>
   857             cat_error msg ("The error above occured in constructor " ^ cname ^
   858               " of datatype " ^ tname);
   859 
   860         val (constrs', constr_syntax', sorts') =
   861           fold prep_constr constrs ([], [], sorts)
   862 
   863       in
   864         case duplicates (op =) (map fst constrs') of
   865            [] =>
   866              (dts' @ [(i, (Sign.full_name tmp_thy (Syntax.type_name tname mx),
   867                 map DtTFree tvs, constrs'))],
   868               constr_syntax @ [constr_syntax'], sorts', i + 1)
   869          | dups => error ("Duplicate constructors " ^ commas dups ^
   870              " in datatype " ^ tname)
   871       end;
   872 
   873     val (dts', constr_syntax, sorts', i) = fold prep_dt_spec dts ([], [], [], 0);
   874     val sorts = sorts' @ (map (rpair (Sign.defaultS tmp_thy)) (tyvars \\ map fst sorts'));
   875     val dt_info = get_datatypes thy;
   876     val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
   877     val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
   878       if err then error ("Nonemptiness check failed for datatype " ^ s)
   879       else raise exn;
   880 
   881     val descr' = flat descr;
   882     val case_names_induct = mk_case_names_induct descr';
   883     val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
   884   in
   885     (if (!quick_and_dirty) then add_datatype_axm else add_datatype_def)
   886       flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
   887       case_names_induct case_names_exhausts thy
   888   end;
   889 
   890 val add_datatype_i = gen_add_datatype cert_typ;
   891 val add_datatype = gen_add_datatype read_typ true;
   892 
   893 
   894 
   895 (** package setup **)
   896 
   897 (* setup theory *)
   898 
   899 val setup =
   900   DatatypeProp.distinctness_limit_setup #>
   901   Method.add_methods tactic_emulations #>
   902   simproc_setup #>
   903   trfun_setup #>
   904   DatatypeInterpretation.init;
   905 
   906 
   907 (* outer syntax *)
   908 
   909 local structure P = OuterParse and K = OuterKeyword in
   910 
   911 val _ = OuterSyntax.keywords ["distinct", "inject", "induction"];
   912 
   913 val datatype_decl =
   914   Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
   915     (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix));
   916 
   917 fun mk_datatype args =
   918   let
   919     val names = map (fn ((((NONE, _), t), _), _) => t | ((((SOME t, _), _), _), _) => t) args;
   920     val specs = map (fn ((((_, vs), t), mx), cons) =>
   921       (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
   922   in snd o add_datatype false names specs end;
   923 
   924 val _ =
   925   OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
   926     (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
   927 
   928 
   929 val rep_datatype_decl =
   930   Scan.option (Scan.repeat1 P.name) --
   931     Scan.optional (P.$$$ "distinct" |-- P.!!! (P.and_list1 SpecParse.xthms1)) [[]] --
   932     Scan.optional (P.$$$ "inject" |-- P.!!! (P.and_list1 SpecParse.xthms1)) [[]] --
   933     (P.$$$ "induction" |-- P.!!! SpecParse.xthm);
   934 
   935 fun mk_rep_datatype (((opt_ts, dss), iss), ind) = #2 o rep_datatype opt_ts dss iss ind;
   936 
   937 val _ =
   938   OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_decl
   939     (rep_datatype_decl >> (Toplevel.theory o mk_rep_datatype));
   940 
   941 end;
   942 
   943 
   944 end;
   945 
   946 structure BasicDatatypePackage: BASIC_DATATYPE_PACKAGE = DatatypePackage;
   947 open BasicDatatypePackage;
   948