src/HOL/BNF/Tools/bnf_fp_n2m_sugar.ML
changeset 54300 bd36da55d825
parent 54288 ce58fb149ff6
child 54301 e0943a2ee400
equal deleted inserted replaced
54299:bc24e1ccfd35 54300:bd36da55d825
   282            is_some (Datatype_Data.get_info (Proof_Context.theory_of lthy) s) then
   282            is_some (Datatype_Data.get_info (Proof_Context.theory_of lthy) s) then
   283           error (qsoty T ^ " is not a new-style datatype (cf. \"datatype_new\")")
   283           error (qsoty T ^ " is not a new-style datatype (cf. \"datatype_new\")")
   284         else
   284         else
   285           not_co_datatype0 T
   285           not_co_datatype0 T
   286       | not_co_datatype T = not_co_datatype0 T;
   286       | not_co_datatype T = not_co_datatype0 T;
   287     fun not_mutually_nested_rec Ts1 Ts2 =
       
   288       error (qsotys Ts1 ^ " is neither mutually recursive with " ^ qsotys Ts2 ^
       
   289         " nor nested recursive via " ^ qsotys Ts2);
       
   290 
   287 
   291     val _ = (case Library.duplicates (op =) actual_Ts of [] => () | T :: _ => duplicate_datatype T);
   288     val _ = (case Library.duplicates (op =) actual_Ts of [] => () | T :: _ => duplicate_datatype T);
   292 
   289 
   293     val perm_actual_Ts =
   290     val perm_actual_Ts =
   294       sort (prod_ord int_ord Term_Ord.typ_ord o pairself (`Term.size_of_typ)) actual_Ts;
   291       sort (prod_ord int_ord Term_Ord.typ_ord o pairself (`Term.size_of_typ)) actual_Ts;
   316     fun gather_types _ _ num_groups seen gen_seen [] = (num_groups, seen, gen_seen)
   313     fun gather_types _ _ num_groups seen gen_seen [] = (num_groups, seen, gen_seen)
   317       | gather_types lthy rho num_groups seen gen_seen ((T as Type (_, tyargs)) :: Ts) =
   314       | gather_types lthy rho num_groups seen gen_seen ((T as Type (_, tyargs)) :: Ts) =
   318         let
   315         let
   319           val {fp_res = {Ts = mutual_Ts0, ...}, ...} = the_fp_sugar_of T;
   316           val {fp_res = {Ts = mutual_Ts0, ...}, ...} = the_fp_sugar_of T;
   320           val mutual_Ts = map (retypargs tyargs) mutual_Ts0;
   317           val mutual_Ts = map (retypargs tyargs) mutual_Ts0;
   321 
       
   322           val _ = seen = [] orelse exists (exists_subtype_in seen) mutual_Ts orelse
       
   323             not_mutually_nested_rec mutual_Ts seen;
       
   324 
   318 
   325           fun fresh_tyargs () =
   319           fun fresh_tyargs () =
   326             let
   320             let
   327               (* The name "'z" is unlikely to clash with the context, yielding more cache hits. *)
   321               (* The name "'z" is unlikely to clash with the context, yielding more cache hits. *)
   328               val (gen_tyargs, lthy') =
   322               val (gen_tyargs, lthy') =