--- a/src/HOL/BNF/Tools/bnf_fp_def_sugar.ML Thu May 02 09:41:29 2013 +0200
+++ b/src/HOL/BNF/Tools/bnf_fp_def_sugar.ML Thu May 02 09:50:58 2013 +0200
@@ -20,7 +20,7 @@
val fp_sugar_of: local_theory -> string -> fp_sugar option
- val build_maps: local_theory -> typ list -> (int -> typ * typ -> term) -> typ * typ -> term
+ val build_map: local_theory -> typ list -> (int -> typ * typ -> term) -> typ * typ -> term
val mk_fp_iter_fun_types: term -> typ list
val mk_fun_arg_typess: int -> int list -> typ -> typ list list
@@ -370,7 +370,7 @@
val TUs' = map dest_funT (fst (strip_typeN live (fastype_of mapx)));
in Term.list_comb (mapx, map build_arg TUs') end;
-fun build_maps lthy Ts build_simple =
+fun build_map lthy Ts build_simple =
let
fun build (TU as (T, U)) =
if T = U then
@@ -381,6 +381,18 @@
| kk => build_simple kk TU);
in build end;
+fun build_map' lthy build_simple =
+ let
+ fun build (TU as (T, U)) =
+ if T = U then
+ id_const T
+ else
+ (case TU of
+ (Type (s, _), Type (s', _)) =>
+ if s = s' then build_map_step lthy build TU else build_simple TU
+ | _ => build_simple TU);
+ in build end;
+
fun build_rel_step lthy build_arg (Type (s, Ts)) =
let
val bnf = the (bnf_of lthy s);
@@ -518,7 +530,7 @@
fun unzip_iters fiters combine (x as Free (_, T)) =
if exists_subtype_in fpTs T then
- combine (x, build_maps lthy fpTs (K o nth fiters) (T, mk_U T) $ x)
+ combine (x, build_map lthy fpTs (K o nth fiters) (T, mk_U T) $ x)
else
([x], []);
@@ -686,7 +698,7 @@
fun intr_coiters fcoiters [] [cf] =
let val T = fastype_of cf in
- if exists_subtype_in Cs T then build_maps lthy Cs (K o nth fcoiters) (T, mk_U T) $ cf
+ if exists_subtype_in Cs T then build_map lthy Cs (K o nth fcoiters) (T, mk_U T) $ cf
else cf
end
| intr_coiters fcoiters [cq] [cf, cf'] =
@@ -1173,14 +1185,7 @@
let
val fpT_to_C = fpT --> C;
- fun build_prod_proj mk_proj (TU as (T, U)) =
- if T = U then
- id_const T
- else
- (case TU of
- (Type (s, _), Type (s', _)) =>
- if s = s' then build_map_step lthy (build_prod_proj mk_proj) TU else mk_proj T
- | _ => mk_proj T);
+ fun build_prod_proj mk_proj = build_map' lthy (mk_proj o fst);
(* TODO: Avoid these complications; cf. corec case *)
fun mk_U proj (Type (s as @{type_name prod}, Ts as [T', U])) =
@@ -1233,15 +1238,7 @@
let
val B_to_fpT = C --> fpT;
- fun build_sum_inj mk_inj (TU as (T, U)) =
- if T = U then
- id_const T
- else
- (case TU of
- (Type (s, _), Type (s', _)) =>
- if s = s' then build_map_step lthy (build_sum_inj mk_inj) TU
- else uncurry mk_inj (dest_sumT U)
- | _ => uncurry mk_inj (dest_sumT U));
+ fun build_sum_inj mk_inj = build_map' lthy (uncurry mk_inj o dest_sumT o snd);
fun build_dtor_coiter_arg _ [] [cf] = cf
| build_dtor_coiter_arg T [cq] [cf, cf'] =