Nominal/Nominal2.thy
branchNominal2-Isabelle2016
changeset 3243 c4f31f1564b7
parent 3239 67370521c09c
child 3245 017e33849f4d
equal deleted inserted replaced
3242:4af8a92396ce 3243:c4f31f1564b7
   176   val raw_distinct_thms = flat (map #distinct dtinfos)
   176   val raw_distinct_thms = flat (map #distinct dtinfos)
   177   val raw_induct_thm = #induct (hd dtinfos)
   177   val raw_induct_thm = #induct (hd dtinfos)
   178   val raw_induct_thms = #inducts (hd dtinfos)
   178   val raw_induct_thms = #inducts (hd dtinfos)
   179   val raw_exhaust_thms = map #exhaust dtinfos
   179   val raw_exhaust_thms = map #exhaust dtinfos
   180   val raw_size_trms = map HOLogic.size_const raw_tys
   180   val raw_size_trms = map HOLogic.size_const raw_tys
   181   val raw_size_thms = these (Option.map ((fn ths => drop (length ths div 2) ths) o fst o snd)
   181   val raw_size_thms = these (Option.map ((fn ths => drop (length ths div 2) ths) o #2 o #2)
   182     (BNF_LFP_Size.size_of lthy1 (hd raw_full_dt_names')))
   182     (BNF_LFP_Size.size_of lthy1 (hd raw_full_dt_names')))
   183 
   183 
   184   val raw_result = RawDtInfo 
   184   val raw_result = RawDtInfo 
   185     {raw_dt_names = raw_full_dt_names',
   185     {raw_dt_names = raw_full_dt_names',
   186      raw_dts = raw_dts,
   186      raw_dts = raw_dts,