joint work with blanchet: intermediate typedef for the input to fp-operations
authortraytel
Tue, 25 Feb 2014 18:14:26 +0100
changeset 55803 74d3fe9031d8
parent 55802 f7ceebe2f1b5
child 55804 341fbb9bdda1
joint work with blanchet: intermediate typedef for the input to fp-operations
src/HOL/BNF_Comp.thy
src/HOL/BNF_Def.thy
src/HOL/BNF_FP_Base.thy
src/HOL/BNF_Util.thy
src/HOL/Lattices_Big.thy
src/HOL/Tools/BNF/bnf_comp.ML
src/HOL/Tools/BNF/bnf_fp_def_sugar.ML
src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML
src/HOL/Tools/BNF/bnf_fp_n2m.ML
src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML
src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML
src/HOL/Tools/BNF/bnf_fp_util.ML
src/HOL/Tools/BNF/bnf_gfp.ML
src/HOL/Tools/BNF/bnf_lfp.ML
src/HOL/Tools/BNF/bnf_util.ML
--- a/src/HOL/BNF_Comp.thy	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/BNF_Comp.thy	Tue Feb 25 18:14:26 2014 +0100
@@ -70,6 +70,49 @@
 lemma OO_Grp_cong: "A = B \<Longrightarrow> (Grp A f)^--1 OO Grp A g = (Grp B f)^--1 OO Grp B g"
 by (rule arg_cong)
 
+lemma vimage2p_relcompp_mono: "R OO S \<le> T \<Longrightarrow>
+  vimage2p f g R OO vimage2p g h S \<le> vimage2p f h T"
+  unfolding vimage2p_def by auto
+
+lemma type_copy_map_cong0: "M (g x) = N (h x) \<Longrightarrow> (f o M o g) x = (f o N o h) x"
+  by auto
+
+lemma type_copy_set_bd: "(\<And>y. |S y| \<le>o bd) \<Longrightarrow> |(S o Rep) x| \<le>o bd"
+  by auto
+
+lemma vimage2p_cong: "R = S \<Longrightarrow> vimage2p f g R = vimage2p f g S"
+  by simp
+
+context
+fixes Rep Abs
+assumes type_copy: "type_definition Rep Abs UNIV"
+begin
+
+lemma type_copy_map_id0: "M = id \<Longrightarrow> Abs o M o Rep = id"
+  using type_definition.Rep_inverse[OF type_copy] by auto
+lemma type_copy_map_comp0: "M = M1 o M2 \<Longrightarrow> f o M o g = (f o M1 o Rep) o (Abs o M2 o g)"
+  using type_definition.Abs_inverse[OF type_copy UNIV_I] by auto
+lemma type_copy_set_map0: "S o M = image f o S' \<Longrightarrow> (S o Rep) o (Abs o M o g) = image f o (S' o g)"
+  using type_definition.Abs_inverse[OF type_copy UNIV_I] by (auto simp: o_def fun_eq_iff)
+lemma type_copy_wit: "x \<in> (S o Rep) (Abs y) \<Longrightarrow> x \<in> S y"
+  using type_definition.Abs_inverse[OF type_copy UNIV_I] by auto
+lemma type_copy_vimage2p_Grp_Rep: "vimage2p f Rep (Grp (Collect P) h) =
+    Grp (Collect (\<lambda>x. P (f x))) (Abs o h o f)"
+  unfolding vimage2p_def Grp_def fun_eq_iff
+  by (auto simp: type_definition.Abs_inverse[OF type_copy UNIV_I]
+   type_definition.Rep_inverse[OF type_copy] dest: sym)
+lemma type_copy_vimage2p_Grp_Abs:
+  "\<And>h. vimage2p g Abs (Grp (Collect P) h) = Grp (Collect (\<lambda>x. P (g x))) (Rep o h o g)"
+  unfolding vimage2p_def Grp_def fun_eq_iff
+  by (auto simp: type_definition.Abs_inverse[OF type_copy UNIV_I]
+   type_definition.Rep_inverse[OF type_copy] dest: sym)
+lemma vimage2p_relcompp_converse:
+  "vimage2p f g (R^--1 OO S) = (vimage2p Rep f R)^--1 OO vimage2p Rep g S"
+  unfolding vimage2p_def relcompp.simps conversep.simps fun_eq_iff image_def
+  by (metis surjD[OF type_definition.Rep_range[OF type_copy]])
+
+end
+
 ML_file "Tools/BNF/bnf_comp_tactics.ML"
 ML_file "Tools/BNF/bnf_comp.ML"
 
--- a/src/HOL/BNF_Def.thy	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/BNF_Def.thy	Tue Feb 25 18:14:26 2014 +0100
@@ -137,9 +137,6 @@
   ((\<lambda>i. if i \<in> g ` C then the_inv_into C g i else x) o g) i = id i"
 unfolding Func_def by (auto elim: the_inv_into_f_f)
 
-definition vimage2p where
-  "vimage2p f g R = (\<lambda>x y. R (f x) (g y))"
-
 lemma vimage2pI: "R (f x) (g y) \<Longrightarrow> vimage2p f g R x y"
   unfolding vimage2p_def by -
 
--- a/src/HOL/BNF_FP_Base.thy	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/BNF_FP_Base.thy	Tue Feb 25 18:14:26 2014 +0100
@@ -64,6 +64,11 @@
 lemma obj_one_pointE: "\<forall>x. s = x \<longrightarrow> P \<Longrightarrow> P"
 by blast
 
+lemma type_copy_obj_one_point_absE:
+  assumes "type_definition Rep Abs UNIV"
+  shows "\<forall>x. s = Abs x \<longrightarrow> P \<Longrightarrow> P"
+  using type_definition.Rep_inverse[OF assms] by metis
+
 lemma obj_sumE_f:
 "\<lbrakk>\<forall>x. s = f (Inl x) \<longrightarrow> P; \<forall>x. s = f (Inr x) \<longrightarrow> P\<rbrakk> \<Longrightarrow> \<forall>x. s = f x \<longrightarrow> P"
 by (rule allI) (metis sumE)
@@ -132,7 +137,7 @@
   unfolding case_sum_o_sum_map id_comp comp_id ..
 
 lemma fun_rel_def_butlast:
-  "(fun_rel R (fun_rel S T)) f g = (\<forall>x y. R x y \<longrightarrow> (fun_rel S T) (f x) (g y))"
+  "fun_rel R (fun_rel S T) f g = (\<forall>x y. R x y \<longrightarrow> (fun_rel S T) (f x) (g y))"
   unfolding fun_rel_def ..
 
 lemma subst_eq_imp: "(\<forall>a b. a = b \<longrightarrow> P a b) \<equiv> (\<forall>a. P a a)"
@@ -148,6 +153,30 @@
    (\<And>x. x \<in> P \<Longrightarrow> f x \<in> Q)"
   unfolding Grp_def by rule auto
 
+lemma vimage2p_mono: "vimage2p f g R x y \<Longrightarrow> R \<le> S \<Longrightarrow> vimage2p f g S x y"
+  unfolding vimage2p_def by blast
+
+lemma vimage2p_refl: "(\<And>x. R x x) \<Longrightarrow> vimage2p f f R x x"
+  unfolding vimage2p_def by auto
+
+lemma
+  assumes "type_definition Rep Abs UNIV"
+  shows type_copy_Rep_o_Abs: "Rep \<circ> Abs = id" and type_copy_Abs_o_Rep: "Abs o Rep = id"
+  unfolding fun_eq_iff comp_apply id_apply
+    type_definition.Abs_inverse[OF assms UNIV_I] type_definition.Rep_inverse[OF assms] by simp_all
+
+lemma type_copy_map_comp0_undo:
+  assumes "type_definition Rep Abs UNIV"
+          "type_definition Rep' Abs' UNIV"
+          "type_definition Rep'' Abs'' UNIV"
+  shows "Abs' o M o Rep'' = (Abs' o M1 o Rep) o (Abs o M2 o Rep'') \<Longrightarrow> M1 o M2 = M"
+  by (rule sym) (auto simp: fun_eq_iff type_definition.Abs_inject[OF assms(2) UNIV_I UNIV_I]
+    type_definition.Abs_inverse[OF assms(1) UNIV_I]
+    type_definition.Abs_inverse[OF assms(3) UNIV_I] dest: spec[of _ "Abs'' x" for x])
+
+lemma vimage2p_comp: "vimage2p (f1 o f2) (g1 o g2) = vimage2p f2 g2 o vimage2p f1 g1"
+  unfolding fun_eq_iff vimage2p_def o_apply by simp
+
 ML_file "Tools/BNF/bnf_fp_util.ML"
 ML_file "Tools/BNF/bnf_fp_def_sugar_tactics.ML"
 ML_file "Tools/BNF/bnf_fp_def_sugar.ML"
--- a/src/HOL/BNF_Util.thy	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/BNF_Util.thy	Tue Feb 25 18:14:26 2014 +0100
@@ -44,6 +44,9 @@
 
 definition "Grp A f = (\<lambda>a b. b = f a \<and> a \<in> A)"
 
+definition vimage2p where
+  "vimage2p f g R = (\<lambda>x y. R (f x) (g y))"
+
 ML_file "Tools/BNF/bnf_util.ML"
 ML_file "Tools/BNF/bnf_tactics.ML"
 
--- a/src/HOL/Lattices_Big.thy	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/Lattices_Big.thy	Tue Feb 25 18:14:26 2014 +0100
@@ -792,4 +792,3 @@
 end
 
 end
-
--- a/src/HOL/Tools/BNF/bnf_comp.ML	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_comp.ML	Tue Feb 25 18:14:26 2014 +0100
@@ -27,8 +27,23 @@
   val normalize_bnfs: (int -> binding -> binding) -> ''a list list -> ''a list ->
     (''a list list -> ''a list) -> BNF_Def.bnf list -> unfold_set -> Proof.context ->
     (int list list * ''a list) * (BNF_Def.bnf list * (unfold_set * Proof.context))
+
+  type absT_info =
+    {absT: typ,
+     repT: typ,
+     abs: term,
+     rep: term,
+     abs_inject: thm,
+     abs_inverse: thm,
+     type_definition: thm}
+
+  val morph_absT_info: morphism -> absT_info -> absT_info
+  val mk_absT: theory -> typ -> typ -> typ -> typ
+  val mk_repT: typ -> typ -> typ -> typ
+  val mk_abs: typ -> term -> term
+  val mk_rep: typ -> term -> term
   val seal_bnf: (binding -> binding) -> unfold_set -> binding -> typ list -> BNF_Def.bnf ->
-    Proof.context -> (BNF_Def.bnf * typ list) * local_theory
+    Proof.context -> (BNF_Def.bnf * (typ list * absT_info)) * local_theory
 end;
 
 structure BNF_Comp : BNF_COMP =
@@ -572,6 +587,41 @@
 
 (* Hide the type of the bound (optimization) and unfold the definitions (nicer to the user) *)
 
+type absT_info =
+  {absT: typ,
+   repT: typ,
+   abs: term,
+   rep: term,
+   abs_inject: thm,
+   abs_inverse: thm,
+   type_definition: thm};
+
+fun morph_absT_info phi
+  {absT, repT, abs, rep, abs_inject, abs_inverse, type_definition} =
+  {absT = Morphism.typ phi absT,
+   repT = Morphism.typ phi repT,
+   abs = Morphism.term phi abs,
+   rep = Morphism.term phi rep,
+   abs_inject = Morphism.thm phi abs_inject,
+   abs_inverse = Morphism.thm phi abs_inverse,
+   type_definition = Morphism.thm phi type_definition};
+
+fun mk_absT thy repT absT repU =
+  let val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (repT, repU) Vartab.empty) [];
+  in Term.typ_subst_TVars rho absT end;
+
+fun mk_repT (t as Type (C, Ts)) repT (u as Type (C', Us)) =
+    if C = C' andalso length Ts = length Us then Term.typ_subst_atomic (Ts ~~ Us) repT
+    else raise Term.TYPE ("mk_repT", [t, repT, u], [])
+  | mk_repT t repT u =  raise Term.TYPE ("mk_repT", [t, repT, u], []);
+
+fun mk_abs_or_rep getT (Type (_, Us)) abs =
+  let val Ts = snd (dest_Type (getT (fastype_of abs)))
+  in Term.subst_atomic_types (Ts ~~ Us) abs end;
+
+val mk_abs = mk_abs_or_rep range_type;
+val mk_rep = mk_abs_or_rep domain_type;
+
 fun seal_bnf qualify (unfold_set : unfold_set) b Ds bnf lthy =
   let
     val live = live_of_bnf bnf;
@@ -582,6 +632,10 @@
     val (Bs, _) = apfst (map TFree)
       (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
 
+    val (((fs, fs'), (Rs, Rs')), _(*names_lthy*)) = lthy
+      |> mk_Frees' "f" (map2 (curry op -->) As Bs)
+      ||>> mk_Frees' "R" (map2 mk_pred2T As Bs)
+
     val map_unfolds = #map_unfolds unfold_set;
     val set_unfoldss = #set_unfoldss unfold_set;
     val rel_unfolds = #rel_unfolds unfold_set;
@@ -596,12 +650,35 @@
     fun unfold_sets ctxt = fold (unfold_thms ctxt) set_unfoldss;
     fun unfold_rels ctxt = unfold_thms ctxt rel_unfolds;
     fun unfold_all ctxt = unfold_sets ctxt o unfold_maps ctxt o unfold_rels ctxt;
-    val bnf_map = expand_maps (mk_map_of_bnf Ds As Bs bnf);
-    val bnf_sets = map (expand_maps o expand_sets)
+
+    val repTA = mk_T_of_bnf Ds As bnf;
+    val repTB = mk_T_of_bnf Ds Bs bnf;
+    val T_bind = qualify b;
+    val TA_params = Term.add_tfreesT repTA [];
+    val TB_params = Term.add_tfreesT repTB [];
+    val ((T_name, (T_glob_info, T_loc_info)), lthy) =
+      typedef (T_bind, TA_params, NoSyn)
+        (HOLogic.mk_UNIV repTA) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
+    val TA = Type (T_name, map TFree TA_params);
+    val TB = Type (T_name, map TFree TB_params);
+    val RepA = Const (#Rep_name T_glob_info, TA --> repTA);
+    val RepB = Const (#Rep_name T_glob_info, TB --> repTB);
+    val AbsA = Const (#Abs_name T_glob_info, repTA --> TA);
+    val AbsB = Const (#Abs_name T_glob_info, repTB --> TB);
+    val typedef_thm = #type_definition T_loc_info;
+    val Abs_inject' = #Abs_inject T_loc_info OF @{thms UNIV_I UNIV_I};
+    val Abs_inverse' = #Abs_inverse T_loc_info OF @{thms UNIV_I};
+
+    val absT_info = {absT = TA, repT = repTA, abs = AbsA, rep = RepA, abs_inject = Abs_inject',
+      abs_inverse = Abs_inverse', type_definition = typedef_thm};
+
+    val bnf_map = fold_rev Term.absfree fs' (HOLogic.mk_comp (HOLogic.mk_comp (AbsB,
+      Term.list_comb (expand_maps (mk_map_of_bnf Ds As Bs bnf), fs)), RepA));
+    val bnf_sets = map ((fn t => HOLogic.mk_comp (t, RepA)) o expand_maps o expand_sets)
       (mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf);
     val bnf_bd = mk_bd_of_bnf Ds As bnf;
-    val bnf_rel = expand_rels (mk_rel_of_bnf Ds As Bs bnf);
-    val T = mk_T_of_bnf Ds As bnf;
+    val bnf_rel = fold_rev Term.absfree Rs' (mk_vimage2p RepA RepB $
+      (Term.list_comb (expand_rels (mk_rel_of_bnf Ds As Bs bnf), Rs)));
 
     (*bd may depend only on dead type variables*)
     val bd_repT = fst (dest_relT (fastype_of bnf_bd));
@@ -626,30 +703,44 @@
     val bd_cinfinite =
       (@{thm Cinfinite_cong} OF [bd_ordIso, bd_Cinfinite_of_bnf bnf]) RS conjunct1;
 
-    val set_bds =
-      map (fn thm => @{thm ordLeq_ordIso_trans} OF [thm, bd_ordIso]) (set_bd_of_bnf bnf);
-
-    fun mk_tac thm ctxt =
-      (rtac (unfold_all ctxt thm) THEN'
-      SOLVE o REPEAT_DETERM o (atac ORELSE' Goal.assume_rule_tac ctxt)) 1;
+    fun map_id0_tac ctxt =
+      rtac (@{thm type_copy_map_id0} OF [typedef_thm, unfold_maps ctxt (map_id0_of_bnf bnf)]) 1;
+    fun map_comp0_tac ctxt =
+      rtac (@{thm type_copy_map_comp0} OF [typedef_thm, unfold_maps ctxt (map_comp0_of_bnf bnf)]) 1;
+    fun map_cong0_tac ctxt =
+      EVERY' (rtac @{thm type_copy_map_cong0} :: rtac (unfold_all ctxt (map_cong0_of_bnf bnf)) ::
+        map (fn i => EVERY' [select_prem_tac live (dtac meta_spec) i, etac meta_mp,
+          etac (o_apply RS equalityD2 RS set_mp)]) (1 upto live)) 1;
+    fun set_map0_tac thm ctxt =
+      rtac (@{thm type_copy_set_map0} OF [typedef_thm, unfold_all ctxt thm]) 1;
+    val set_bd_tacs = map (fn thm => fn ctxt => rtac (@{thm ordLeq_ordIso_trans} OF
+        [unfold_sets ctxt thm, bd_ordIso] RS @{thm type_copy_set_bd}) 1)
+      (set_bd_of_bnf bnf);
+    fun le_rel_OO_tac ctxt =
+      rtac (unfold_rels ctxt (le_rel_OO_of_bnf bnf) RS @{thm vimage2p_relcompp_mono}) 1;
+    fun rel_OO_Grp_tac ctxt =
+      (rtac (unfold_all ctxt (rel_OO_Grp_of_bnf bnf) RS @{thm vimage2p_cong} RS trans) THEN'
+      SELECT_GOAL (unfold_thms_tac ctxt [o_apply, typedef_thm RS @{thm type_copy_vimage2p_Grp_Rep},
+        typedef_thm RS @{thm vimage2p_relcompp_converse}]) THEN' rtac refl) 1;
 
-    val tacs = zip_axioms (mk_tac (map_id0_of_bnf bnf)) (mk_tac (map_comp0_of_bnf bnf))
-      (mk_tac (map_cong0_of_bnf bnf)) (map mk_tac (set_map0_of_bnf bnf))
-      (K (rtac bd_card_order 1)) (K (rtac bd_cinfinite 1)) (map mk_tac set_bds)
-      (mk_tac (le_rel_OO_of_bnf bnf))
-      (mk_tac (rel_OO_Grp_of_bnf bnf));
+    val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac
+      (map set_map0_tac (set_map0_of_bnf bnf)) (K (rtac bd_card_order 1)) (K (rtac bd_cinfinite 1))
+      set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;
 
-    val bnf_wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
+    val bnf_wits = map (fn (I, t) =>
+        fold Term.absdummy (map (nth As) I)
+          (AbsA $ Term.list_comb (t, map Bound (0 upto length I - 1))))
+      (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
 
-    fun wit_tac ctxt =
+    fun wit_tac ctxt = ALLGOALS (dtac (typedef_thm RS @{thm type_copy_wit})) THEN
       mk_simple_wit_tac (map (unfold_all ctxt) (wit_thms_of_bnf bnf));
 
     val (bnf', lthy') =
       bnf_def Hardly_Inline (user_policy Dont_Note) qualify tacs wit_tac (SOME all_deads)
         Binding.empty Binding.empty []
-        ((((((b, T), bnf_map), bnf_sets), bnf_bd'), bnf_wits), SOME bnf_rel) lthy;
+        ((((((b, TA), bnf_map), bnf_sets), bnf_bd'), bnf_wits), SOME bnf_rel) lthy;
   in
-    ((bnf', all_deads), lthy')
+    ((bnf', (all_deads, absT_info)), lthy')
   end;
 
 fun key_of_types Ts = Type ("", Ts);
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML	Tue Feb 25 18:14:26 2014 +0100
@@ -14,6 +14,7 @@
      fp_res_index: int,
      fp_res: BNF_FP_Util.fp_result,
      pre_bnf: BNF_Def.bnf,
+     absT_info: BNF_Comp.absT_info,
      nested_bnfs: BNF_Def.bnf list,
      nesting_bnfs: BNF_Def.bnf list,
      ctrXs_Tss: typ list list,
@@ -56,7 +57,7 @@
   val transfer_gfp_sugar_thms: Proof.context -> gfp_sugar_thms -> gfp_sugar_thms
 
   val mk_co_iters_prelims: BNF_Util.fp_kind -> typ list list list -> typ list -> typ list ->
-    int list -> int list list -> term list list -> Proof.context ->
+    typ list -> typ list -> int list -> int list list -> term list list -> Proof.context ->
     (term list list
      * (typ list list * typ list list list list * term list list
         * term list list list list) list option
@@ -65,26 +66,30 @@
     * Proof.context
   val repair_nullary_single_ctr: typ list list -> typ list list
   val mk_coiter_p_pred_types: typ list -> int list -> typ list list
+  val mk_coiter_fun_arg_types: typ list list list -> typ list -> typ list -> typ list -> int list ->
+    int list list -> term ->
+    typ list list
+    * (typ list list list list * typ list list list * typ list list list list * typ list)
   val define_iters: string list ->
     (typ list list * typ list list list list * term list list * term list list list list) list ->
-    (string -> binding) -> typ list -> typ list -> term list -> Proof.context ->
+    (string -> binding) -> typ list -> typ list -> term list -> term list -> Proof.context ->
     (term list * thm list) * Proof.context
   val define_coiters: string list -> string * term list * term list list
     * ((term list list * term list list list) * typ list) list ->
-    (string -> binding) -> typ list -> typ list -> term list -> Proof.context ->
+    (string -> binding) -> typ list -> typ list -> term list -> term list -> Proof.context ->
     (term list * thm list) * Proof.context
   val derive_induct_iters_thms_for_types: BNF_Def.bnf list ->
     (typ list list * typ list list list list * term list list * term list list list list) list ->
     thm -> thm list list -> BNF_Def.bnf list -> BNF_Def.bnf list -> typ list -> typ list ->
-    typ list -> typ list list list -> term list list -> thm list list -> term list list ->
-    thm list list -> local_theory -> lfp_sugar_thms
+    typ list -> typ list list list -> thm list -> thm list -> thm list -> term list list ->
+    thm list list -> term list list -> thm list list -> local_theory -> lfp_sugar_thms
   val derive_coinduct_coiters_thms_for_types: BNF_Def.bnf list ->
     string * term list * term list list * ((term list list * term list list list)
       * typ list) list ->
     thm -> thm list -> thm list -> thm list list -> BNF_Def.bnf list -> typ list -> typ list ->
-    typ list -> typ list list list -> int list list -> int list list -> int list -> thm list list ->
-    Ctr_Sugar.ctr_sugar list -> term list list -> thm list list -> (thm list -> thm list) ->
-    local_theory -> gfp_sugar_thms
+    typ list -> typ list list list -> int list list -> int list list -> int list -> thm list ->
+    thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> term list list ->
+    thm list list -> (thm list -> thm list) -> local_theory -> gfp_sugar_thms
 
   type co_datatype_spec =
     ((((binding option * (typ * sort)) list * binding) * (binding * binding)) * mixfix)
@@ -92,12 +97,14 @@
 
   val co_datatypes: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list ->
       binding list list -> binding list -> (string * sort) list -> typ list * typ list list ->
-      BNF_Def.bnf list -> local_theory -> BNF_FP_Util.fp_result * local_theory) ->
+      BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory ->
+      BNF_FP_Util.fp_result * local_theory) ->
     (bool * bool) * co_datatype_spec list ->
     local_theory -> local_theory
   val parse_co_datatype_cmd: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list ->
       binding list list -> binding list -> (string * sort) list -> typ list * typ list list ->
-      BNF_Def.bnf list -> local_theory -> BNF_FP_Util.fp_result * local_theory) ->
+      BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory ->
+      BNF_FP_Util.fp_result * local_theory) ->
     (local_theory -> local_theory) parser
 end;
 
@@ -121,6 +128,7 @@
    fp_res_index: int,
    fp_res: fp_result,
    pre_bnf: bnf,
+   absT_info: absT_info,
    nested_bnfs: bnf list,
    nesting_bnfs: bnf list,
    ctrXs_Tss: typ list list,
@@ -134,15 +142,16 @@
    disc_co_iterss: thm list list,
    sel_co_itersss: thm list list list};
 
-fun morph_fp_sugar phi ({T, X, fp, fp_res, fp_res_index, pre_bnf, nested_bnfs, nesting_bnfs,
-    ctrXs_Tss, ctr_defs, ctr_sugar, co_iters, maps, common_co_inducts, co_inducts, co_iter_thmss,
-    disc_co_iterss, sel_co_itersss} : fp_sugar) =
+fun morph_fp_sugar phi ({T, X, fp, fp_res, fp_res_index, pre_bnf, absT_info, nested_bnfs,
+    nesting_bnfs, ctrXs_Tss, ctr_defs, ctr_sugar, co_iters, maps, common_co_inducts, co_inducts,
+    co_iter_thmss, disc_co_iterss, sel_co_itersss} : fp_sugar) =
   {T = Morphism.typ phi T,
    X = Morphism.typ phi X,
    fp = fp,
    fp_res = morph_fp_result phi fp_res,
    fp_res_index = fp_res_index,
    pre_bnf = morph_bnf phi pre_bnf,
+   absT_info = morph_absT_info phi absT_info,
    nested_bnfs = map (morph_bnf phi) nested_bnfs,
    nesting_bnfs = map (morph_bnf phi) nesting_bnfs,
    ctrXs_Tss = map (map (Morphism.typ phi)) ctrXs_Tss,
@@ -183,17 +192,18 @@
   Local_Theory.declaration {syntax = false, pervasive = true}
     (fn phi => Data.map (Symtab.update (key, morph_fp_sugar phi fp_sugar)));
 
-fun register_fp_sugars Xs fp pre_bnfs nested_bnfs nesting_bnfs (fp_res as {Ts, ...}) ctrXs_Tsss
-    ctr_defss ctr_sugars co_iterss mapss common_co_inducts co_inductss co_iter_thmsss
+fun register_fp_sugars Xs fp pre_bnfs absT_infos nested_bnfs nesting_bnfs (fp_res as {Ts, ...})
+    ctrXs_Tsss ctr_defss ctr_sugars co_iterss mapss common_co_inducts co_inductss co_iter_thmsss
     disc_co_itersss sel_co_iterssss lthy =
   (0, lthy)
   |> fold (fn T as Type (s, _) => fn (kk, lthy) => (kk + 1,
     register_fp_sugar s {T = T, X = nth Xs kk, fp = fp, fp_res = fp_res, fp_res_index = kk,
-        pre_bnf = nth pre_bnfs kk, nested_bnfs = nested_bnfs, nesting_bnfs = nesting_bnfs,
-        ctrXs_Tss = nth ctrXs_Tsss kk, ctr_defs = nth ctr_defss kk, ctr_sugar = nth ctr_sugars kk,
-        co_iters = nth co_iterss kk, maps = nth mapss kk, common_co_inducts = common_co_inducts,
-        co_inducts = nth co_inductss kk, co_iter_thmss = nth co_iter_thmsss kk,
-        disc_co_iterss = nth disc_co_itersss kk, sel_co_itersss = nth sel_co_iterssss kk}
+        pre_bnf = nth pre_bnfs kk, absT_info = nth absT_infos kk, nested_bnfs = nested_bnfs,
+        nesting_bnfs = nesting_bnfs, ctrXs_Tss = nth ctrXs_Tsss kk, ctr_defs = nth ctr_defss kk,
+        ctr_sugar = nth ctr_sugars kk, co_iters = nth co_iterss kk, maps = nth mapss kk,
+        common_co_inducts = common_co_inducts, co_inducts = nth co_inductss kk,
+        co_iter_thmss = nth co_iter_thmsss kk, disc_co_iterss = nth disc_co_itersss kk,
+        sel_co_itersss = nth sel_co_iterssss kk}
       lthy)) Ts
   |> snd;
 
@@ -227,12 +237,6 @@
   | flat_corec_preds_predsss_gettersss (p :: ps) (qss :: qsss) (fss :: fsss) =
     p :: flat_corec_predss_getterss qss fss @ flat_corec_preds_predsss_gettersss ps qsss fsss;
 
-fun mk_tupled_fun x f xs =
-  if xs = [x] then f else HOLogic.tupled_lambda x (Term.list_comb (f, xs));
-
-fun mk_uncurried2_fun f xss =
-  mk_tupled_fun (HOLogic.mk_tuple (map HOLogic.mk_tuple xss)) f (flat_rec_arg_args xss);
-
 fun mk_flip (x, Type (_, [T1, Type (_, [T2, T3])])) =
   Abs ("x", T1, Abs ("y", T2, Var (x, T2 --> T1 --> T3) $ Bound 0 $ Bound 1));
 
@@ -350,12 +354,12 @@
 val transfer_gfp_sugar_thms =
   morph_gfp_sugar_thms o Morphism.transfer_morphism o Proof_Context.theory_of;
 
-fun mk_iter_fun_arg_types n ms = map2 dest_tupleT ms o dest_sumTN_balanced n o domain_type;
-
-fun mk_iters_args_types ctr_Tsss Cs ns mss ctor_iter_fun_Tss lthy =
+fun mk_iters_args_types ctr_Tsss Cs absTs repTs ns mss ctor_iter_fun_Tss lthy =
   let
     val Css = map2 replicate ns Cs;
-    val y_Tsss = map3 mk_iter_fun_arg_types ns mss (map un_fold_of ctor_iter_fun_Tss);
+    val y_Tsss = map5 (fn absT => fn repT => fn n => fn ms =>
+        dest_absumprodT absT repT n ms o domain_type)
+      absTs repTs ns mss (map un_fold_of ctor_iter_fun_Tss);
     val g_Tss = map2 (fn C => map (fn y_Ts => y_Ts ---> C)) Cs y_Tsss;
 
     val ((gss, ysss), lthy) =
@@ -367,11 +371,10 @@
     val yssss = map (map (map single)) ysss;
 
     val z_Tssss =
-      map4 (fn n => fn ms => fn ctr_Tss => fn ctor_iter_fun_Ts =>
-          map3 (fn m => fn ctr_Ts => fn ctor_iter_fun_T =>
-              map2 unzip_recT ctr_Ts (dest_tupleT m ctor_iter_fun_T))
-            ms ctr_Tss (dest_sumTN_balanced n (domain_type (co_rec_of ctor_iter_fun_Ts))))
-        ns mss ctr_Tsss ctor_iter_fun_Tss;
+      map6 (fn absT => fn repT => fn n => fn ms => fn ctr_Tss => fn ctor_iter_fun_Ts =>
+          map2 (map2 unzip_recT)
+            ctr_Tss (dest_absumprodT absT repT n ms (domain_type (co_rec_of ctor_iter_fun_Ts))))
+        absTs repTs ns mss ctr_Tsss ctor_iter_fun_Tss;
 
     val z_Tsss' = map (map flat_rec_arg_args) z_Tssss;
     val h_Tss = map2 (map2 (curry (op --->))) z_Tsss' Css;
@@ -390,31 +393,36 @@
 fun repair_nullary_single_ctr [[]] = [[@{typ unit}]]
   | repair_nullary_single_ctr Tss = Tss;
 
-fun mk_coiter_fun_arg_types0 ctr_Tsss Cs ns fun_Ts =
+fun mk_coiter_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss fun_Ts =
   let
     val ctr_Tsss' = map repair_nullary_single_ctr ctr_Tsss;
-    val f_sum_prod_Ts = map range_type fun_Ts;
-    val f_prod_Tss = map2 dest_sumTN_balanced ns f_sum_prod_Ts;
-    val f_Tsss = map2 (map2 (dest_tupleT o length)) ctr_Tsss' f_prod_Tss;
+    val f_absTs = map range_type fun_Ts;
+    val f_Tsss = map repair_nullary_single_ctr (map5 dest_absumprodT absTs repTs ns mss f_absTs);
     val f_Tssss = map3 (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT)))
       Cs ctr_Tsss' f_Tsss;
     val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) f_Tssss;
   in
-    (q_Tssss, f_Tsss, f_Tssss, f_sum_prod_Ts)
+    (q_Tssss, f_Tsss, f_Tssss, f_absTs)
   end;
 
 fun mk_coiter_p_pred_types Cs ns = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs;
 
-fun mk_coiters_args_types ctr_Tsss Cs ns dtor_coiter_fun_Tss lthy =
+fun mk_coiter_fun_arg_types ctr_Tsss Cs absTs repTs ns mss dtor_coiter =
+  (mk_coiter_p_pred_types Cs ns,
+   mk_coiter_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss
+     (binder_fun_types (fastype_of dtor_coiter)));
+
+fun mk_coiters_args_types ctr_Tsss Cs absTs repTs ns mss dtor_coiter_fun_Tss lthy =
   let
     val p_Tss = mk_coiter_p_pred_types Cs ns;
 
     fun mk_types get_Ts =
       let
         val fun_Ts = map get_Ts dtor_coiter_fun_Tss;
-        val (q_Tssss, f_Tsss, f_Tssss, f_sum_prod_Ts) = mk_coiter_fun_arg_types0 ctr_Tsss Cs ns fun_Ts;
+        val (q_Tssss, f_Tsss, f_Tssss, f_repTs) =
+          mk_coiter_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss fun_Ts;
       in
-        (q_Tssss, f_Tsss, f_Tssss, f_sum_prod_Ts)
+        (q_Tssss, f_Tsss, f_Tssss, f_repTs)
       end;
 
     val (r_Tssss, g_Tsss, g_Tssss, unfold_types) = mk_types un_fold_of;
@@ -458,7 +466,7 @@
     ((z, cs, cpss, [(unfold_args, unfold_types), (corec_args, corec_types)]), lthy)
   end;
 
-fun mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy =
+fun mk_co_iters_prelims fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_iterss0 lthy =
   let
     val thy = Proof_Context.theory_of lthy;
 
@@ -469,17 +477,21 @@
 
     val ((iters_args_types, coiters_args_types), lthy') =
       if fp = Least_FP then
-        mk_iters_args_types ctr_Tsss Cs ns mss xtor_co_iter_fun_Tss lthy |>> (rpair NONE o SOME)
+        mk_iters_args_types ctr_Tsss Cs absTs repTs ns mss xtor_co_iter_fun_Tss lthy
+        |>> (rpair NONE o SOME)
       else
-        mk_coiters_args_types ctr_Tsss Cs ns xtor_co_iter_fun_Tss lthy |>> (pair NONE o SOME);
+        mk_coiters_args_types ctr_Tsss Cs absTs repTs ns mss xtor_co_iter_fun_Tss lthy
+        |>> (pair NONE o SOME);
   in
     ((xtor_co_iterss, iters_args_types, coiters_args_types), lthy')
   end;
 
-fun mk_preds_getterss_join c cps sum_prod_T cqfss =
-  let val n = length cqfss in
-    Term.lambda c (mk_IfN sum_prod_T cps
-      (map2 (mk_InN_balanced sum_prod_T n) (map HOLogic.mk_tuple cqfss) (1 upto n)))
+fun mk_preds_getterss_join c cps absT abs cqfss =
+  let
+    val n = length cqfss;
+    val ts = map2 (mk_absumprod absT abs n) (1 upto n) cqfss;
+  in
+    Term.lambda c (mk_IfN absT cps ts)
   end;
 
 fun define_co_iters fp fpT Cs binding_specs lthy0 =
@@ -503,39 +515,43 @@
     ((csts', defs'), lthy')
   end;
 
-fun define_iters iterNs iter_args_typess' mk_binding fpTs Cs ctor_iters lthy =
+fun define_iters iterNs iter_args_typess' mk_binding fpTs Cs reps ctor_iters lthy =
   let
     val nn = length fpTs;
-
-    val Type (_, [fpT, _]) = snd (strip_typeN nn (fastype_of (hd ctor_iters)));
+    val fpT = domain_type (snd (strip_typeN nn (fastype_of (un_fold_of ctor_iters))));
 
     fun generate_iter pre (_, _, fss, xssss) ctor_iter =
-      (mk_binding pre,
-       fold_rev (fold_rev Term.lambda) fss (Term.list_comb (ctor_iter,
-         map2 (mk_case_sumN_balanced oo map2 mk_uncurried2_fun) fss xssss)));
+      let val ctor_iter_absTs = map domain_type (fst (strip_typeN nn (fastype_of ctor_iter))) in
+        (mk_binding pre,
+         fold_rev (fold_rev Term.lambda) fss (Term.list_comb (ctor_iter,
+           map4 (fn ctor_iter_absT => fn rep => fn fs => fn xsss =>
+               mk_case_absumprod ctor_iter_absT rep fs
+                 (map (HOLogic.mk_tuple o map HOLogic.mk_tuple) xsss) (map flat_rec_arg_args xsss))
+             ctor_iter_absTs reps fss xssss)))
+      end;
   in
     define_co_iters Least_FP fpT Cs (map3 generate_iter iterNs iter_args_typess' ctor_iters) lthy
   end;
 
-fun define_coiters coiterNs (_, cs, cpss, coiter_args_typess') mk_binding fpTs Cs dtor_coiters
+fun define_coiters coiterNs (_, cs, cpss, coiter_args_typess') mk_binding fpTs Cs abss dtor_coiters
     lthy =
   let
     val nn = length fpTs;
 
     val Type (_, [_, fpT]) = snd (strip_typeN nn (fastype_of (hd dtor_coiters)));
 
-    fun generate_coiter pre ((pfss, cqfsss), f_sum_prod_Ts) dtor_coiter =
+    fun generate_coiter pre ((pfss, cqfsss), f_absTs) dtor_coiter =
       (mk_binding pre,
        fold_rev (fold_rev Term.lambda) pfss (Term.list_comb (dtor_coiter,
-         map4 mk_preds_getterss_join cs cpss f_sum_prod_Ts cqfsss)));
+         map5 mk_preds_getterss_join cs cpss f_absTs abss cqfsss)));
   in
     define_co_iters Greatest_FP fpT Cs
       (map3 generate_coiter coiterNs coiter_args_typess' dtor_coiters) lthy
   end;
 
 fun derive_induct_iters_thms_for_types pre_bnfs [fold_args_types, rec_args_types] ctor_induct
-    ctor_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss iterss iter_defss
-    lthy =
+    ctor_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss fp_abs_inverses
+    fp_type_definitions abs_inverses ctrss ctr_defss iterss iter_defss lthy =
   let
     val iterss' = transpose iterss;
     val iter_defss' = transpose iter_defss;
@@ -632,12 +648,12 @@
 
         val kksss = map (map (map (fst o snd) o #2)) raw_premss;
 
-        val ctor_induct' = ctor_induct OF (map mk_sumEN_tupled_balanced mss);
+        val ctor_induct' = ctor_induct OF (map2 mk_absumprodE fp_type_definitions mss);
 
         val thm =
           Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-            mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' nested_set_maps
-              pre_set_defss)
+            mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' fp_abs_inverses
+              abs_inverses nested_set_maps pre_set_defss)
           |> singleton (Proof_Context.export names_lthy lthy)
           (* for "datatype_realizer.ML": *)
           |> Thm.name_derivation (fst (dest_Type (hd fpTs)) ^ Long_Name.separator ^
@@ -678,8 +694,8 @@
         val goalss = map5 (map4 o mk_goal fss) fiters xctrss fss xsss fxsss;
 
         val tacss =
-          map2 (map o mk_iter_tac pre_map_defs (nested_map_idents @ nesting_map_idents) iter_defs)
-            ctor_iter_thms ctr_defss;
+          map4 (map ooo mk_iter_tac pre_map_defs (nested_map_idents @ nesting_map_idents) iter_defs)
+            ctor_iter_thms fp_abs_inverses abs_inverses ctr_defss;
 
         fun prove goal tac =
           Goal.prove_sorry lthy [] [] goal (tac o #context)
@@ -698,7 +714,8 @@
 fun derive_coinduct_coiters_thms_for_types pre_bnfs (z, cs, cpss,
       coiters_args_types as [((pgss, crgsss), _), ((phss, cshsss), _)])
     dtor_coinduct dtor_injects dtor_ctors dtor_coiter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss
-    mss ns ctr_defss (ctr_sugars : ctr_sugar list) coiterss coiter_defss export_args lthy =
+    mss ns fp_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list)
+    coiterss coiter_defss export_args lthy =
   let
     fun mk_ctor_dtor_coiter_thm dtor_inject dtor_ctor coiter =
       iffD1 OF [dtor_inject, trans OF [coiter, dtor_ctor RS sym]];
@@ -787,8 +804,8 @@
 
         fun prove dtor_coinduct' goal =
           Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-            mk_coinduct_tac ctxt nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs dtor_ctors
-              exhausts ctr_defss disc_thmsss sel_thmsss)
+            mk_coinduct_tac ctxt nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs fp_abs_inverses
+              abs_inverses dtor_ctors exhausts ctr_defss disc_thmsss sel_thmsss)
           |> singleton (Proof_Context.export names_lthy lthy)
           |> Thm.close_derivation;
 
@@ -801,7 +818,7 @@
         val rel_eqs = map rel_eq_of_bnf pre_bnfs;
         val rel_monos = map rel_mono_of_bnf pre_bnfs;
         val dtor_coinducts =
-          [dtor_coinduct, mk_strong_coinduct_thm dtor_coinduct rel_eqs rel_monos lthy];
+          [dtor_coinduct, mk_strong_coinduct_thm dtor_coinduct rel_eqs rel_monos mk_vimage2p lthy]
       in
         map2 (postproc nn oo prove) dtor_coinducts goals
       end;
@@ -860,11 +877,11 @@
         val corec_goalss = map8 (map4 oooo mk_goal phss) cs cpss hcorecs ns kss ctrss mss cshsss';
 
         val unfold_tacss =
-          map3 (map oo mk_coiter_tac unfold_defs nesting_map_idents)
-            (map un_fold_of ctor_dtor_coiter_thmss) pre_map_defs ctr_defss;
+          map4 (map ooo mk_coiter_tac unfold_defs nesting_map_idents)
+            (map un_fold_of ctor_dtor_coiter_thmss) pre_map_defs abs_inverses ctr_defss;
         val corec_tacss =
-          map3 (map oo mk_coiter_tac corec_defs nesting_map_idents)
-            (map co_rec_of ctor_dtor_coiter_thmss) pre_map_defs ctr_defss;
+          map4 (map ooo mk_coiter_tac corec_defs nesting_map_idents)
+            (map co_rec_of ctor_dtor_coiter_thmss) pre_map_defs abs_inverses ctr_defss;
 
         fun prove goal tac =
           Goal.prove_sorry lthy [] [] goal (tac o #context)
@@ -1048,10 +1065,10 @@
     val unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fake_Ts);
 
     val ctrXs_Tsss = map (map (map freeze_fp)) fake_ctr_Tsss;
-    val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
+    val ctrXs_repTs = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
 
     val fp_eqs =
-      map dest_TFree Xs ~~ map (Term.typ_subst_atomic (As ~~ unsorted_As)) ctrXs_sum_prod_Ts;
+      map dest_TFree Xs ~~ map (Term.typ_subst_atomic (As ~~ unsorted_As)) ctrXs_repTs;
 
     val rhsXs_As' = fold (fold (fold Term.add_tfreesT)) ctrXs_Tsss [];
     val _ = (case subtract (op =) rhsXs_As' As' of [] => ()
@@ -1062,9 +1079,10 @@
       map_filter (fn (A, set_bos) => if exists is_none set_bos then SOME A else NONE)
         (unsorted_As ~~ transpose set_boss);
 
-    val (pre_bnfs, (fp_res as {bnfs = fp_bnfs as any_fp_bnf :: _, ctors = ctors0, dtors = dtors0,
-           xtor_co_iterss = xtor_co_iterss0, xtor_co_induct, dtor_ctors, ctor_dtors, ctor_injects,
-           dtor_injects, xtor_map_thms, xtor_set_thmss, xtor_rel_thms, xtor_co_iter_thmss, ...},
+    val ((pre_bnfs, absT_infos), (fp_res as {bnfs = fp_bnfs as any_fp_bnf :: _, ctors = ctors0,
+             dtors = dtors0, xtor_co_iterss = xtor_co_iterss0, xtor_co_induct, dtor_ctors,
+             ctor_dtors, ctor_injects, dtor_injects, xtor_map_thms, xtor_set_thmss, xtor_rel_thms,
+             xtor_co_iter_thmss, ...},
            lthy)) =
       fp_bnf (construct_fp mixfixes map_bs rel_bs set_bss) fp_bs (map dest_TFree unsorted_As)
         (map dest_TFree killed_As) fp_eqs no_defs_lthy0
@@ -1094,6 +1112,14 @@
                 register_hint ())
           end);
 
+    val abss = map #abs absT_infos;
+    val reps = map #rep absT_infos;
+    val absTs = map #absT absT_infos;
+    val repTs = map #repT absT_infos;
+    val abs_injects = map #abs_inject absT_infos;
+    val abs_inverses = map #abs_inverse absT_infos;
+    val type_definitions = map #type_definition absT_infos;
+
     val time = time lthy;
     val timer = time (Timer.startRealTimer ());
 
@@ -1142,34 +1168,30 @@
     val mss = map (map length) ctr_Tsss;
 
     val ((xtor_co_iterss, iters_args_types, coiters_args_types), lthy') =
-      mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
+      mk_co_iters_prelims fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_iterss0 lthy;
 
-    fun define_ctrs_dtrs_for_type (((((((((((((((((((((((fp_bnf, fp_b), fpT), ctor), dtor),
-            xtor_co_iters), ctor_dtor), dtor_ctor), ctor_inject), pre_map_def), pre_set_defs),
-          pre_rel_def), fp_map_thm), fp_set_thms), fp_rel_thm), n), ks), ms), ctr_bindings),
-        ctr_mixfixes), ctr_Tss), disc_bindings), sel_bindingss), raw_sel_defaultss) no_defs_lthy =
+    fun define_ctrs_dtrs_for_type (((((((((((((((((((((((((((fp_bnf, fp_b), fpT), ctor), dtor),
+              xtor_co_iters), ctor_dtor), dtor_ctor), ctor_inject), pre_map_def), pre_set_defs),
+            pre_rel_def), fp_map_thm), fp_set_thms), fp_rel_thm), n), ks), ms), abs),
+          abs_inject), abs_inverse), type_definition), ctr_bindings), ctr_mixfixes), ctr_Tss),
+        disc_bindings), sel_bindingss), raw_sel_defaultss) no_defs_lthy =
       let
         val fp_b_name = Binding.name_of fp_b;
 
-        val dtorT = domain_type (fastype_of ctor);
-        val ctr_prod_Ts = map HOLogic.mk_tupleT ctr_Tss;
-        val ctr_sum_prod_T = mk_sumTN_balanced ctr_prod_Ts;
+        val ctr_absT = domain_type (fastype_of ctor);
 
         val ((((w, xss), yss), u'), names_lthy) =
           no_defs_lthy
-          |> yield_singleton (mk_Frees "w") dtorT
+          |> yield_singleton (mk_Frees "w") ctr_absT
           ||>> mk_Freess "x" ctr_Tss
           ||>> mk_Freess "y" (map (map B_ify) ctr_Tss)
           ||>> yield_singleton Variable.variant_fixes fp_b_name;
 
         val u = Free (u', fpT);
 
-        val tuple_xs = map HOLogic.mk_tuple xss;
-        val tuple_ys = map HOLogic.mk_tuple yss;
-
         val ctr_rhss =
-          map3 (fn k => fn xs => fn tuple_x => fold_rev Term.lambda xs (ctor $
-            mk_InN_balanced ctr_sum_prod_T n tuple_x k)) ks xss tuple_xs;
+          map2 (fn k => fn xs => fold_rev Term.lambda xs (ctor $ mk_absumprod ctr_absT abs n k xs))
+            ks xss;
 
         val maybe_conceal_def_binding = Thm.def_binding
           #> Config.get no_defs_lthy bnf_note_all = false ? Binding.conceal;
@@ -1200,28 +1222,27 @@
                         (mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w)));
                   in
                     Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-                      mk_ctor_iff_dtor_tac ctxt (map (SOME o certifyT lthy) [dtorT, fpT])
+                      mk_ctor_iff_dtor_tac ctxt (map (SOME o certifyT lthy) [ctr_absT, fpT])
                         (certify lthy ctor) (certify lthy dtor) ctor_dtor dtor_ctor)
                     |> Morphism.thm phi
                     |> Thm.close_derivation
                   end;
 
                 val sumEN_thm' =
-                  unfold_thms lthy @{thms unit_all_eq1}
-                    (Drule.instantiate' (map (SOME o certifyT lthy) ctr_prod_Ts) []
-                       (mk_sumEN_balanced n))
+                  unfold_thms lthy @{thms unit_all_eq1} (mk_absumprodE type_definition ms)
                   |> Morphism.thm phi;
               in
                 mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm'
               end;
 
             val inject_tacss =
-              map2 (fn 0 => K [] | _ => fn ctr_def => [fn {context = ctxt, ...} =>
-                mk_inject_tac ctxt ctr_def ctor_inject]) ms ctr_defs;
+              map2 (fn ctr_def => fn 0 => [] | _ => [fn {context = ctxt, ...} =>
+                mk_inject_tac ctxt ctr_def ctor_inject abs_inject]) ctr_defs ms;
 
             val half_distinct_tacss =
               map (map (fn (def, def') => fn {context = ctxt, ...} =>
-                mk_half_distinct_tac ctxt ctor_inject [def, def'])) (mk_half_pairss (`I ctr_defs));
+                  mk_half_distinct_tac ctxt ctor_inject abs_inject [def, def']))
+                (mk_half_pairss (`I ctr_defs));
 
             val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss;
 
@@ -1249,17 +1270,21 @@
                     cterm_instantiate_pos [NONE, NONE, SOME (certify lthy ctor')] arg_cong
                   end;
 
-              fun mk_cIn ify =
-                certify lthy o (fp = Greatest_FP ? curry (op $) (map_types ify ctor)) oo
-                mk_InN_balanced (ify ctr_sum_prod_T) n;
+              fun mk_cIn ctor k xs =
+                let val absT = domain_type (fastype_of ctor) in
+                  mk_absumprod absT abs n k xs
+                  |> fp = Greatest_FP ? curry (op $) ctor
+                  |> certify lthy
+                end;
 
-              val cxIns = map2 (mk_cIn I) tuple_xs ks;
-              val cyIns = map2 (mk_cIn B_ify) tuple_ys ks;
+              val cxIns = map2 (mk_cIn ctor) ks xss;
+              val cyIns = map2 (mk_cIn (map_types B_ify ctor)) ks yss;
 
               fun mk_map_thm ctr_def' cxIn =
                 fold_thms lthy [ctr_def']
-                  (unfold_thms lthy (pre_map_def ::
-                       (if fp = Least_FP then [] else [ctor_dtor, dtor_ctor]) @ sum_prod_thms_map)
+                  (unfold_thms lthy (o_apply :: pre_map_def ::
+                       (if fp = Least_FP then [] else [ctor_dtor, dtor_ctor]) @ sum_prod_thms_map @
+                       abs_inverses)
                      (cterm_instantiate_pos (nones @ [SOME cxIn])
                         (if fp = Least_FP then fp_map_thm else fp_map_thm RS ctor_cong)))
                 |> singleton (Proof_Context.export names_lthy no_defs_lthy);
@@ -1267,7 +1292,8 @@
               fun mk_set_thm fp_set_thm ctr_def' cxIn =
                 fold_thms lthy [ctr_def']
                   (unfold_thms lthy (pre_set_defs @ nested_set_maps @ nesting_set_maps @
-                       (if fp = Least_FP then [] else [dtor_ctor]) @ sum_prod_thms_set)
+                       (if fp = Least_FP then [] else [dtor_ctor]) @ sum_prod_thms_set @
+                       abs_inverses)
                      (cterm_instantiate_pos [SOME cxIn] fp_set_thm))
                 |> singleton (Proof_Context.export names_lthy no_defs_lthy);
 
@@ -1281,8 +1307,9 @@
 
               fun mk_rel_thm postproc ctr_defs' cxIn cyIn =
                 fold_thms lthy ctr_defs'
-                  (unfold_thms lthy (@{thm Inl_Inr_False} :: pre_rel_def ::
-                       (if fp = Least_FP then [] else [dtor_ctor]) @ sum_prod_thms_rel)
+                  (unfold_thms lthy (pre_rel_def :: abs_inverse ::
+                       (if fp = Least_FP then [] else [dtor_ctor]) @ sum_prod_thms_rel @
+                       @{thms vimage2p_def Inl_Inr_False})
                      (cterm_instantiate_pos (nones @ [SOME cxIn, SOME cyIn]) fp_rel_thm))
                 |> postproc
                 |> singleton (Proof_Context.export names_lthy no_defs_lthy);
@@ -1341,9 +1368,11 @@
         (wrap_ctrs
          #> derive_maps_sets_rels
          ##>>
-           (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
-            else define_coiters [unfoldN, corecN] (the coiters_args_types))
-             mk_binding fpTs Cs xtor_co_iters
+           (if fp = Least_FP then
+             define_iters [foldN, recN] (the iters_args_types) mk_binding fpTs Cs reps xtor_co_iters
+           else
+             define_coiters [unfoldN, corecN] (the coiters_args_types) mk_binding fpTs Cs abss
+               xtor_co_iters)
          #> massage_res, lthy')
       end;
 
@@ -1363,8 +1392,8 @@
       let
         val ((induct_thms, induct_thm, induct_attrs), (fold_thmss, rec_thmss, iter_attrs)) =
           derive_induct_iters_thms_for_types pre_bnfs (the iters_args_types) xtor_co_induct
-            xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss iterss
-            iter_defss lthy;
+            xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss abs_inverses
+            type_definitions abs_inverses ctrss ctr_defss iterss iter_defss lthy;
 
         val induct_type_attr = Attrib.internal o K o Induct.induct_type;
 
@@ -1386,8 +1415,8 @@
         |> Spec_Rules.add Spec_Rules.Equational (map un_fold_of iterss, flat fold_thmss)
         |> Spec_Rules.add Spec_Rules.Equational (map co_rec_of iterss, flat rec_thmss)
         |> Local_Theory.notes (common_notes @ notes) |> snd
-        |> register_fp_sugars Xs Least_FP pre_bnfs nested_bnfs nesting_bnfs fp_res ctrXs_Tsss
-          ctr_defss ctr_sugars iterss mapss [induct_thm] (map single induct_thms)
+        |> register_fp_sugars Xs Least_FP pre_bnfs absT_infos nested_bnfs nesting_bnfs fp_res
+          ctrXs_Tsss ctr_defss ctr_sugars iterss mapss [induct_thm] (map single induct_thms)
           (transpose [fold_thmss, rec_thmss]) (replicate nn []) (replicate nn [])
       end;
 
@@ -1403,8 +1432,8 @@
              (sel_unfold_thmsss, sel_corec_thmsss, sel_coiter_attrs)) =
           derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
             dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns
-            ctr_defss ctr_sugars coiterss coiter_defss (Proof_Context.export lthy' no_defs_lthy)
-            lthy;
+            abs_inverses abs_inverses I ctr_defss ctr_sugars coiterss coiter_defss
+            (Proof_Context.export lthy' no_defs_lthy) lthy;
 
         val sel_unfold_thmss = map flat sel_unfold_thmsss;
         val sel_corec_thmss = map flat sel_corec_thmsss;
@@ -1451,8 +1480,8 @@
         |> add_spec_rules un_fold_of sel_unfold_thmss unfold_thmss
         |> add_spec_rules co_rec_of sel_corec_thmss corec_thmss
         |> Local_Theory.notes (common_notes @ notes) |> snd
-        |> register_fp_sugars Xs Greatest_FP pre_bnfs nested_bnfs nesting_bnfs fp_res ctrXs_Tsss
-          ctr_defss ctr_sugars coiterss mapss [coinduct_thm, strong_coinduct_thm]
+        |> register_fp_sugars Xs Greatest_FP pre_bnfs absT_infos nested_bnfs nesting_bnfs fp_res
+          ctrXs_Tsss ctr_defss ctr_sugars coiterss mapss [coinduct_thm, strong_coinduct_thm]
           (transpose [coinduct_thms, strong_coinduct_thms])
           (transpose [unfold_thmss, corec_thmss]) (transpose [disc_unfold_thmss, disc_corec_thmss])
           (transpose [sel_unfold_thmsss, sel_corec_thmsss])
@@ -1462,8 +1491,9 @@
       |> fold_map define_ctrs_dtrs_for_type (fp_bnfs ~~ fp_bs ~~ fpTs ~~ ctors ~~ dtors ~~
         xtor_co_iterss ~~ ctor_dtors ~~ dtor_ctors ~~ ctor_injects ~~ pre_map_defs ~~
         pre_set_defss ~~ pre_rel_defs ~~ xtor_map_thms ~~ xtor_set_thmss ~~ xtor_rel_thms ~~ ns ~~
-        kss ~~ mss ~~ ctr_bindingss ~~ ctr_mixfixess ~~ ctr_Tsss ~~ disc_bindingss ~~
-        sel_bindingsss ~~ raw_sel_defaultsss)
+        kss ~~ mss ~~ abss ~~ abs_injects ~~ abs_inverses ~~ type_definitions ~~
+        ctr_bindingss ~~ ctr_mixfixess ~~ ctr_Tsss ~~ disc_bindingss ~~ sel_bindingsss ~~
+        raw_sel_defaultsss)
       |> wrap_types_etc
       |> fp_case fp derive_note_induct_iters_thms_for_types
            derive_note_coinduct_coiters_thms_for_types;
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML	Tue Feb 25 18:14:26 2014 +0100
@@ -12,17 +12,19 @@
   val sum_prod_thms_rel: thm list
 
   val mk_coinduct_tac: Proof.context -> thm list -> int -> int list -> thm -> thm list ->
-    thm list -> thm list -> thm list list -> thm list list list -> thm list list list -> tactic
-  val mk_coiter_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
+    thm list -> thm list -> thm list -> thm list -> thm list list -> thm list list list ->
+    thm list list list -> tactic
+  val mk_coiter_tac: thm list -> thm list -> thm -> thm -> thm -> thm -> Proof.context -> tactic
   val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
     tactic
   val mk_disc_coiter_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
   val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
-  val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
+  val mk_half_distinct_tac: Proof.context -> thm -> thm -> thm list -> tactic
   val mk_induct_tac: Proof.context -> int -> int list -> int list list -> int list list list ->
-    thm list -> thm -> thm list -> thm list list -> tactic
-  val mk_inject_tac: Proof.context -> thm -> thm -> tactic
-  val mk_iter_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
+    thm list -> thm -> thm list -> thm list -> thm list -> thm list list -> tactic
+  val mk_inject_tac: Proof.context -> thm -> thm -> thm -> tactic
+  val mk_iter_tac: thm list -> thm list -> thm list -> thm -> thm -> thm -> thm -> Proof.context ->
+    tactic
 end;
 
 structure BNF_FP_Def_Sugar_Tactics : BNF_FP_DEF_SUGAR_TACTICS =
@@ -75,31 +77,35 @@
       SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
       atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor]));
 
-fun mk_half_distinct_tac ctxt ctor_inject ctr_defs =
-  unfold_thms_tac ctxt (ctor_inject :: @{thms sum.inject} @ ctr_defs) THEN
+fun mk_half_distinct_tac ctxt ctor_inject abs_inject ctr_defs =
+  unfold_thms_tac ctxt (ctor_inject :: abs_inject :: @{thms sum.inject} @ ctr_defs) THEN
   HEADGOAL (rtac @{thm sum.distinct(1)});
 
-fun mk_inject_tac ctxt ctr_def ctor_inject =
-  unfold_thms_tac ctxt [ctr_def] THEN HEADGOAL (rtac (ctor_inject RS ssubst)) THEN
-  unfold_thms_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN HEADGOAL (rtac refl);
+fun mk_inject_tac ctxt ctr_def ctor_inject abs_inject =
+  unfold_thms_tac ctxt [ctr_def] THEN
+  HEADGOAL (rtac (ctor_inject RS ssubst)) THEN
+  unfold_thms_tac ctxt (abs_inject :: @{thms sum.inject Pair_eq conj_assoc}) THEN
+  HEADGOAL (rtac refl);
 
 val iter_unfold_thms =
   @{thms comp_def convol_def fst_conv id_def case_prod_Pair_iden snd_conv split_conv
       case_unit_Unity} @ sum_prod_thms_map;
 
-fun mk_iter_tac pre_map_defs map_idents iter_defs ctor_iter ctr_def ctxt =
-  unfold_thms_tac ctxt (ctr_def :: ctor_iter :: iter_defs @ pre_map_defs @ map_idents @
-    iter_unfold_thms) THEN HEADGOAL (rtac refl);
+fun mk_iter_tac pre_map_defs map_idents iter_defs ctor_iter fp_abs_inverse abs_inverse ctr_def ctxt =
+  unfold_thms_tac ctxt (ctr_def :: ctor_iter :: fp_abs_inverse :: abs_inverse :: iter_defs @
+    pre_map_defs @ map_idents @ iter_unfold_thms) THEN HEADGOAL (rtac refl);
 
 val coiter_unfold_thms = @{thms id_def} @ sum_prod_thms_map;
-val ss_if_True_False = simpset_of (ss_only @{thms if_True if_False} @{context});
 
-fun mk_coiter_tac coiter_defs map_idents ctor_dtor_coiter pre_map_def ctr_def ctxt =
-  unfold_thms_tac ctxt (ctr_def :: coiter_defs) THEN
-  HEADGOAL (rtac (ctor_dtor_coiter RS trans) THEN'
-    asm_simp_tac (put_simpset ss_if_True_False ctxt)) THEN_MAYBE
-  (unfold_thms_tac ctxt (pre_map_def :: map_idents @ coiter_unfold_thms) THEN
-   HEADGOAL (rtac refl ORELSE' rtac (@{thm unit_eq} RS arg_cong)));
+fun mk_coiter_tac coiter_defs map_idents ctor_dtor_coiter pre_map_def abs_inverse ctr_def ctxt =
+  let
+    val ss = ss_only (pre_map_def :: abs_inverse :: map_idents @ coiter_unfold_thms @
+      @{thms o_apply vimage2p_def if_True if_False}) ctxt;
+  in
+    unfold_thms_tac ctxt (ctr_def :: coiter_defs) THEN
+    HEADGOAL (rtac (ctor_dtor_coiter RS trans) THEN' asm_simp_tac ss) THEN_MAYBE
+    HEADGOAL (rtac refl ORELSE' rtac (@{thm unit_eq} RS arg_cong))
+  end;
 
 fun mk_disc_coiter_iff_tac case_splits' coiters discs ctxt =
   EVERY (map3 (fn case_split_tac => fn coiter_thm => fn disc =>
@@ -113,35 +119,43 @@
     hyp_subst_tac ctxt ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN'
   (rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI});
 
-fun mk_induct_leverage_prem_prems_tac ctxt nn kks set_maps pre_set_defs =
+fun mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps
+    pre_set_defs =
   HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp,
-    SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ set_maps @ sum_prod_thms_set0)),
+    SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ fp_abs_inverses @ abs_inverses @ set_maps @
+      sum_prod_thms_set0)),
     solve_prem_prem_tac ctxt]) (rev kks)));
 
-fun mk_induct_discharge_prem_tac ctxt nn n set_maps pre_set_defs m k kks =
+fun mk_induct_discharge_prem_tac ctxt nn n fp_abs_inverses abs_inverses set_maps pre_set_defs m k
+    kks =
   let val r = length kks in
     HEADGOAL (EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac ctxt,
       REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)]) THEN
     EVERY [REPEAT_DETERM_N r
         (HEADGOAL (rotate_tac ~1 THEN' dtac meta_mp THEN' rotate_tac 1) THEN prefer_tac 2),
       if r > 0 then ALLGOALS (Goal.norm_hhf_tac ctxt) else all_tac, HEADGOAL atac,
-      mk_induct_leverage_prem_prems_tac ctxt nn kks set_maps pre_set_defs]
+      mk_induct_leverage_prem_prems_tac ctxt nn kks fp_abs_inverses abs_inverses set_maps
+        pre_set_defs]
   end;
 
-fun mk_induct_tac ctxt nn ns mss kkss ctr_defs ctor_induct' set_maps pre_set_defss =
+fun mk_induct_tac ctxt nn ns mss kkss ctr_defs ctor_induct' fp_abs_inverses abs_inverses set_maps
+    pre_set_defss =
   let val n = Integer.sum ns in
     unfold_thms_tac ctxt ctr_defs THEN HEADGOAL (rtac ctor_induct') THEN
     co_induct_inst_as_projs_tac ctxt 0 THEN
-    EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_maps) pre_set_defss
-      mss (unflat mss (1 upto n)) kkss)
+    EVERY (map4 (EVERY oooo map3 o
+        mk_induct_discharge_prem_tac ctxt nn n fp_abs_inverses abs_inverses set_maps)
+      pre_set_defss mss (unflat mss (1 upto n)) kkss)
   end;
 
-fun mk_coinduct_same_ctr_tac ctxt rel_eqs pre_rel_def dtor_ctor ctr_def discs sels =
+fun mk_coinduct_same_ctr_tac ctxt rel_eqs pre_rel_def fp_abs_inverse abs_inverse dtor_ctor ctr_def
+    discs sels =
   hyp_subst_tac ctxt THEN'
   CONVERSION (hhf_concl_conv
     (Conv.top_conv (K (Conv.try_conv (Conv.rewr_conv ctr_def))) ctxt) ctxt) THEN'
   SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels)) THEN'
-  SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels @ sum_prod_thms_rel)) THEN'
+  SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: fp_abs_inverse :: abs_inverse :: dtor_ctor ::
+    sels @ sum_prod_thms_rel @ @{thms o_apply vimage2p_def})) THEN'
   (atac ORELSE' REPEAT o etac conjE THEN'
      full_simp_tac
        (ss_only (@{thm prod.inject} :: no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN'
@@ -157,8 +171,8 @@
     full_simp_tac (ss_only (refl :: no_refl discs'' @ basic_simp_thms) ctxt)
   end;
 
-fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def dtor_ctor exhaust ctr_defs
-    discss selss =
+fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def fp_abs_inverse abs_inverse
+    dtor_ctor exhaust ctr_defs discss selss =
   let val ks = 1 upto n in
     EVERY' ([rtac allI, rtac allI, rtac impI, select_prem_tac nn (dtac meta_spec) kk,
         dtac meta_spec, dtac meta_mp, atac, rtac exhaust, K (co_induct_inst_as_projs_tac ctxt 0),
@@ -167,15 +181,17 @@
         EVERY' ([rtac exhaust, K (co_induct_inst_as_projs_tac ctxt 1)] @
           map2 (fn k' => fn discs' =>
             if k' = k then
-              mk_coinduct_same_ctr_tac ctxt rel_eqs' pre_rel_def dtor_ctor ctr_def discs sels
+              mk_coinduct_same_ctr_tac ctxt rel_eqs' pre_rel_def fp_abs_inverse abs_inverse
+                dtor_ctor ctr_def discs sels
             else
               mk_coinduct_distinct_ctrs_tac ctxt discs discs') ks discss)) ks ctr_defs discss selss)
   end;
 
-fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs dtor_ctors exhausts ctr_defss
-    discsss selsss =
+fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs fp_abs_inverses abs_inverses
+    dtor_ctors exhausts ctr_defss discsss selsss =
   HEADGOAL (rtac dtor_coinduct' THEN'
-    EVERY' (map8 (mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn)
-      (1 upto nn) ns pre_rel_defs dtor_ctors exhausts ctr_defss discsss selsss));
+    EVERY' (map10 (mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn)
+      (1 upto nn) ns pre_rel_defs fp_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss discsss
+      selsss));
 
 end;
--- a/src/HOL/Tools/BNF/bnf_fp_n2m.ML	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m.ML	Tue Feb 25 18:14:26 2014 +0100
@@ -9,7 +9,7 @@
 sig
   val construct_mutualized_fp: BNF_Util.fp_kind -> typ list -> BNF_FP_Def_Sugar.fp_sugar list ->
     binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
-    local_theory -> BNF_FP_Util.fp_result * local_theory
+    BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory
 end;
 
 structure BNF_FP_N2M : BNF_FP_N2M =
@@ -17,6 +17,7 @@
 
 open BNF_Def
 open BNF_Util
+open BNF_Comp
 open BNF_FP_Util
 open BNF_FP_Def_Sugar
 open BNF_Tactics
@@ -45,28 +46,16 @@
     Const (@{const_name sum_map}, fT --> gT --> mk_sumT (fAT, gAT) --> mk_sumT (fBT, gBT)) $ f $ g
   end;
 
-fun construct_mutualized_fp fp fpTs (fp_sugars : fp_sugar list) bs resBs (resDs, Dss) bnfs lthy =
+fun construct_mutualized_fp fp fpTs (fp_sugars : fp_sugar list) bs resBs (resDs, Dss) bnfs
+    absT_infos lthy =
   let
-    fun steal_fp_res get =
+    fun of_fp_res get =
       map (fn {fp_res, fp_res_index, ...} => nth (get fp_res) fp_res_index) fp_sugars;
 
-    val n = length bnfs;
-    val deads = fold (union (op =)) Dss resDs;
-    val As = subtract (op =) deads (map TFree resBs);
-    val names_lthy = fold Variable.declare_typ (As @ deads) lthy;
-    val m = length As;
-    val live = m + n;
-    val ((Xs, Bs), names_lthy) = names_lthy
-      |> mk_TFrees n
-      ||>> mk_TFrees m;
-    val allAs = As @ Xs;
-    val phiTs = map2 mk_pred2T As Bs;
-    val theta = As ~~ Bs;
-    val fpTs' = map (Term.typ_subst_atomic theta) fpTs;
-    val pre_phiTs = map2 mk_pred2T fpTs fpTs';
-
     fun mk_co_algT T U = fp_case fp (T --> U) (U --> T);
     fun co_swap pair = fp_case fp I swap pair;
+    val mk_co_comp = HOLogic.mk_comp o co_swap;
+
     val dest_co_algT = co_swap o dest_funT;
     val co_alg_argT = fp_case fp range_type domain_type;
     val co_alg_funT = fp_case fp domain_type range_type;
@@ -75,30 +64,78 @@
     val co_proj1_const = fp_case fp (fst_const o fst) (uncurry Inl_const o dest_sumT o snd);
     val mk_co_productT = curry (fp_case fp HOLogic.mk_prodT mk_sumT);
     val dest_co_productT = fp_case fp HOLogic.dest_prodT dest_sumT;
+    val rewrite_comp_comp = fp_case fp @{thm rewriteL_comp_comp} @{thm rewriteR_comp_comp};
+
+    val fp_absT_infos = map #absT_info fp_sugars;
+    val fp_bnfs = of_fp_res #bnfs;
+    val pre_bnfs = map #pre_bnf fp_sugars;
+    val nesty_bnfss = map (fn sugar => #nested_bnfs sugar @ #nesting_bnfs sugar) fp_sugars;
+    val fp_nesty_bnfss = fp_bnfs :: nesty_bnfss;
+    val fp_nesty_bnfs = distinct (op = o pairself T_of_bnf) (flat fp_nesty_bnfss);
+
+    val fp_absTs = map #absT fp_absT_infos;
+    val fp_repTs = map #repT fp_absT_infos;
+    val fp_abss = map #abs fp_absT_infos;
+    val fp_reps = map #rep fp_absT_infos;
+    val fp_type_definitions = map #type_definition fp_absT_infos;
+
+    val absTs = map #absT absT_infos;
+    val repTs = map #repT absT_infos;
+    val absTs' = map (Logic.type_map (singleton (Variable.polymorphic lthy))) absTs;
+    val repTs' = map (Logic.type_map (singleton (Variable.polymorphic lthy))) repTs;
+    val abss = map #abs absT_infos;
+    val reps = map #rep absT_infos;
+    val abs_inverses = map #abs_inverse absT_infos;
+    val type_definitions = map #type_definition absT_infos;
+
+    val n = length bnfs;
+    val deads = fold (union (op =)) Dss resDs;
+    val As = subtract (op =) deads (map TFree resBs);
+    val names_lthy = fold Variable.declare_typ (As @ deads) lthy;
+    val m = length As;
+    val live = m + n;
+
+    val ((Xs, Bs), names_lthy) = names_lthy
+      |> mk_TFrees n
+      ||>> mk_TFrees m;
+
+    val allAs = As @ Xs;
+    val allBs = Bs @ Xs;
+    val phiTs = map2 mk_pred2T As Bs;
+    val thetaBs = As ~~ Bs;
+    val fpTs' = map (Term.typ_subst_atomic thetaBs) fpTs;
+    val fold_thetaAs = Xs ~~ fpTs;
+    val fold_thetaBs = Xs ~~ fpTs';
+    val rec_theta = Xs ~~ map2 mk_co_productT fpTs Xs;
+    val pre_phiTs = map2 mk_pred2T fpTs fpTs';
 
     val ((ctors, dtors), (xtor's, xtors)) =
       let
-        val ctors = map2 (force_typ names_lthy o (fn T => dummyT --> T)) fpTs (steal_fp_res #ctors);
-        val dtors = map2 (force_typ names_lthy o (fn T => T --> dummyT)) fpTs (steal_fp_res #dtors);
+        val ctors = map2 (force_typ names_lthy o (fn T => dummyT --> T)) fpTs (of_fp_res #ctors);
+        val dtors = map2 (force_typ names_lthy o (fn T => T --> dummyT)) fpTs (of_fp_res #dtors);
       in
-        ((ctors, dtors), `(map (Term.subst_atomic_types theta)) (fp_case fp ctors dtors))
+        ((ctors, dtors), `(map (Term.subst_atomic_types thetaBs)) (fp_case fp ctors dtors))
       end;
 
+    val absATs = map (domain_type o fastype_of) ctors;
+    val absBTs = map (Term.typ_subst_atomic thetaBs) absATs;
     val xTs = map (domain_type o fastype_of) xtors;
     val yTs = map (domain_type o fastype_of) xtor's;
 
+    fun abs_of allAs Ds bnf = mk_abs (mk_T_of_bnf Ds allAs bnf) o #abs;
+    fun rep_of absAT = mk_rep absAT o #rep;
+
+    val absAs = map3 (abs_of allAs) Dss bnfs absT_infos;
+    val absBs = map3 (abs_of allBs) Dss bnfs absT_infos;
+    val fp_repAs = map2 rep_of absATs fp_absT_infos;
+    val fp_repBs = map2 rep_of absBTs fp_absT_infos;
+
     val (((((phis, phis'), pre_phis), xs), ys), names_lthy) = names_lthy
       |> mk_Frees' "R" phiTs
       ||>> mk_Frees "S" pre_phiTs
       ||>> mk_Frees "x" xTs
       ||>> mk_Frees "y" yTs;
 
-    val fp_bnfs = steal_fp_res #bnfs;
-    val pre_bnfs = map #pre_bnf fp_sugars;
-    val nesty_bnfss = map (fn sugar => #nested_bnfs sugar @ #nesting_bnfs sugar) fp_sugars;
-    val fp_nesty_bnfss = fp_bnfs :: nesty_bnfss;
-    val fp_nesty_bnfs = distinct (op = o pairself T_of_bnf) (flat fp_nesty_bnfss);
-
     val rels =
       let
         fun find_rel T As Bs = fp_nesty_bnfss
@@ -127,15 +164,28 @@
     val pre_rels = map2 (fn Ds => mk_rel_of_bnf Ds (As @ fpTs) (Bs @ fpTs')) Dss bnfs;
 
     val rel_unfolds = maps (no_refl o single o rel_def_of_bnf) pre_bnfs;
-    val rel_xtor_co_inducts = steal_fp_res (split_conj_thm o #rel_xtor_co_induct_thm)
+    val rel_xtor_co_inducts = of_fp_res (split_conj_thm o #rel_xtor_co_induct_thm)
       |> map (unfold_thms lthy (id_apply :: rel_unfolds));
 
     val rel_defs = map rel_def_of_bnf bnfs;
     val rel_monos = map rel_mono_of_bnf bnfs;
 
+    fun cast castA castB pre_rel =
+      let
+        val castAB = mk_vimage2p (Term.subst_atomic_types fold_thetaAs castA)
+          (Term.subst_atomic_types fold_thetaBs castB);
+      in
+        fold_rev (fold_rev Term.absdummy) [phiTs, pre_phiTs]
+          (castAB $ Term.list_comb (pre_rel, map Bound (live - 1 downto 0)))
+      end;
+
+    val castAs = map2 (curry HOLogic.mk_comp) absAs fp_repAs;
+    val castBs = map2 (curry HOLogic.mk_comp) absBs fp_repBs;
+
     val rel_xtor_co_induct_thm =
-      mk_rel_xtor_co_induct_thm fp pre_rels pre_phis rels phis xs ys xtors xtor's
-        (mk_rel_xtor_co_induct_tactic fp rel_xtor_co_inducts rel_defs rel_monos) lthy;
+      mk_rel_xtor_co_induct_thm fp (map3 cast castAs castBs pre_rels) pre_phis rels phis xs ys xtors
+        xtor's (mk_rel_xtor_co_induct_tactic fp abs_inverses rel_xtor_co_inducts rel_defs rel_monos)
+        lthy;
 
     val rel_eqs = no_refl (map rel_eq_of_bnf fp_nesty_bnfs);
     val map_id0s = no_refl (map map_id0_of_bnf bnfs);
@@ -150,16 +200,22 @@
               let val T = domain_type (fastype_of P);
               in mk_Grp (HOLogic.Collect_const T $ P) (HOLogic.id_const T) end;
             val cts = map (SOME o certify lthy) (map HOLogic.eq_const As @ map mk_Grp_id Ps);
+            fun mk_fp_type_copy_thms thm = map (curry op RS thm)
+              @{thms type_copy_Abs_o_Rep type_copy_vimage2p_Grp_Rep};
+            fun mk_type_copy_thms thm = map (curry op RS thm)
+              @{thms type_copy_Rep_o_Abs type_copy_vimage2p_Grp_Abs};
           in
             cterm_instantiate_pos cts rel_xtor_co_induct_thm
             |> singleton (Proof_Context.export names_lthy lthy)
             |> unfold_thms lthy (@{thms eq_le_Grp_id_iff all_simps(1,2)[symmetric]} @ rel_eqs)
             |> funpow n (fn thm => thm RS spec)
             |> unfold_thms lthy (@{thm eq_alt} :: map rel_Grp_of_bnf bnfs @ map_id0s)
-            |> unfold_thms lthy @{thms Grp_id_mono_subst eqTrueI[OF subset_UNIV] simp_thms(22)}
+            |> unfold_thms lthy (@{thms vimage2p_comp comp_apply comp_id
+               Grp_id_mono_subst eqTrueI[OF subset_UNIV] simp_thms(22)} @
+               maps mk_fp_type_copy_thms fp_type_definitions @
+               maps mk_type_copy_thms type_definitions)
             |> unfold_thms lthy @{thms subset_iff mem_Collect_eq
                atomize_conjL[symmetric] atomize_all[symmetric] atomize_imp[symmetric]}
-            |> unfold_thms lthy (maps set_defs_of_bnf bnfs)
           end
       | Greatest_FP =>
           let
@@ -173,8 +229,6 @@
           end);
 
     val fold_preTs = map2 (fn Ds => mk_T_of_bnf Ds allAs) Dss bnfs;
-    val fold_pre_deads_only_Ts = map2 (fn Ds => mk_T_of_bnf Ds (replicate live dummyT)) Dss bnfs;
-    val rec_theta = Xs ~~ map2 mk_co_productT fpTs Xs;
     val rec_preTs = map (Term.typ_subst_atomic rec_theta) fold_preTs;
 
     val fold_strTs = map2 mk_co_algT fold_preTs Xs;
@@ -185,32 +239,52 @@
       |> mk_Frees' "s" fold_strTs
       ||>> mk_Frees' "s" rec_strTs;
 
-    val co_iters = steal_fp_res #xtor_co_iterss;
+    val co_iters = of_fp_res #xtor_co_iterss;
     val ns = map (length o #Ts o #fp_res) fp_sugars;
 
     fun substT rho (Type (@{type_name "fun"}, [T, U])) = substT rho T --> substT rho U
       | substT rho (Type (s, Ts)) = Type (s, map (typ_subst_nonatomic rho) Ts)
       | substT _ T = T;
 
+    val typ_subst_nonatomic_sorted = fold_rev (typ_subst_nonatomic o single);
+
     fun force_iter is_rec i TU TU_rec raw_iters =
       let
+        val thy = Proof_Context.theory_of lthy;
+
         val approx_fold = un_fold_of raw_iters
           |> force_typ names_lthy
             (replicate (nth ns i) dummyT ---> (if is_rec then TU_rec else TU));
-        val subst = Term.typ_subst_atomic (Xs ~~ fpTs);
+        val subst = Term.typ_subst_atomic fold_thetaAs;
+
+        fun mk_fp_absT_repT fp_repT fp_absT = mk_absT thy fp_repT fp_absT ooo mk_repT;
+        val mk_fp_absT_repTs = map5 mk_fp_absT_repT fp_repTs fp_absTs absTs repTs;
+
+        val fold_preTs' = mk_fp_absT_repTs (map subst fold_preTs);
+
+        val fold_pre_deads_only_Ts =
+          map (typ_subst_nonatomic_sorted (map (rpair dummyT) (As @ fpTs))) fold_preTs';
+
         val TUs = map_split dest_co_algT (binder_fun_types (fastype_of approx_fold))
           |>> map subst
           |> uncurry (map2 mk_co_algT);
-        val js = find_indices Type.could_unify TUs
-          (map2 (fn T => fn U => mk_co_algT (subst T) U) fold_preTs Xs);
+        val cands = map2 mk_co_algT fold_preTs' Xs;
+
+        val js = find_indices Type.could_unify TUs cands;
         val Tpats = map (fn j => mk_co_algT (nth fold_pre_deads_only_Ts j) (nth Xs j)) js;
         val iter = raw_iters |> (if is_rec then co_rec_of else un_fold_of);
       in
         force_typ names_lthy (Tpats ---> TU) iter
       end;
 
+    fun mk_co_comp_abs_rep fp_absT absT fp_abs fp_rep abs rep t =
+      fp_case fp (HOLogic.mk_comp (HOLogic.mk_comp (t, mk_abs absT abs), mk_rep fp_absT fp_rep))
+        (HOLogic.mk_comp (mk_abs fp_absT fp_abs, HOLogic.mk_comp (mk_rep absT rep, t)));
+
     fun mk_iter b_opt is_rec iters lthy TU =
       let
+        val thy = Proof_Context.theory_of lthy;
+
         val x = co_alg_argT TU;
         val i = find_index (fn T => x = T) Xs;
         val TUiter =
@@ -220,20 +294,32 @@
                 (TU |> (is_none b_opt andalso not is_rec) ? substT (fpTs ~~ Xs))
                 (TU |> (is_none b_opt) ? substT (map2 mk_co_productT fpTs Xs ~~ Xs))
           | SOME f => f);
+
         val TUs = binder_fun_types (fastype_of TUiter);
         val iter_preTs = if is_rec then rec_preTs else fold_preTs;
         val iter_strs = if is_rec then rec_strs else fold_strs;
+
         fun mk_s TU' =
           let
+            fun mk_absT_fp_repT repT absT = mk_absT thy repT absT ooo mk_repT;
+
             val i = find_index (fn T => co_alg_argT TU' = T) Xs;
+            val fp_abs = nth fp_abss i;
+            val fp_rep = nth fp_reps i;
+            val abs = nth abss i;
+            val rep = nth reps i;
             val sF = co_alg_funT TU';
+            val sF' =
+              mk_absT_fp_repT (nth repTs' i) (nth absTs' i) (nth fp_absTs i) (nth fp_repTs i) sF
+                handle Term.TYPE _ => sF;
             val F = nth iter_preTs i;
             val s = nth iter_strs i;
           in
-            (if sF = F then s
+            if sF = F then s
+            else if sF' = F then mk_co_comp_abs_rep sF sF' fp_abs fp_rep abs rep s
             else
               let
-                val smapT = replicate live dummyT ---> mk_co_algT sF F;
+                val smapT = replicate live dummyT ---> mk_co_algT sF' F;
                 fun hidden_to_unit t =
                   Term.subst_TVars (map (rpair HOLogic.unitT) (Term.add_tvar_names t [])) t;
                 val smap = map_of_bnf (nth bnfs i)
@@ -260,8 +346,9 @@
                     fst (fst (mk_iter NONE is_rec iters lthy TU)))
                 val smap_args = map mk_smap_arg smap_argTs;
               in
-                HOLogic.mk_comp (co_swap (s, Term.list_comb (smap, smap_args)))
-              end)
+                mk_co_comp_abs_rep sF sF' fp_abs fp_rep abs rep
+                  (mk_co_comp (s, Term.list_comb (smap, smap_args)))
+              end
           end;
         val t = Term.list_comb (TUiter, map mk_s TUs);
       in
@@ -306,13 +393,19 @@
               map HOLogic.id_const As @ map2 (mk_co_product o HOLogic.id_const) fpTs recs))
           Dss bnfs;
 
-        fun mk_goals f xtor s smap =
-          ((f, xtor), (s, smap))
-          |> pairself (HOLogic.mk_comp o co_swap)
-          |> HOLogic.mk_eq;
+        fun mk_goals f xtor s smap fp_abs fp_rep abs rep =
+          let
+            val lhs = mk_co_comp (f, xtor);
+            val rhs = mk_co_comp (s, smap);
+          in
+            HOLogic.mk_eq (lhs,
+              mk_co_comp_abs_rep (co_alg_funT (fastype_of lhs)) (co_alg_funT (fastype_of rhs))
+                fp_abs fp_rep abs rep rhs)
+          end;
 
-        val fold_goals = map4 mk_goals folds xtors fold_strs pre_fold_maps
-        val rec_goals = map4 mk_goals recs xtors rec_strs pre_rec_maps;
+        val fold_goals =
+          map8 mk_goals folds xtors fold_strs pre_fold_maps fp_abss fp_reps abss reps;
+        val rec_goals = map8 mk_goals recs xtors rec_strs pre_rec_maps fp_abss fp_reps abss reps;
 
         fun mk_thms ss goals tac =
           Library.foldr1 HOLogic.mk_conj goals
@@ -327,27 +420,42 @@
         val pre_map_defs = no_refl (map map_def_of_bnf bnfs);
         val fp_pre_map_defs = no_refl (map map_def_of_bnf pre_bnfs);
 
-        val map_unfolds = maps (fn bnf => no_refl [map_def_of_bnf bnf]) pre_bnfs;
-        val unfold_map = map (unfold_thms lthy (id_apply :: map_unfolds));
+        val unfold_map = map (unfold_thms lthy (id_apply :: pre_map_defs));
+
+        val fp_xtor_co_iterss = of_fp_res #xtor_co_iter_thmss;
+        val fp_xtor_un_folds = map (mk_pointfree lthy o un_fold_of) fp_xtor_co_iterss;
+        val fp_xtor_co_recs = map (mk_pointfree lthy o co_rec_of) fp_xtor_co_iterss;
 
-        val fp_xtor_co_iterss = steal_fp_res #xtor_co_iter_thmss;
-        val fp_xtor_un_folds = map (mk_pointfree lthy o un_fold_of) fp_xtor_co_iterss |> unfold_map;
-        val fp_xtor_co_recs = map (mk_pointfree lthy o co_rec_of) fp_xtor_co_iterss |> unfold_map;
-
-        val fp_co_iter_o_mapss = steal_fp_res #xtor_co_iter_o_map_thmss;
-        val fp_fold_o_maps = map un_fold_of fp_co_iter_o_mapss |> unfold_map;
-        val fp_rec_o_maps = map co_rec_of fp_co_iter_o_mapss |> unfold_map;
-        val fold_thms = fp_case fp @{thm comp_assoc} @{thm comp_assoc[symmetric]} :: @{thms id_apply
-          o_apply comp_id id_comp map_pair.comp map_pair.id sum_map.comp sum_map.id};
+        val fp_co_iter_o_mapss = of_fp_res #xtor_co_iter_o_map_thmss;
+        val fp_fold_o_maps = map un_fold_of fp_co_iter_o_mapss;
+        val fp_rec_o_maps = map co_rec_of fp_co_iter_o_mapss;
+        val fold_thms = fp_case fp @{thm comp_assoc} @{thm comp_assoc[symmetric]} ::
+          map (fn thm => thm RS rewrite_comp_comp) @{thms map_pair.comp sum_map.comp} @
+          @{thms id_apply comp_id id_comp map_pair.comp map_pair.id sum_map.comp sum_map.id};
         val rec_thms = fold_thms @ fp_case fp
           @{thms fst_convol map_pair_o_convol convol_o}
           @{thms case_sum_o_inj(1) case_sum_o_sum_map o_case_sum};
+
+        val eq_thm_prop_untyped = op Term.aconv_untyped o pairself Thm.full_prop_of;
+
         val map_thms = no_refl (maps (fn bnf =>
-          [map_comp0_of_bnf bnf RS sym, map_id0_of_bnf bnf]) fp_nesty_bnfs);
+           let val map_comp0 = map_comp0_of_bnf bnf RS sym
+           in [map_comp0, map_comp0 RS rewrite_comp_comp, map_id0_of_bnf bnf] end) fp_nesty_bnfs) @
+          remove eq_thm_prop_untyped (fp_case fp @{thm comp_assoc[symmetric]} @{thm comp_assoc})
+          (map2 (fn thm => fn bnf =>
+            @{thm type_copy_map_comp0_undo} OF
+              (replicate 3 thm @ unfold_map [map_comp0_of_bnf bnf]) RS
+              rewrite_comp_comp)
+          type_definitions bnfs);
+
+        fun mk_Rep_o_Abs thm = thm RS @{thm type_copy_Rep_o_Abs} RS rewrite_comp_comp;
+
+        val fp_Rep_o_Abss = map mk_Rep_o_Abs fp_type_definitions;
+        val Rep_o_Abss = map mk_Rep_o_Abs type_definitions;
 
         fun mk_tac defs o_map_thms xtor_thms thms {context = ctxt, prems = _} =
-          unfold_thms_tac ctxt
-            (flat [thms, defs, pre_map_defs, fp_pre_map_defs, xtor_thms, o_map_thms, map_thms]) THEN
+          unfold_thms_tac ctxt (flat [thms, defs, pre_map_defs, fp_pre_map_defs,
+            xtor_thms, o_map_thms, map_thms, fp_Rep_o_Abss, Rep_o_Abss]) THEN
           CONJ_WRAP (K (HEADGOAL (rtac refl))) bnfs;
 
         val fold_tac = mk_tac raw_un_fold_defs fp_fold_o_maps fp_xtor_un_folds fold_thms;
@@ -360,20 +468,20 @@
        used by "primrec", "primcorecursive", and "datatype_compat". *)
     val fp_res =
       ({Ts = fpTs,
-        bnfs = steal_fp_res #bnfs,
+        bnfs = of_fp_res #bnfs,
         dtors = dtors,
         ctors = ctors,
         xtor_co_iterss = transpose [un_folds, co_recs],
         xtor_co_induct = xtor_co_induct_thm,
-        dtor_ctors = steal_fp_res #dtor_ctors (*too general types*),
-        ctor_dtors = steal_fp_res #ctor_dtors (*too general types*),
-        ctor_injects = steal_fp_res #ctor_injects (*too general types*),
-        dtor_injects = steal_fp_res #dtor_injects (*too general types*),
-        xtor_map_thms = steal_fp_res #xtor_map_thms (*too general types and terms*),
-        xtor_set_thmss = steal_fp_res #xtor_set_thmss (*too general types and terms*),
-        xtor_rel_thms = steal_fp_res #xtor_rel_thms (*too general types and terms*),
+        dtor_ctors = of_fp_res #dtor_ctors (*too general types*),
+        ctor_dtors = of_fp_res #ctor_dtors (*too general types*),
+        ctor_injects = of_fp_res #ctor_injects (*too general types*),
+        dtor_injects = of_fp_res #dtor_injects (*too general types*),
+        xtor_map_thms = of_fp_res #xtor_map_thms (*too general types and terms*),
+        xtor_set_thmss = of_fp_res #xtor_set_thmss (*too general types and terms*),
+        xtor_rel_thms = of_fp_res #xtor_rel_thms (*too general types and terms*),
         xtor_co_iter_thmss = transpose [xtor_un_fold_thms, xtor_co_rec_thms],
-        xtor_co_iter_o_map_thmss = steal_fp_res #xtor_co_iter_o_map_thmss
+        xtor_co_iter_o_map_thmss = of_fp_res #xtor_co_iter_o_map_thmss
           (*theorem about old constant*),
         rel_xtor_co_induct_thm = rel_xtor_co_induct_thm}
        |> morph_fp_result (Morphism.term_morphism "BNF" (singleton (Variable.polymorphic lthy))));
--- a/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML	Tue Feb 25 18:14:26 2014 +0100
@@ -191,13 +191,13 @@
 
     val ctr_Tsss = map (map binder_types) ctr_Tss;
     val ctrXs_Tsss = map4 (map2 o map2 oo freeze_fpTs_call) kks fpTs callssss ctr_Tsss;
-    val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
+    val ctrXs_repTs = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
 
     val ns = map length ctr_Tsss;
     val kss = map (fn n => 1 upto n) ns;
     val mss = map (map length) ctr_Tsss;
 
-    val fp_eqs = map dest_TFree Xs ~~ ctrXs_sum_prod_Ts;
+    val fp_eqs = map dest_TFree Xs ~~ ctrXs_repTs;
     val key = key_of_fp_eqs fp fpTs fp_eqs;
   in
     (case n2m_sugar_of no_defs_lthy key of
@@ -221,32 +221,45 @@
                 map_filter (fn (T, U) => if member (op =) deads T then SOME U else NONE) (Ts ~~ Us);
             in fold Term.add_tfreesT dead_Us [] end);
 
-        val (pre_bnfs, (fp_res as {xtor_co_iterss = xtor_co_iterss0, xtor_co_induct, dtor_injects,
-               dtor_ctors, xtor_co_iter_thmss, ...}, lthy)) =
+        val fp_absT_infos = map #absT_info fp_sugars0;
+
+        val ((pre_bnfs, absT_infos), (fp_res as {xtor_co_iterss = xtor_co_iterss0, xtor_co_induct,
+               dtor_injects, dtor_ctors, xtor_co_iter_thmss, ...}, lthy)) =
           fp_bnf (construct_mutualized_fp fp fpTs fp_sugars0) fp_bs As' killed_As' fp_eqs
-            no_defs_lthy;
+            no_defs_lthy0;
+
+        val fp_abs_inverses = map #abs_inverse fp_absT_infos;
+        val fp_type_definitions = map #type_definition fp_absT_infos;
+
+        val abss = map #abs absT_infos;
+        val reps = map #rep absT_infos;
+        val absTs = map #absT absT_infos;
+        val repTs = map #repT absT_infos;
+        val abs_inverses = map #abs_inverse absT_infos;
 
         val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
         val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
 
         val ((xtor_co_iterss, iters_args_types, coiters_args_types), _) =
-          mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
+          mk_co_iters_prelims fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_iterss0 lthy;
 
         fun mk_binding b suf = Binding.suffix_name ("_" ^ suf) b;
 
         val ((co_iterss, co_iter_defss), lthy) =
           fold_map2 (fn b =>
-            (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
-             else define_coiters [unfoldN, corecN] (the coiters_args_types))
-              (mk_binding b) fpTs Cs) fp_bs xtor_co_iterss lthy
+            if fp = Least_FP then
+              define_iters [foldN, recN] (the iters_args_types) (mk_binding b) fpTs Cs reps
+            else
+              define_coiters [unfoldN, corecN] (the coiters_args_types) (mk_binding b) fpTs Cs abss)
+            fp_bs xtor_co_iterss lthy
           |>> split_list;
 
         val ((common_co_inducts, co_inductss, un_fold_thmss, co_rec_thmss, disc_unfold_thmss,
               disc_corec_thmss, sel_unfold_thmsss, sel_corec_thmsss), fp_sugar_thms) =
           if fp = Least_FP then
             derive_induct_iters_thms_for_types pre_bnfs (the iters_args_types) xtor_co_induct
-              xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss
-              co_iterss co_iter_defss lthy
+              xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss fp_abs_inverses
+              fp_type_definitions abs_inverses ctrss ctr_defss co_iterss co_iter_defss lthy
             |> `(fn ((inducts, induct, _), (fold_thmss, rec_thmss, _)) =>
               ([induct], [inducts], fold_thmss, rec_thmss, replicate nn [],
                replicate nn [], replicate nn [], replicate nn []))
@@ -254,8 +267,8 @@
           else
             derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
               dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss
-              ns ctr_defss ctr_sugars co_iterss co_iter_defss
-              (Proof_Context.export lthy no_defs_lthy) lthy
+              ns fp_abs_inverses abs_inverses (fn thm => thm RS @{thm vimage2p_refl}) ctr_defss
+              ctr_sugars co_iterss co_iter_defss (Proof_Context.export lthy no_defs_lthy) lthy
             |> `(fn ((coinduct_thms_pairs, _), (unfold_thmss, corec_thmss, _),
                     (disc_unfold_thmss, disc_corec_thmss, _), _,
                     (sel_unfold_thmsss, sel_corec_thmsss, _)) =>
@@ -266,21 +279,21 @@
 
         val phi = Proof_Context.export_morphism no_defs_lthy no_defs_lthy0;
 
-        fun mk_target_fp_sugar T X kk pre_bnf ctrXs_Tss ctr_defs ctr_sugar co_iters maps co_inducts
-            un_fold_thms co_rec_thms disc_unfold_thms disc_corec_thms sel_unfold_thmss
+        fun mk_target_fp_sugar T X kk pre_bnf absT_info ctrXs_Tss ctr_defs ctr_sugar co_iters maps
+            co_inducts un_fold_thms co_rec_thms disc_unfold_thms disc_corec_thms sel_unfold_thmss
             sel_corec_thmss =
           {T = T, X = X, fp = fp, fp_res = fp_res, fp_res_index = kk, pre_bnf = pre_bnf,
-           nested_bnfs = nested_bnfs, nesting_bnfs = nesting_bnfs, ctrXs_Tss = ctrXs_Tss,
-           ctr_defs = ctr_defs, ctr_sugar = ctr_sugar, co_iters = co_iters, maps = maps,
-           common_co_inducts = common_co_inducts, co_inducts = co_inducts,
+           absT_info = absT_info, nested_bnfs = nested_bnfs, nesting_bnfs = nesting_bnfs,
+           ctrXs_Tss = ctrXs_Tss, ctr_defs = ctr_defs, ctr_sugar = ctr_sugar, co_iters = co_iters,
+           maps = maps, common_co_inducts = common_co_inducts, co_inducts = co_inducts,
            co_iter_thmss = [un_fold_thms, co_rec_thms],
            disc_co_iterss = [disc_unfold_thms, disc_corec_thms],
            sel_co_itersss = [sel_unfold_thmss, sel_corec_thmss]}
           |> morph_fp_sugar phi;
 
         val target_fp_sugars =
-          map16 mk_target_fp_sugar fpTs Xs kks pre_bnfs ctrXs_Tsss ctr_defss ctr_sugars co_iterss
-            mapss (transpose co_inductss) un_fold_thmss co_rec_thmss disc_unfold_thmss
+          map17 mk_target_fp_sugar fpTs Xs kks pre_bnfs absT_infos ctrXs_Tsss ctr_defss ctr_sugars
+            co_iterss mapss (transpose co_inductss) un_fold_thmss co_rec_thmss disc_unfold_thmss
             disc_corec_thmss sel_unfold_thmsss sel_corec_thmsss;
 
         val n2m_sugar = (target_fp_sugars, fp_sugar_thms);
--- a/src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML	Tue Feb 25 18:14:26 2014 +0100
@@ -8,29 +8,35 @@
 signature BNF_FP_N2M_TACTICS =
 sig
   val mk_rel_xtor_co_induct_tactic: BNF_Util.fp_kind -> thm list -> thm list -> thm list ->
-    {prems: thm list, context: Proof.context} -> tactic
+    thm list -> {prems: thm list, context: Proof.context} -> tactic
 end;
 
 structure BNF_FP_N2M_Tactics : BNF_FP_N2M_TACTICS =
 struct
 
 open BNF_Util
+open BNF_Tactics
 open BNF_FP_Util
 
-fun mk_rel_xtor_co_induct_tactic fp co_inducts rel_defs rel_monos
+val vimage2p_unfolds = o_apply :: @{thms vimage2p_def};
+
+fun mk_rel_xtor_co_induct_tactic fp abs_inverses co_inducts0 rel_defs rel_monos
   {context = ctxt, prems = raw_C_IHs} =
   let
-    val unfolds = map (fn def => unfold_thms ctxt (id_apply :: no_reflexive [def])) rel_defs;
+    val co_inducts = map (unfold_thms ctxt vimage2p_unfolds) co_inducts0;
+    val unfolds = map (fn def =>
+      unfold_thms ctxt (id_apply :: vimage2p_unfolds @ abs_inverses @ no_reflexive [def])) rel_defs;
     val folded_C_IHs = map (fn thm => thm RS @{thm spec2} RS mp) raw_C_IHs;
     val C_IHs = map2 (curry op |>) folded_C_IHs unfolds;
     val C_IH_monos =
       map3 (fn C_IH => fn mono => fn unfold =>
-        (mono RSN (2, @{thm rev_predicate2D}), C_IH)
+        (mono RSN (2, @{thm vimage2p_mono}), C_IH)
         |> fp = Greatest_FP ? swap
         |> op RS
         |> unfold)
       folded_C_IHs rel_monos unfolds;
   in
+    unfold_thms_tac ctxt vimage2p_unfolds THEN
     HEADGOAL (CONJ_WRAP_GEN' (rtac @{thm context_conjI})
       (fn thm => rtac thm THEN_ALL_NEW (rotate_tac ~1 THEN'
          REPEAT_ALL_NEW (FIRST' [eresolve_tac C_IHs, eresolve_tac C_IH_monos,
--- a/src/HOL/Tools/BNF/bnf_fp_util.ML	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_util.ML	Tue Feb 25 18:14:26 2014 +0100
@@ -142,16 +142,15 @@
 
   val mk_case_sum: term * term -> term
   val mk_case_sumN: term list -> term
-  val mk_case_sumN_balanced: term list -> term
+  val mk_case_absumprod: typ -> term -> term list -> term list -> term list list -> term
+
   val mk_Inl: typ -> term -> term
   val mk_Inr: typ -> term -> term
   val mk_InN: typ list -> term -> int -> term
-  val mk_InN_balanced: typ -> int -> term -> int -> term
+  val mk_absumprod: typ -> term -> int -> int -> term list -> term
 
   val dest_sumT: typ -> typ * typ
-  val dest_sumTN: int -> typ -> typ list
-  val dest_sumTN_balanced: int -> typ -> typ list
-  val dest_tupleT: int -> typ -> typ list
+  val dest_absumprodT: typ -> typ -> int -> int list -> typ -> typ list list
 
   val If_const: typ -> term
 
@@ -160,8 +159,8 @@
   val mk_union: term * term -> term
 
   val mk_sumEN: int -> thm
-  val mk_sumEN_balanced: int -> thm
-  val mk_sumEN_tupled_balanced: int list -> thm
+  val mk_absumprodE: thm -> int list -> thm
+
   val mk_sum_caseN: int -> int -> thm
   val mk_sum_caseN_balanced: int -> int -> thm
 
@@ -176,12 +175,12 @@
   val mk_xtor_un_fold_o_map_thms: BNF_Util.fp_kind -> bool -> int -> thm -> thm list -> thm list ->
     thm list -> thm list -> thm list
 
-  val mk_strong_coinduct_thm: thm -> thm list -> thm list -> Proof.context -> thm
+  val mk_strong_coinduct_thm: thm -> thm list -> thm list -> (thm -> thm) -> Proof.context -> thm
 
   val fp_bnf: (binding list -> (string * sort) list -> typ list * typ list list ->
-      BNF_Def.bnf list -> local_theory -> 'a) ->
+      BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> 'a) ->
     binding list -> (string * sort) list -> (string * sort) list -> ((string * sort) * typ) list ->
-    local_theory -> BNF_Def.bnf list * 'a
+    local_theory -> (BNF_Def.bnf list * BNF_Comp.absT_info list) * 'a
 end;
 
 structure BNF_FP_Util : BNF_FP_UTIL =
@@ -347,9 +346,6 @@
 
 fun dest_sumT (Type (@{type_name sum}, [T, T'])) = (T, T');
 
-fun dest_sumTN 1 T = [T]
-  | dest_sumTN n (Type (@{type_name sum}, [T, T'])) = T :: dest_sumTN (n - 1) T';
-
 val dest_sumTN_balanced = Balanced_Tree.dest dest_sumT;
 
 (* TODO: move something like this to "HOLogic"? *)
@@ -357,6 +353,8 @@
   | dest_tupleT 1 T = [T]
   | dest_tupleT n (Type (@{type_name prod}, [T, T'])) = T :: dest_tupleT (n - 1) T';
 
+fun dest_absumprodT absT repT n ms = map2 dest_tupleT ms o dest_sumTN_balanced n o mk_repT absT repT;
+
 val mk_sumTN = Library.foldr1 mk_sumT;
 val mk_sumTN_balanced = Balanced_Tree.make mk_sumT;
 
@@ -397,6 +395,10 @@
     |> repair_types sum_T
   end;
 
+fun mk_absumprod absT abs0 n k ts =
+  let val abs = mk_abs absT abs0;
+  in abs $ mk_InN_balanced (domain_type (fastype_of abs)) n (HOLogic.mk_tuple ts) k end;
+
 fun mk_case_sum (f, g) =
   let
     val fT = fastype_of f;
@@ -409,6 +411,12 @@
 val mk_case_sumN = Library.foldr1 mk_case_sum;
 val mk_case_sumN_balanced = Balanced_Tree.make mk_case_sum;
 
+fun mk_tupled_fun f x xs =
+  if xs = [x] then f else HOLogic.tupled_lambda x (Term.list_comb (f, xs));
+
+fun mk_case_absumprod absT rep fs xs xss =
+  HOLogic.mk_comp (mk_case_sumN_balanced (map3 mk_tupled_fun fs xs xss), mk_rep absT rep);
+
 fun If_const T = Const (@{const_name If}, HOLogic.boolT --> T --> T --> T);
 fun mk_If p t f = let val T = fastype_of t in If_const T $ p $ t $ f end;
 
@@ -441,21 +449,15 @@
   Balanced_Tree.make (fn (thm1, thm2) => thm1 RSN (1, thm2 RSN (2, @{thm obj_sumE_f})))
     (replicate n asm_rl);
 
-fun mk_sumEN_balanced' n all_impIs = mk_obj_sumEN_balanced n OF all_impIs RS @{thm obj_one_pointE};
-
-fun mk_sumEN_balanced 1 = @{thm one_pointE} (*optimization*)
-  | mk_sumEN_balanced 2 = @{thm sumE} (*optimization*)
-  | mk_sumEN_balanced n = mk_sumEN_balanced' n (replicate n (impI RS allI));
-
 fun mk_tupled_allIN 0 = @{thm unit_all_impI}
   | mk_tupled_allIN 1 = @{thm impI[THEN allI]}
   | mk_tupled_allIN 2 = @{thm prod_all_impI} (*optimization*)
   | mk_tupled_allIN n = mk_tupled_allIN (n - 1) RS @{thm prod_all_impI_step};
 
-fun mk_sumEN_tupled_balanced ms =
+fun mk_absumprodE type_definition ms =
   let val n = length ms in
-    if forall (curry op = 1) ms then mk_sumEN_balanced n
-    else mk_sumEN_balanced' n (map mk_tupled_allIN ms)
+    mk_obj_sumEN_balanced n OF map mk_tupled_allIN ms RS
+      (type_definition RS @{thm type_copy_obj_one_point_absE})
   end;
 
 fun mk_sum_caseN 1 1 = refl
@@ -543,7 +545,7 @@
     split_conj_thm (un_fold_unique OF map (fp_case fp I mk_sym) unique_prems)
   end;
 
-fun mk_strong_coinduct_thm coind rel_eqs rel_monos ctxt =
+fun mk_strong_coinduct_thm coind rel_eqs rel_monos mk_vimage2p ctxt =
   let
     val n = Thm.nprems_of coind;
     val m = Thm.nprems_of (hd rel_monos) - n;
@@ -554,7 +556,7 @@
       let
         val eq = iffD2 OF [rel_eq RS @{thm predicate2_eqD}, refl];
         val mono = rel_mono OF (replicate m @{thm order_refl} @ replicate n @{thm eq_subset});
-      in eq RS (mono RS @{thm predicate2D}) RS @{thm eqTrueI} end;
+      in mk_vimage2p (eq RS (mono RS @{thm predicate2D})) RS @{thm eqTrueI} end;
     val unfolds = map2 mk_unfold rel_eqs rel_monos @ @{thms sup_fun_def sup_bool_def
       imp_disjL all_conj_distrib subst_eq_imp simp_thms(18,21,35)};
   in
@@ -603,18 +605,19 @@
     fun pre_qualify b = Binding.qualify false (Binding.name_of b)
       #> Config.get lthy' bnf_note_all = false ? Binding.conceal;
 
-    val ((pre_bnfs, deadss), lthy'') =
+    val ((pre_bnfs, (deadss, absT_infos)), lthy'') =
       fold_map3 (fn b => seal_bnf (pre_qualify b) unfold_set' (Binding.prefix_name preN b))
         bs Dss bnfs' lthy'
-      |>> split_list;
+      |>> split_list
+      |>> apsnd split_list;
 
     val timer = time (timer "Normalization & sealing of BNFs");
 
-    val res = construct_fp bs resBs (map TFree resDs, deadss) pre_bnfs lthy'';
+    val res = construct_fp bs resBs (map TFree resDs, deadss) pre_bnfs absT_infos lthy'';
 
     val timer = time (timer "FP construction in total");
   in
-    timer; (pre_bnfs, res)
+    timer; ((pre_bnfs, absT_infos), res)
   end;
 
 end;
--- a/src/HOL/Tools/BNF/bnf_gfp.ML	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp.ML	Tue Feb 25 18:14:26 2014 +0100
@@ -11,7 +11,7 @@
 sig
   val construct_gfp: mixfix list -> binding list -> binding list -> binding list list ->
     binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
-    local_theory -> BNF_FP_Util.fp_result * local_theory
+    BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory
 end;
 
 structure BNF_GFP : BNF_GFP =
@@ -55,7 +55,7 @@
      ((i, I), nth (nth lwitss i) nwit) :: maps (tree_to_coind_wits lwitss) subtrees;
 
 (*all BNFs have the same lives*)
-fun construct_gfp mixfixes map_bs rel_bs set_bss0 bs resBs (resDs, Dss) bnfs lthy =
+fun construct_gfp mixfixes map_bs rel_bs set_bss0 bs resBs (resDs, Dss) bnfs _ lthy =
   let
     val time = time lthy;
     val timer = time (Timer.startRealTimer ());
--- a/src/HOL/Tools/BNF/bnf_lfp.ML	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp.ML	Tue Feb 25 18:14:26 2014 +0100
@@ -10,7 +10,7 @@
 sig
   val construct_lfp: mixfix list -> binding list -> binding list -> binding list list ->
     binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
-    local_theory -> BNF_FP_Util.fp_result * local_theory
+    BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory
 end;
 
 structure BNF_LFP : BNF_LFP =
@@ -26,7 +26,7 @@
 open BNF_LFP_Tactics
 
 (*all BNFs have the same lives*)
-fun construct_lfp mixfixes map_bs rel_bs set_bss0 bs resBs (resDs, Dss) bnfs lthy =
+fun construct_lfp mixfixes map_bs rel_bs set_bss0 bs resBs (resDs, Dss) bnfs _ lthy =
   let
     val time = time lthy;
     val timer = time (Timer.startRealTimer ());
@@ -701,6 +701,7 @@
       let
         val i_field = HOLogic.mk_mem (idx, field_suc_bd);
         val min_algs = mk_min_algs ss;
+
         val min_algss = map (fn k => mk_nthN n (min_algs $ idx) k) ks;
 
         val concl = HOLogic.mk_Trueprop
--- a/src/HOL/Tools/BNF/bnf_util.ML	Fri Feb 28 12:04:40 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_util.ML	Tue Feb 25 18:14:26 2014 +0100
@@ -43,6 +43,11 @@
       'o -> 'p -> 'q) ->
     'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
     'i list -> 'j list -> 'k list -> 'l list -> 'm list -> 'n list -> 'o list -> 'p list -> 'q list
+  val map17: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l -> 'm -> 'n ->
+      'o -> 'p -> 'q -> 'r) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+    'i list -> 'j list -> 'k list -> 'l list -> 'm list -> 'n list -> 'o list -> 'p list ->
+    'q list -> 'r list
   val fold_map4: ('a -> 'b -> 'c -> 'd -> 'e -> 'f * 'e) ->
     'a list -> 'b list -> 'c list -> 'd list -> 'e -> 'f list * 'e
   val fold_map5: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g * 'f) ->
@@ -109,6 +114,7 @@
   val mk_ordLeq: term -> term -> term
   val mk_rel_comp: term * term -> term
   val mk_rel_compp: term * term -> term
+  val mk_vimage2p: term -> term -> term
 
   (*parameterized terms*)
   val mk_nthN: int -> term -> int -> term
@@ -225,6 +231,14 @@
       map16 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s x12s x13s x14s x15s x16s
   | map16 _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
 
+fun map17 _ [] [] [] [] [] [] [] [] [] [] [] [] [] [] [] [] [] = []
+  | map17 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
+      (x9::x9s) (x10::x10s) (x11::x11s) (x12::x12s) (x13::x13s) (x14::x14s) (x15::x15s)
+      (x16::x16s) (x17::x17s) =
+    f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 x14 x15 x16 x17 ::
+      map17 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s x12s x13s x14s x15s x16s x17s
+  | map17 _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
 fun fold_map4 _ [] [] [] [] acc = ([], acc)
   | fold_map4 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) acc =
     let
@@ -498,6 +512,15 @@
   let val T = (case xs of [] => defT | (x::_) => fastype_of x);
   in Const (@{const_name collect}, HOLogic.mk_setT T --> T) $ (HOLogic.mk_set T xs) end;
 
+fun mk_vimage2p f g =
+  let
+    val (T1, T2) = dest_funT (fastype_of f);
+    val (U1, U2) = dest_funT (fastype_of g);
+  in
+    Const (@{const_name vimage2p},
+      (T1 --> T2) --> (U1 --> U2) --> mk_pred2T T2 U2 --> mk_pred2T T1 U1) $ f $ g
+  end;
+
 fun mk_trans thm1 thm2 = trans OF [thm1, thm2];
 fun mk_sym thm = thm RS sym;