--- a/Nominal/Nominal2.thy Wed Jun 22 13:40:25 2011 +0100
+++ b/Nominal/Nominal2.thy Wed Jun 22 14:14:54 2011 +0100
@@ -189,23 +189,21 @@
val (raw_dt_names, raw_dts, raw_bclauses, raw_bn_funs, raw_bn_eqs, cnstr_names, lthy0) =
define_raw_dts dts bn_funs bn_eqs bclauses lthy
- val dtinfo = Datatype.the_info (ProofContext.theory_of lthy0) (hd raw_dt_names)
- val {descr, sorts, ...} = dtinfo
+ val dtinfos = map (Datatype.the_info (ProofContext.theory_of lthy0)) raw_dt_names
+ val {descr, sorts, ...} = hd dtinfos
- val raw_tys = all_dtyps descr sorts
+ val raw_tys = Datatype_Aux.get_rec_types descr sorts
val tvs = hd raw_tys
|> snd o dest_Type
|> map dest_TFree
- val dtinfos = map (Datatype.the_info (ProofContext.theory_of lthy0)) raw_dt_names
-
val raw_cns_info = all_dtyp_constrs_types descr sorts
val raw_constrs = (map o map) (fn (c, _, _, _) => c) raw_cns_info
val raw_inject_thms = flat (map #inject dtinfos)
val raw_distinct_thms = flat (map #distinct dtinfos)
- val raw_induct_thm = #induct dtinfo
- val raw_induct_thms = #inducts dtinfo
+ val raw_induct_thm = #induct (hd dtinfos)
+ val raw_induct_thms = #inducts (hd dtinfos)
val raw_exhaust_thms = map #exhaust dtinfos
val raw_size_trms = map HOLogic.size_const raw_tys
val raw_size_thms = Size.size_thms (ProofContext.theory_of lthy0) (hd raw_dt_names)