src/HOL/Tools/datatype_package.ML
changeset 31614 ef6d67b1ad10
parent 31593 dc65b2f78664
parent 31613 78ac5c304db7
child 31615 0b807e04f1f8
child 31626 fe35b72b9ef0
equal deleted inserted replaced
31593:dc65b2f78664 31614:ef6d67b1ad10
     1 (*  Title:      HOL/Tools/datatype_package.ML
       
     2     Author:     Stefan Berghofer, TU Muenchen
       
     3 
       
     4 Datatype package for Isabelle/HOL.
       
     5 *)
       
     6 
       
     7 signature DATATYPE_PACKAGE =
       
     8 sig
       
     9   val quiet_mode : bool ref
       
    10   val get_datatypes : theory -> DatatypeAux.datatype_info Symtab.table
       
    11   val print_datatypes : theory -> unit
       
    12   val get_datatype : theory -> string -> DatatypeAux.datatype_info option
       
    13   val the_datatype : theory -> string -> DatatypeAux.datatype_info
       
    14   val datatype_of_constr : theory -> string -> DatatypeAux.datatype_info option
       
    15   val datatype_of_case : theory -> string -> DatatypeAux.datatype_info option
       
    16   val the_datatype_spec : theory -> string -> (string * sort) list * (string * typ list) list
       
    17   val get_datatype_constrs : theory -> string -> (string * typ) list option
       
    18   val construction_interpretation : theory
       
    19     -> {atom : typ -> 'a, dtyp : string -> 'a, rtyp : string * typ list -> 'a list -> 'a}
       
    20     -> (string * sort) list -> string list
       
    21     -> (string * (string * 'a list) list) list
       
    22       * ((string * typ list) * (string * 'a list) list) list
       
    23   val distinct_simproc : simproc
       
    24   val make_case :  Proof.context -> bool -> string list -> term ->
       
    25     (term * term) list -> term * (term * (int * bool)) list
       
    26   val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
       
    27   val read_typ: theory ->
       
    28     (typ list * (string * sort) list) * string -> typ list * (string * sort) list
       
    29   val interpretation : (string list -> theory -> theory) -> theory -> theory
       
    30   val rep_datatype : ({distinct : thm list list,
       
    31        inject : thm list list,
       
    32        exhaustion : thm list,
       
    33        rec_thms : thm list,
       
    34        case_thms : thm list list,
       
    35        split_thms : (thm * thm) list,
       
    36        induction : thm,
       
    37        simps : thm list} -> Proof.context -> Proof.context) -> string list option -> term list
       
    38     -> theory -> Proof.state;
       
    39   val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state;
       
    40   val add_datatype : bool -> bool -> string list -> (string list * binding * mixfix *
       
    41     (binding * typ list * mixfix) list) list -> theory ->
       
    42       {distinct : thm list list,
       
    43        inject : thm list list,
       
    44        exhaustion : thm list,
       
    45        rec_thms : thm list,
       
    46        case_thms : thm list list,
       
    47        split_thms : (thm * thm) list,
       
    48        induction : thm,
       
    49        simps : thm list} * theory
       
    50   val add_datatype_cmd : bool -> string list -> (string list * binding * mixfix *
       
    51     (binding * string list * mixfix) list) list -> theory ->
       
    52       {distinct : thm list list,
       
    53        inject : thm list list,
       
    54        exhaustion : thm list,
       
    55        rec_thms : thm list,
       
    56        case_thms : thm list list,
       
    57        split_thms : (thm * thm) list,
       
    58        induction : thm,
       
    59        simps : thm list} * theory
       
    60   val setup: theory -> theory
       
    61 end;
       
    62 
       
    63 structure DatatypePackage : DATATYPE_PACKAGE =
       
    64 struct
       
    65 
       
    66 open DatatypeAux;
       
    67 
       
    68 val quiet_mode = quiet_mode;
       
    69 
       
    70 
       
    71 (* theory data *)
       
    72 
       
    73 structure DatatypesData = TheoryDataFun
       
    74 (
       
    75   type T =
       
    76     {types: datatype_info Symtab.table,
       
    77      constrs: datatype_info Symtab.table,
       
    78      cases: datatype_info Symtab.table};
       
    79 
       
    80   val empty =
       
    81     {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
       
    82   val copy = I;
       
    83   val extend = I;
       
    84   fun merge _
       
    85     ({types = types1, constrs = constrs1, cases = cases1},
       
    86      {types = types2, constrs = constrs2, cases = cases2}) =
       
    87     {types = Symtab.merge (K true) (types1, types2),
       
    88      constrs = Symtab.merge (K true) (constrs1, constrs2),
       
    89      cases = Symtab.merge (K true) (cases1, cases2)};
       
    90 );
       
    91 
       
    92 val get_datatypes = #types o DatatypesData.get;
       
    93 val map_datatypes = DatatypesData.map;
       
    94 
       
    95 fun print_datatypes thy =
       
    96   Pretty.writeln (Pretty.strs ("datatypes:" ::
       
    97     map #1 (NameSpace.extern_table (Sign.type_space thy, get_datatypes thy))));
       
    98 
       
    99 
       
   100 (** theory information about datatypes **)
       
   101 
       
   102 fun put_dt_infos (dt_infos : (string * datatype_info) list) =
       
   103   map_datatypes (fn {types, constrs, cases} =>
       
   104     {types = fold Symtab.update dt_infos types,
       
   105      constrs = fold Symtab.default (*conservative wrt. overloaded constructors*)
       
   106        (maps (fn (_, info as {descr, index, ...}) => map (rpair info o fst)
       
   107           (#3 (the (AList.lookup op = descr index)))) dt_infos) constrs,
       
   108      cases = fold Symtab.update
       
   109        (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)
       
   110        cases});
       
   111 
       
   112 val get_datatype = Symtab.lookup o get_datatypes;
       
   113 
       
   114 fun the_datatype thy name = (case get_datatype thy name of
       
   115       SOME info => info
       
   116     | NONE => error ("Unknown datatype " ^ quote name));
       
   117 
       
   118 val datatype_of_constr = Symtab.lookup o #constrs o DatatypesData.get;
       
   119 val datatype_of_case = Symtab.lookup o #cases o DatatypesData.get;
       
   120 
       
   121 fun get_datatype_descr thy dtco =
       
   122   get_datatype thy dtco
       
   123   |> Option.map (fn info as { descr, index, ... } =>
       
   124        (info, (((fn SOME (_, dtys, cos) => (dtys, cos)) o AList.lookup (op =) descr) index)));
       
   125 
       
   126 fun the_datatype_spec thy dtco =
       
   127   let
       
   128     val info as { descr, index, sorts = raw_sorts, ... } = the_datatype thy dtco;
       
   129     val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index;
       
   130     val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
       
   131       o DatatypeAux.dest_DtTFree) dtys;
       
   132     val cos = map
       
   133       (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos;
       
   134   in (sorts, cos) end;
       
   135 
       
   136 fun get_datatype_constrs thy dtco =
       
   137   case try (the_datatype_spec thy) dtco
       
   138    of SOME (sorts, cos) =>
       
   139         let
       
   140           fun subst (v, sort) = TVar ((v, 0), sort);
       
   141           fun subst_ty (TFree v) = subst v
       
   142             | subst_ty ty = ty;
       
   143           val dty = Type (dtco, map subst sorts);
       
   144           fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
       
   145         in SOME (map mk_co cos) end
       
   146     | NONE => NONE;
       
   147 
       
   148 fun construction_interpretation thy { atom, dtyp, rtyp } sorts tycos =
       
   149   let
       
   150     val descr = (#descr o the_datatype thy o hd) tycos;
       
   151     val k = length tycos;
       
   152     val descr_of = the o AList.lookup (op =) descr;
       
   153     val typ_of_dtyp = typ_of_dtyp descr sorts;
       
   154     fun interpT (dT as DtTFree _) = atom (typ_of_dtyp dT)
       
   155       | interpT (dT as DtType (tyco, dTs)) = 
       
   156           let
       
   157             val Ts = map typ_of_dtyp dTs;
       
   158           in if is_rec_type dT
       
   159             then rtyp (tyco, Ts) (map interpT dTs)
       
   160             else atom (Type (tyco, Ts))
       
   161           end
       
   162       | interpT (DtRec l) = if l < k then (dtyp o #1 o descr_of) l
       
   163           else let
       
   164             val (tyco, dTs, _) = descr_of l;
       
   165             val Ts = map typ_of_dtyp dTs;
       
   166           in rtyp (tyco, Ts) (map interpT dTs) end;
       
   167     fun interpC (c, dTs) = (c, map interpT dTs);
       
   168     fun interpD (_, (tyco, _, cs)) = (tyco, map interpC cs);
       
   169     fun interpR (_, (tyco, dTs, cs)) = ((tyco, map typ_of_dtyp dTs), map interpC cs);
       
   170   in chop k descr |> (apfst o map) interpD |> (apsnd o map) interpR end;
       
   171 
       
   172 
       
   173 
       
   174 (** induct method setup **)
       
   175 
       
   176 (* case names *)
       
   177 
       
   178 local
       
   179 
       
   180 fun dt_recs (DtTFree _) = []
       
   181   | dt_recs (DtType (_, dts)) = maps dt_recs dts
       
   182   | dt_recs (DtRec i) = [i];
       
   183 
       
   184 fun dt_cases (descr: descr) (_, args, constrs) =
       
   185   let
       
   186     fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
       
   187     val bnames = map the_bname (distinct (op =) (maps dt_recs args));
       
   188   in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
       
   189 
       
   190 
       
   191 fun induct_cases descr =
       
   192   DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
       
   193 
       
   194 fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
       
   195 
       
   196 in
       
   197 
       
   198 fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
       
   199 
       
   200 fun mk_case_names_exhausts descr new =
       
   201   map (RuleCases.case_names o exhaust_cases descr o #1)
       
   202     (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
       
   203 
       
   204 end;
       
   205 
       
   206 fun add_rules simps case_thms rec_thms inject distinct
       
   207                   weak_case_congs cong_att =
       
   208   PureThy.add_thmss [((Binding.name "simps", simps), []),
       
   209     ((Binding.empty, flat case_thms @
       
   210           flat distinct @ rec_thms), [Simplifier.simp_add]),
       
   211     ((Binding.empty, rec_thms), [Code.add_default_eqn_attribute]),
       
   212     ((Binding.empty, flat inject), [iff_add]),
       
   213     ((Binding.empty, map (fn th => th RS notE) (flat distinct)), [Classical.safe_elim NONE]),
       
   214     ((Binding.empty, weak_case_congs), [cong_att])]
       
   215   #> snd;
       
   216 
       
   217 
       
   218 (* add_cases_induct *)
       
   219 
       
   220 fun add_cases_induct infos induction thy =
       
   221   let
       
   222     val inducts = ProjectRule.projections (ProofContext.init thy) induction;
       
   223 
       
   224     fun named_rules (name, {index, exhaustion, ...}: datatype_info) =
       
   225       [((Binding.empty, nth inducts index), [Induct.induct_type name]),
       
   226        ((Binding.empty, exhaustion), [Induct.cases_type name])];
       
   227     fun unnamed_rule i =
       
   228       ((Binding.empty, nth inducts i), [Thm.kind_internal, Induct.induct_type ""]);
       
   229   in
       
   230     thy |> PureThy.add_thms
       
   231       (maps named_rules infos @
       
   232         map unnamed_rule (length infos upto length inducts - 1)) |> snd
       
   233     |> PureThy.add_thmss [((Binding.name "inducts", inducts), [])] |> snd
       
   234   end;
       
   235 
       
   236 
       
   237 
       
   238 (**** simplification procedure for showing distinctness of constructors ****)
       
   239 
       
   240 fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
       
   241   | stripT p = p;
       
   242 
       
   243 fun stripC (i, f $ x) = stripC (i + 1, f)
       
   244   | stripC p = p;
       
   245 
       
   246 val distinctN = "constr_distinct";
       
   247 
       
   248 fun distinct_rule thy ss tname eq_t = case #distinct (the_datatype thy tname) of
       
   249     FewConstrs thms => Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
       
   250       (EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
       
   251         atac 2, resolve_tac thms 1, etac FalseE 1]))
       
   252   | ManyConstrs (thm, simpset) =>
       
   253       let
       
   254         val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
       
   255           map (PureThy.get_thm (ThyInfo.the_theory "Datatype" thy))
       
   256             ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"];
       
   257       in
       
   258         Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
       
   259         (EVERY [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
       
   260           full_simp_tac (Simplifier.inherit_context ss simpset) 1,
       
   261           REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
       
   262           eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
       
   263           etac FalseE 1]))
       
   264       end;
       
   265 
       
   266 fun distinct_proc thy ss (t as Const ("op =", _) $ t1 $ t2) =
       
   267   (case (stripC (0, t1), stripC (0, t2)) of
       
   268      ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
       
   269          (case (stripT (0, T1), stripT (0, T2)) of
       
   270             ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
       
   271                 if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
       
   272                    (case (get_datatype_descr thy) tname1 of
       
   273                       SOME (_, (_, constrs)) => let val cnames = map fst constrs
       
   274                         in if cname1 mem cnames andalso cname2 mem cnames then
       
   275                              SOME (distinct_rule thy ss tname1
       
   276                                (Logic.mk_equals (t, Const ("False", HOLogic.boolT))))
       
   277                            else NONE
       
   278                         end
       
   279                     | NONE => NONE)
       
   280                 else NONE
       
   281           | _ => NONE)
       
   282    | _ => NONE)
       
   283   | distinct_proc _ _ _ = NONE;
       
   284 
       
   285 val distinct_simproc =
       
   286   Simplifier.simproc @{theory HOL} distinctN ["s = t"] distinct_proc;
       
   287 
       
   288 val dist_ss = HOL_ss addsimprocs [distinct_simproc];
       
   289 
       
   290 val simproc_setup =
       
   291   Simplifier.map_simpset (fn ss => ss addsimprocs [distinct_simproc]);
       
   292 
       
   293 
       
   294 (**** translation rules for case ****)
       
   295 
       
   296 fun make_case ctxt = DatatypeCase.make_case
       
   297   (datatype_of_constr (ProofContext.theory_of ctxt)) ctxt;
       
   298 
       
   299 fun strip_case ctxt = DatatypeCase.strip_case
       
   300   (datatype_of_case (ProofContext.theory_of ctxt));
       
   301 
       
   302 fun add_case_tr' case_names thy =
       
   303   Sign.add_advanced_trfuns ([], [],
       
   304     map (fn case_name =>
       
   305       let val case_name' = Sign.const_syntax_name thy case_name
       
   306       in (case_name', DatatypeCase.case_tr' datatype_of_case case_name')
       
   307       end) case_names, []) thy;
       
   308 
       
   309 val trfun_setup =
       
   310   Sign.add_advanced_trfuns ([],
       
   311     [("_case_syntax", DatatypeCase.case_tr true datatype_of_constr)],
       
   312     [], []);
       
   313 
       
   314 
       
   315 (* prepare types *)
       
   316 
       
   317 fun read_typ thy ((Ts, sorts), str) =
       
   318   let
       
   319     val ctxt = ProofContext.init thy
       
   320       |> fold (Variable.declare_typ o TFree) sorts;
       
   321     val T = Syntax.read_typ ctxt str;
       
   322   in (Ts @ [T], Term.add_tfreesT T sorts) end;
       
   323 
       
   324 fun cert_typ sign ((Ts, sorts), raw_T) =
       
   325   let
       
   326     val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
       
   327       TYPE (msg, _, _) => error msg;
       
   328     val sorts' = Term.add_tfreesT T sorts;
       
   329   in (Ts @ [T],
       
   330       case duplicates (op =) (map fst sorts') of
       
   331          [] => sorts'
       
   332        | dups => error ("Inconsistent sort constraints for " ^ commas dups))
       
   333   end;
       
   334 
       
   335 
       
   336 (**** make datatype info ****)
       
   337 
       
   338 fun make_dt_info alt_names descr sorts induct reccomb_names rec_thms
       
   339     (((((((((i, (_, (tname, _, _))), case_name), case_thms),
       
   340       exhaustion_thm), distinct_thm), inject), nchotomy), case_cong), weak_case_cong) =
       
   341   (tname,
       
   342    {index = i,
       
   343     alt_names = alt_names,
       
   344     descr = descr,
       
   345     sorts = sorts,
       
   346     rec_names = reccomb_names,
       
   347     rec_rewrites = rec_thms,
       
   348     case_name = case_name,
       
   349     case_rewrites = case_thms,
       
   350     induction = induct,
       
   351     exhaustion = exhaustion_thm,
       
   352     distinct = distinct_thm,
       
   353     inject = inject,
       
   354     nchotomy = nchotomy,
       
   355     case_cong = case_cong,
       
   356     weak_case_cong = weak_case_cong});
       
   357 
       
   358 structure DatatypeInterpretation = InterpretationFun(type T = string list val eq = op =);
       
   359 val interpretation = DatatypeInterpretation.interpretation;
       
   360 
       
   361 
       
   362 (******************* definitional introduction of datatypes *******************)
       
   363 
       
   364 fun add_datatype_def flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
       
   365     case_names_induct case_names_exhausts thy =
       
   366   let
       
   367     val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
       
   368 
       
   369     val ((inject, distinct, dist_rewrites, simproc_dists, induct), thy2) = thy |>
       
   370       DatatypeRepProofs.representation_proofs flat_names dt_info new_type_names descr sorts
       
   371         types_syntax constr_syntax case_names_induct;
       
   372 
       
   373     val (casedist_thms, thy3) = DatatypeAbsProofs.prove_casedist_thms new_type_names descr
       
   374       sorts induct case_names_exhausts thy2;
       
   375     val ((reccomb_names, rec_thms), thy4) = DatatypeAbsProofs.prove_primrec_thms
       
   376       flat_names new_type_names descr sorts dt_info inject dist_rewrites
       
   377       (Simplifier.theory_context thy3 dist_ss) induct thy3;
       
   378     val ((case_thms, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
       
   379       flat_names new_type_names descr sorts reccomb_names rec_thms thy4;
       
   380     val (split_thms, thy7) = DatatypeAbsProofs.prove_split_thms new_type_names
       
   381       descr sorts inject dist_rewrites casedist_thms case_thms thy6;
       
   382     val (nchotomys, thy8) = DatatypeAbsProofs.prove_nchotomys new_type_names
       
   383       descr sorts casedist_thms thy7;
       
   384     val (case_congs, thy9) = DatatypeAbsProofs.prove_case_congs new_type_names
       
   385       descr sorts nchotomys case_thms thy8;
       
   386     val (weak_case_congs, thy10) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
       
   387       descr sorts thy9;
       
   388 
       
   389     val dt_infos = map
       
   390       (make_dt_info (SOME new_type_names) (flat descr) sorts induct reccomb_names rec_thms)
       
   391       ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
       
   392         casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
       
   393 
       
   394     val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
       
   395 
       
   396     val thy12 =
       
   397       thy10
       
   398       |> add_case_tr' case_names
       
   399       |> Sign.add_path (space_implode "_" new_type_names)
       
   400       |> add_rules simps case_thms rec_thms inject distinct
       
   401           weak_case_congs (Simplifier.attrib (op addcongs))
       
   402       |> put_dt_infos dt_infos
       
   403       |> add_cases_induct dt_infos induct
       
   404       |> Sign.parent_path
       
   405       |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms) |> snd
       
   406       |> DatatypeInterpretation.data (map fst dt_infos);
       
   407   in
       
   408     ({distinct = distinct,
       
   409       inject = inject,
       
   410       exhaustion = casedist_thms,
       
   411       rec_thms = rec_thms,
       
   412       case_thms = case_thms,
       
   413       split_thms = split_thms,
       
   414       induction = induct,
       
   415       simps = simps}, thy12)
       
   416   end;
       
   417 
       
   418 
       
   419 (*********************** declare existing type as datatype *********************)
       
   420 
       
   421 fun prove_rep_datatype alt_names new_type_names descr sorts induct inject half_distinct thy =
       
   422   let
       
   423     val ((_, [induct']), _) =
       
   424       Variable.importT_thms [induct] (Variable.thm_context induct);
       
   425 
       
   426     fun err t = error ("Ill-formed predicate in induction rule: " ^
       
   427       Syntax.string_of_term_global thy t);
       
   428 
       
   429     fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
       
   430           ((tname, map (fst o dest_TFree) Ts) handle TERM _ => err t)
       
   431       | get_typ t = err t;
       
   432     val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct')));
       
   433 
       
   434     val dt_info = get_datatypes thy;
       
   435 
       
   436     val distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
       
   437     val (case_names_induct, case_names_exhausts) =
       
   438       (mk_case_names_induct descr, mk_case_names_exhausts descr (map #1 dtnames));
       
   439 
       
   440     val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
       
   441 
       
   442     val (casedist_thms, thy2) = thy |>
       
   443       DatatypeAbsProofs.prove_casedist_thms new_type_names [descr] sorts induct
       
   444         case_names_exhausts;
       
   445     val ((reccomb_names, rec_thms), thy3) = DatatypeAbsProofs.prove_primrec_thms
       
   446       false new_type_names [descr] sorts dt_info inject distinct
       
   447       (Simplifier.theory_context thy2 dist_ss) induct thy2;
       
   448     val ((case_thms, case_names), thy4) = DatatypeAbsProofs.prove_case_thms false
       
   449       new_type_names [descr] sorts reccomb_names rec_thms thy3;
       
   450     val (split_thms, thy5) = DatatypeAbsProofs.prove_split_thms
       
   451       new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
       
   452     val (nchotomys, thy6) = DatatypeAbsProofs.prove_nchotomys new_type_names
       
   453       [descr] sorts casedist_thms thy5;
       
   454     val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
       
   455       [descr] sorts nchotomys case_thms thy6;
       
   456     val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
       
   457       [descr] sorts thy7;
       
   458 
       
   459     val ((_, [induct']), thy10) =
       
   460       thy8
       
   461       |> store_thmss "inject" new_type_names inject
       
   462       ||>> store_thmss "distinct" new_type_names distinct
       
   463       ||> Sign.add_path (space_implode "_" new_type_names)
       
   464       ||>> PureThy.add_thms [((Binding.name "induct", induct), [case_names_induct])];
       
   465 
       
   466     val dt_infos = map (make_dt_info alt_names descr sorts induct' reccomb_names rec_thms)
       
   467       ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
       
   468         map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
       
   469 
       
   470     val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
       
   471 
       
   472     val thy11 =
       
   473       thy10
       
   474       |> add_case_tr' case_names
       
   475       |> add_rules simps case_thms rec_thms inject distinct
       
   476            weak_case_congs (Simplifier.attrib (op addcongs))
       
   477       |> put_dt_infos dt_infos
       
   478       |> add_cases_induct dt_infos induct'
       
   479       |> Sign.parent_path
       
   480       |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)
       
   481       |> snd
       
   482       |> DatatypeInterpretation.data (map fst dt_infos);
       
   483   in
       
   484     ({distinct = distinct,
       
   485       inject = inject,
       
   486       exhaustion = casedist_thms,
       
   487       rec_thms = rec_thms,
       
   488       case_thms = case_thms,
       
   489       split_thms = split_thms,
       
   490       induction = induct',
       
   491       simps = simps}, thy11)
       
   492   end;
       
   493 
       
   494 fun gen_rep_datatype prep_term after_qed alt_names raw_ts thy =
       
   495   let
       
   496     fun constr_of_term (Const (c, T)) = (c, T)
       
   497       | constr_of_term t =
       
   498           error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
       
   499     fun no_constr (c, T) = error ("Bad constructor: "
       
   500       ^ Sign.extern_const thy c ^ "::"
       
   501       ^ Syntax.string_of_typ_global thy T);
       
   502     fun type_of_constr (cT as (_, T)) =
       
   503       let
       
   504         val frees = OldTerm.typ_tfrees T;
       
   505         val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
       
   506           handle TYPE _ => no_constr cT
       
   507         val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
       
   508         val _ = if length frees <> length vs then no_constr cT else ();
       
   509       in (tyco, (vs, cT)) end;
       
   510 
       
   511     val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
       
   512     val _ = case map_filter (fn (tyco, _) =>
       
   513         if Symtab.defined (get_datatypes thy) tyco then SOME tyco else NONE) raw_cs
       
   514      of [] => ()
       
   515       | tycos => error ("Type(s) " ^ commas (map quote tycos)
       
   516           ^ " already represented inductivly");
       
   517     val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
       
   518     val ms = case distinct (op =) (map length raw_vss)
       
   519      of [n] => 0 upto n - 1
       
   520       | _ => error ("Different types in given constructors");
       
   521     fun inter_sort m = map (fn xs => nth xs m) raw_vss
       
   522       |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy))
       
   523     val sorts = map inter_sort ms;
       
   524     val vs = Name.names Name.context Name.aT sorts;
       
   525 
       
   526     fun norm_constr (raw_vs, (c, T)) = (c, map_atyps
       
   527       (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
       
   528 
       
   529     val cs = map (apsnd (map norm_constr)) raw_cs;
       
   530     val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs))
       
   531       o fst o strip_type;
       
   532     val new_type_names = map Long_Name.base_name (the_default (map fst cs) alt_names);
       
   533 
       
   534     fun mk_spec (i, (tyco, constr)) = (i, (tyco,
       
   535       map (DtTFree o fst) vs,
       
   536       (map o apsnd) dtyps_of_typ constr))
       
   537     val descr = map_index mk_spec cs;
       
   538     val injs = DatatypeProp.make_injs [descr] vs;
       
   539     val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs);
       
   540     val ind = DatatypeProp.make_ind [descr] vs;
       
   541     val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
       
   542 
       
   543     fun after_qed' raw_thms =
       
   544       let
       
   545         val [[[induct]], injs, half_distincts] =
       
   546           unflat rules (map Drule.zero_var_indexes_list raw_thms);
       
   547             (*FIXME somehow dubious*)
       
   548       in
       
   549         ProofContext.theory_result
       
   550           (prove_rep_datatype alt_names new_type_names descr vs induct injs half_distincts)
       
   551         #-> after_qed
       
   552       end;
       
   553   in
       
   554     thy
       
   555     |> ProofContext.init
       
   556     |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
       
   557   end;
       
   558 
       
   559 val rep_datatype = gen_rep_datatype Sign.cert_term;
       
   560 val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global (K I);
       
   561 
       
   562 
       
   563 
       
   564 (******************************** add datatype ********************************)
       
   565 
       
   566 fun gen_add_datatype prep_typ err flat_names new_type_names dts thy =
       
   567   let
       
   568     val _ = Theory.requires thy "Datatype" "datatype definitions";
       
   569 
       
   570     (* this theory is used just for parsing *)
       
   571 
       
   572     val tmp_thy = thy |>
       
   573       Theory.copy |>
       
   574       Sign.add_types (map (fn (tvs, tname, mx, _) =>
       
   575         (tname, length tvs, mx)) dts);
       
   576 
       
   577     val (tyvars, _, _, _)::_ = dts;
       
   578     val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
       
   579       let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname)
       
   580       in (case duplicates (op =) tvs of
       
   581             [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
       
   582                   else error ("Mutually recursive datatypes must have same type parameters")
       
   583           | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^
       
   584               " : " ^ commas dups))
       
   585       end) dts);
       
   586 
       
   587     val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
       
   588       [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
       
   589 
       
   590     fun prep_dt_spec ((tvs, tname, mx, constrs), tname') (dts', constr_syntax, sorts, i) =
       
   591       let
       
   592         fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
       
   593           let
       
   594             val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
       
   595             val _ = (case fold (curry OldTerm.add_typ_tfree_names) cargs' [] \\ tvs of
       
   596                 [] => ()
       
   597               | vs => error ("Extra type variables on rhs: " ^ commas vs))
       
   598           in (constrs @ [((if flat_names then Sign.full_name tmp_thy else
       
   599                 Sign.full_name_path tmp_thy tname')
       
   600                   (Binding.map_name (Syntax.const_name mx') cname),
       
   601                    map (dtyp_of_typ new_dts) cargs')],
       
   602               constr_syntax' @ [(cname, mx')], sorts'')
       
   603           end handle ERROR msg => cat_error msg
       
   604            ("The error above occured in constructor " ^ quote (Binding.str_of cname) ^
       
   605             " of datatype " ^ quote (Binding.str_of tname));
       
   606 
       
   607         val (constrs', constr_syntax', sorts') =
       
   608           fold prep_constr constrs ([], [], sorts)
       
   609 
       
   610       in
       
   611         case duplicates (op =) (map fst constrs') of
       
   612            [] =>
       
   613              (dts' @ [(i, (Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname),
       
   614                 map DtTFree tvs, constrs'))],
       
   615               constr_syntax @ [constr_syntax'], sorts', i + 1)
       
   616          | dups => error ("Duplicate constructors " ^ commas dups ^
       
   617              " in datatype " ^ quote (Binding.str_of tname))
       
   618       end;
       
   619 
       
   620     val (dts', constr_syntax, sorts', i) =
       
   621       fold prep_dt_spec (dts ~~ new_type_names) ([], [], [], 0);
       
   622     val sorts = sorts' @ (map (rpair (Sign.defaultS tmp_thy)) (tyvars \\ map fst sorts'));
       
   623     val dt_info = get_datatypes thy;
       
   624     val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
       
   625     val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
       
   626       if err then error ("Nonemptiness check failed for datatype " ^ s)
       
   627       else raise exn;
       
   628 
       
   629     val descr' = flat descr;
       
   630     val case_names_induct = mk_case_names_induct descr';
       
   631     val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
       
   632   in
       
   633     add_datatype_def
       
   634       flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
       
   635       case_names_induct case_names_exhausts thy
       
   636   end;
       
   637 
       
   638 val add_datatype = gen_add_datatype cert_typ;
       
   639 val add_datatype_cmd = gen_add_datatype read_typ true;
       
   640 
       
   641 
       
   642 
       
   643 (** package setup **)
       
   644 
       
   645 (* setup theory *)
       
   646 
       
   647 val setup =
       
   648   DatatypeRepProofs.distinctness_limit_setup #>
       
   649   simproc_setup #>
       
   650   trfun_setup #>
       
   651   DatatypeInterpretation.init;
       
   652 
       
   653 
       
   654 (* outer syntax *)
       
   655 
       
   656 local structure P = OuterParse and K = OuterKeyword in
       
   657 
       
   658 val datatype_decl =
       
   659   Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_infix --
       
   660     (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix));
       
   661 
       
   662 fun mk_datatype args =
       
   663   let
       
   664     val names = map
       
   665       (fn ((((NONE, _), t), _), _) => Binding.name_of t | ((((SOME t, _), _), _), _) => t) args;
       
   666     val specs = map (fn ((((_, vs), t), mx), cons) =>
       
   667       (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
       
   668   in snd o add_datatype_cmd false names specs end;
       
   669 
       
   670 val _ =
       
   671   OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
       
   672     (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
       
   673 
       
   674 val _ =
       
   675   OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal
       
   676     (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term
       
   677       >> (fn (alt_names, ts) => Toplevel.print
       
   678            o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
       
   679 
       
   680 end;
       
   681 
       
   682 
       
   683 (* document antiquotation *)
       
   684 
       
   685 val _ = ThyOutput.antiquotation "datatype" Args.tyname
       
   686   (fn {source = src, context = ctxt, ...} => fn dtco =>
       
   687     let
       
   688       val thy = ProofContext.theory_of ctxt;
       
   689       val (vs, cos) = the_datatype_spec thy dtco;
       
   690       val ty = Type (dtco, map TFree vs);
       
   691       fun pretty_typ_bracket (ty as Type (_, _ :: _)) =
       
   692             Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty]
       
   693         | pretty_typ_bracket ty =
       
   694             Syntax.pretty_typ ctxt ty;
       
   695       fun pretty_constr (co, tys) =
       
   696         (Pretty.block o Pretty.breaks)
       
   697           (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
       
   698             map pretty_typ_bracket tys);
       
   699       val pretty_datatype =
       
   700         Pretty.block
       
   701           (Pretty.command "datatype" :: Pretty.brk 1 ::
       
   702            Syntax.pretty_typ ctxt ty ::
       
   703            Pretty.str " =" :: Pretty.brk 1 ::
       
   704            flat (separate [Pretty.brk 1, Pretty.str "| "]
       
   705              (map (single o pretty_constr) cos)));
       
   706     in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end);
       
   707 
       
   708 end;
       
   709