Theory EquivalenceOfBicategories

(*  Title:       EquivalenceOfBicategories
    Author:      Eugene W. Stark <stark@cs.stonybrook.edu>, 2020
    Maintainer:  Eugene W. Stark <stark@cs.stonybrook.edu>
*)

section "Equivalence of Bicategories"

text ‹
  In this section, we define the notion ``equivalence of bicategories'', which generalizes
  the notion of equivalence of categories, and we establish various of its properties.
  In particular, we show that ``equivalent bicategories'' is an equivalence relation,
  and that a pseudofunctor is part of an equivalence of bicategories if and only if it
  is an equivalence pseudofunctor (\emph{i.e.}~it is faithful, locally full, locally
  essentially surjective, and biessentially surjective on objects).
›

theory EquivalenceOfBicategories
imports InternalAdjunction PseudonaturalTransformation
begin

  subsection "Definition of Equivalence of Bicategories"

  text ‹
    An equivalence of bicategories between bicategories C› and D› consists of pseudofunctors
    F : D → C› and G : C → D› and pseudonatural equivalences η: ID ≈ GF› and
    ε: FG ≈ IC.
  ›

  locale equivalence_of_bicategories =  (* 25 sec *)
    C: bicategory VC HC 𝖺C 𝗂C srcC trgC +
    D: bicategory VD HD 𝖺D 𝗂D srcD trgD +
    F: pseudofunctor VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC F ΦF +
    G: pseudofunctor VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD G ΦG +
    FG: composite_pseudofunctor VC HC 𝖺C 𝗂C srcC trgC
          VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC G ΦG F ΦF +
    GF: composite_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD
          VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F ΦF G ΦG +
    IC: identity_pseudofunctor VC HC 𝖺C 𝗂C srcC trgC +
    ID: identity_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD +
    η: pseudonatural_equivalence VD HD 𝖺D 𝗂D srcD trgD VD HD 𝖺D 𝗂D srcD trgD
          ID.map ID.cmp GF.map GF.cmp η0 η1 +
    ε: pseudonatural_equivalence VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
          FG.map FG.cmp IC.map IC.cmp ε0 ε1
  for VC :: "'c comp"                    (infixr "C" 55)
  and HC :: "'c comp"                   (infixr "C" 53)
  and 𝖺C :: "'c  'c  'c  'c"       ("𝖺C[_, _, _]")
  and 𝗂C :: "'c  'c"                   ("𝗂C[_]")
  and srcC :: "'c  'c"
  and trgC :: "'c  'c"
  and VD :: "'d comp"                    (infixr "D" 55)
  and HD :: "'d comp"                   (infixr "D" 53)
  and 𝖺D :: "'d  'd  'd  'd"       ("𝖺D[_, _, _]")
  and 𝗂D :: "'d  'd"                   ("𝗂D[_]")
  and srcD :: "'d  'd"
  and trgD :: "'d  'd"
  and F :: "'d  'c"
  and ΦF :: "'d * 'd  'c"
  and G :: "'c  'd"
  and ΦG :: "'c * 'c  'd"
  and η0 :: "'d  'd"
  and η1 :: "'d  'd"
  and ε0 :: "'c  'c"
  and ε1 :: "'c  'c"
  begin

    notation C.isomorphic (infix "C" 50)
    notation D.isomorphic (infix "D" 50)

    notation C.iso_in_hom ("«_ : _ C _»")
    notation D.iso_in_hom ("«_ : _ D _»")

    notation C.some_quasi_inverse  ("_~C" [1000] 1000)
    notation D.some_quasi_inverse  ("_~D" [1000] 1000)

    interpretation η': converse_pseudonatural_equivalence
                         VD HD 𝖺D 𝗂D srcD trgD VD HD 𝖺D 𝗂D srcD trgD
                         ID.map ID.cmp GF.map GF.cmp η0 η1
      ..
    interpretation ε': converse_pseudonatural_equivalence
                         VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                         FG.map FG.cmp IC.map IC.cmp ε0 ε1
      ..

    text ‹
      Although it will be some trouble yet to prove that F› is an equivalence pseudofunctor,
      it is not as difficult to prove that the composites GF› and FG› are equivalence
      pseudofunctors, by virtue of their being pseudonaturally equivalent to identity
      pseudofunctors.
    ›

    interpretation GF: equivalence_pseudofunctor
                          VD HD 𝖺D 𝗂D srcD trgD VD HD 𝖺D 𝗂D srcD trgD GF.map GF.cmp
    proof -
      interpret GF: pseudofunctor_pseudonaturally_equivalent_to_equivalence_pseudofunctor
                       VD HD 𝖺D 𝗂D srcD trgD VD HD 𝖺D 𝗂D srcD trgD
                       ID.map ID.cmp GF.map GF.cmp η0 η1
        ..
      show "equivalence_pseudofunctor
              VD HD 𝖺D 𝗂D srcD trgD VD HD 𝖺D 𝗂D srcD trgD GF.map GF.cmp"
        using GF.is_equivalence_pseudofunctor by simp
    qed

    interpretation FG: equivalence_pseudofunctor
                          VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC FG.map FG.cmp
    proof -
      interpret FG: pseudofunctor_pseudonaturally_equivalent_to_equivalence_pseudofunctor
                       VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                       IC.map IC.cmp FG.map FG.cmp ε'.map0 ε'.map1
        ..
      show "equivalence_pseudofunctor
              VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC FG.map FG.cmp"
        using FG.is_equivalence_pseudofunctor by simp
    qed

    text ‹
      It is also easy to establish that F› and G› are faithful.
      We will use the facts, that GF› is locally full and G› is faithful,
      to prove that F› is locally full.
    ›

    interpretation F: faithful_functor VD VC F
    proof
      fix μ μ'
      assume par: "D.par μ μ'" and eq: "F μ = F μ'"
      have "srcD μ = srcD μ'  trgD μ = trgD μ'"
        using par D.src_dom D.trg_cod by (metis D.src_cod D.trg_cod)
      hence "η0 (trgD μ) D μ = η0 (trgD μ) D μ'"
        using par eq η.iso_map1_ide D.comp_arr_inv' D.trg.preserves_dom
              D.comp_arr_dom D.comp_assoc
        by (metis GF.is_faithful o_apply)
      thus "μ = μ'"
        using par η.ide_map0_obj η.components_are_equivalences
              D.equivalence_cancel_left [of "η0 (trgD μ)" μ μ']
        by simp
    qed

    interpretation G: faithful_functor VC VD G
    proof
      show "μ μ'. C.par μ μ'; G μ = G μ'  μ = μ'"
      proof -
        fix μ μ'
        assume par: "C.par μ μ'" and eq: "G μ = G μ'"
        have "srcC μ = srcC μ'  trgC μ = trgC μ'"
          using par by (metis C.src_cod C.trg_cod)
        hence "μ C ε0 (srcC μ) = μ' C ε0 (srcC μ)"
            using par eq ε.iso_map1_ide C.comp_inv_arr' C.src.preserves_dom
                  C.comp_arr_dom C.comp_assoc
            by (metis FG.is_faithful o_apply)
        thus "μ = μ'"
          using par ε.ide_map0_obj ε.components_are_equivalences
                C.equivalence_cancel_right [of "ε0 (srcC μ)" μ μ']
          by simp
      qed
    qed

    text ‹
      It is perhaps not so easy to discover a proof that F› is locally essentially surjective.
      Here we follow the proof of cite"johnson-yau-2d-categories", Lemma 6.2.13, except we
      have expressed it in a way that explicitly shows its constructive nature, given that
      we have already chosen an extension of each component of η› and ε› to an adjoint
      equivalence.
    ›

    abbreviation Φ
    where "Φ ψ f f'  𝗋C[f'] C
                       (f' C ε'.counit (srcC f)) C
                       𝖺C[f', ε0 (srcC f), ε'.map0 (srcC f)] C
                       (C.inv (ε1 f') C ε'.map0 (srcC f)) C
                       𝖺C-1[ε0 (trgC f), F (G f'), ε'.map0 (srcC f)] C
                       (ε0 (trgC f) C F ψ C ε'.map0 (srcC f)) C
                       (ε0 (trgC f) C C.inv (ε'.map1 f)) C
                       𝖺C[ε0 (trgC f), ε'.map0 (trgC f), f] C
                       (C.inv (ε'.counit (trgC f)) C f) C
                       𝗅C-1[f]"

    lemma G_reflects_isomorphic:
    assumes "C.ide f" and "C.ide f'" and "srcC f = srcC f'" and "trgC f = trgC f'"
    and "«ψ : G f D G f'»"
    shows "«Φ ψ f f' : f C f'»"
    proof -
      let ?a = "srcC f" and ?a' = "trgC f"
      have f: "«f : ?a C ?a'»  C.ide f"
        using assms by simp
      have f': "«f' : ?a C ?a'»  C.ide f'"
        using assms by simp
      have ψ: "«ψ : G.map0 ?a D G.map0 ?a'»"
        using assms D.vconn_implies_hpar(1-2)
        by (elim D.iso_in_homE) auto
      hence : "«F ψ : F.map0 (G.map0 ?a) C F.map0 (G.map0 ?a')»"
        by auto
      show "«Φ ψ f f' : f C f'»"
      proof (intro C.comp_iso_in_hom)
        show "«𝗅C-1[f] : f C ?a' C f»"
          using f by auto
        show "«C.inv (ε'.counit ?a') C f : ?a' C f C (ε0 ?a' C ε'.map0 ?a') C f»"
          using f ε'.counit_in_hom [of ?a'] ε'.iso_counit [of ?a']
          by (intro C.hcomp_iso_in_hom) auto
        show "«𝖺C[ε0 ?a', ε'.map0 ?a', f] :
                 (ε0 ?a' C ε'.map0 ?a') C f C ε0 ?a' C ε'.map0 ?a' C f»"
          using f ε'.counit_in_hom [of ?a] ε'.ide_map0_obj
          by (intro C.iso_in_homI) auto
        show "«ε0 ?a' C C.inv (ε'.map1 f) :
                 ε0 ?a' C ε'.map0 ?a' C f C ε0 ?a' C F (G f) C ε'.map0 ?a»"
          using f ε'.map1_in_hom [of f] ε'.iso_map1_ide [of f] C.ide_iso_in_hom
          by (intro C.hcomp_iso_in_hom) auto
        show "«ε0 ?a' C F ψ C ε'.map0 ?a :
                 ε0 ?a' C F (G f) C ε'.map0 ?a C ε0 ?a' C F (G f') C ε'.map0 ?a»"
          using assms f f'  F.preserves_iso C.ide_iso_in_hom
          by (intro C.hcomp_iso_in_hom) auto
        show "«𝖺C-1[ε0 ?a', F (G f'), ε'.map0 ?a] :
                 ε0 ?a' C F (G f') C ε'.map0 ?a C (ε0 ?a' C F (G f')) C ε'.map0 ?a»"
          using assms f' ε.map0_in_hom(1) [of ?a'] ε.ide_map0_obj [of ?a']
                ε'.map0_in_hom(1) [of ?a] ε'.ide_map0_obj [of ?a]
                C.assoc'_in_hom C.iso_assoc'
          by auto
        show "«C.inv (ε1 f') C ε'.map0 ?a :
                 (ε0 ?a' C F (G f')) C ε'.map0 ?a C (f' C ε0 ?a)  C ε'.map0 ?a»"
          using assms f' ε.map1_in_hom ε.iso_map1_ide ε'.map0_in_hom(1) [of ?a]
                C.ide_iso_in_hom
          by (intro C.hcomp_iso_in_hom) auto
        show "«𝖺C[f', ε0 ?a, ε'.map0 ?a] :
                (f' C ε0 ?a) C ε'.map0 ?a C f' C ε0 ?a C ε'.map0 ?a»"
          using assms f' ε.map0_in_hom(1) [of ?a] ε'.map0_in_hom(1) [of ?a]
                ε.ide_map0_obj ε'.ide_map0_obj C.assoc_in_hom
          by auto
        show "«f' C ε'.counit ?a : f' C ε0 ?a C ε'.map0 ?a C f' C ?a»"
          using f f' ε'.counit_in_hom
          by (intro C.hcomp_iso_in_hom) auto
        show "«𝗋C[f'] : f' C ?a C f'»"
          using assms f' by auto
      qed
    qed

    abbreviation Ψ
    where "Ψ f b b'  𝗋D[G (F (η'.map0 b' D G f D η0 b))] D
                       (G (F (η'.map0 b' D G f D η0 b)) D η'.ε b) D
                       𝖺D[G (F (η'.map0 b' D G f D η0 b)), η0 b, η'.map0 b] D
                       (D.inv (η1 (η'.map0 b' D G f D η0 b)) D η'.map0 b) D
                       (𝖺D[η0 b', η'.map0 b', G f D η0 b] D η'.map0 b) D
                       𝖺D-1[η0 b' D η'.map0 b', G f D η0 b, η'.map0 b] D
                       ((η0 b' D η'.map0 b') D 𝖺D-1[G f, η0 b, η'.map0 b]) D
                       (D.inv (η'.counit b') D G f D D.inv (η'.counit b)) D
                       𝗅D-1[G f D G.map0 (F.map0 b)] D
                       𝗋D-1[G f]"

    lemma F_is_locally_essentially_surjective:
    assumes "D.obj b" and "D.obj b'" and "«f : F.map0 b C F.map0 b'»" and "C.ide f"
    shows "«Φ (Ψ f b b') f (F (η'.map0 b' D G f D η0 b)) :
              f C F (η'.map0 b' D G f D η0 b)»"
    proof -
      let ?g = "η'.map0 b' D G f D η0 b"
      have g_in_hhom: "«?g : b D b'»"
        using assms η.ide_map0_obj η.map0_in_hhom by auto
      have ide_g: "D.ide ?g"
        using assms g_in_hhom η.ide_map0_obj η'.ide_map0_obj G.preserves_ide by blast
      let  = "Ψ f b b'"
      have "« : G f D G (F ?g)»"
      proof (intro D.comp_iso_in_hom)
        show "«𝗋D-1[G f] : G f D G f D G.map0 (F.map0 b)»"
          using assms
          by (intro D.iso_in_homI) auto
        show "«𝗅D-1[G f D G.map0 (F.map0 b)] :
                  G f D G.map0 (F.map0 b)
                     D G.map0 (F.map0 b') D G f D G.map0 (F.map0 b)»"
        proof -
          have "D.ide (G f D G.map0 (F.map0 b))"
            using assms D.ide_hcomp [of "G f" "G.map0 (F.map0 b)"]
            by (metis D.hcomp_in_hhomE D.hseqE D.objE F.map0_simps(1) G.map0_simps(1)
                G.preserves_ide GF.map0_simp η.map0_simps(3) g_in_hhom)
          thus ?thesis
            using assms η.ide_map0_obj
            by (intro D.iso_in_homI) auto
        qed
        show "«D.inv (η'.counit b') D G f D D.inv (η'.counit b) :
                 G.map0 (F.map0 b') D G f D G.map0 (F.map0 b)
                   D (η0 b' D η'.map0 b') D G f D (η0 b D η'.map0 b)»"
          using assms η'.iso_counit η'.counit_in_hom(2) D.vconn_implies_hpar(4) D.ide_iso_in_hom
          by (intro D.hcomp_iso_in_hom) auto
        show "«(η0 b' D η'.map0 b') D 𝖺D-1[G f, η0 b, η'.map0 b] :
                 (η0 b' D η'.map0 b') D G f D (η0 b D η'.map0 b)
                   D (η0 b' D η'.map0 b') D (G f D η0 b) D η'.map0 b»"
        proof -
          have "«𝖺D-1[G f, η0 b, η'.map0 b] :
                   G f D (η0 b D η'.map0 b) D (G f D η0 b) D η'.map0 b»"
            using assms η.ide_map0_obj η'.ide_map0_obj η.map0_in_hom η'.map0_in_hom
                  D.assoc'_in_hom D.iso_assoc'
            by (intro D.iso_in_homI) auto
          thus ?thesis
            using assms
            by (intro D.hcomp_iso_in_hom) auto
        qed
        show "«𝖺D-1[η0 b' D η'.map0 b', G f D η0 b, η'.map0 b] :
                (η0 b' D η'.map0 b') D (G f D η0 b) D η'.map0 b
                   D ((η0 b' D η'.map0 b') D (G f D η0 b)) D η'.map0 b»"
          using assms η.ide_map0_obj η'.ide_map0_obj η.map0_in_hom η'.map0_in_hom
                D.assoc'_in_hom D.iso_assoc'
          by (intro D.iso_in_homI) auto
        show "«𝖺D[η0 b', η'.map0 b', G f D η0 b] D η'.map0 b :
                 ((η0 b' D η'.map0 b') D (G f D η0 b)) D η'.map0 b
                   D (η0 b' D ?g) D η'.map0 b»"
          using assms η.ide_map0_obj η'.ide_map0_obj η.map0_in_hom η'.map0_in_hom
                D.assoc_in_hom D.iso_assoc
          by (intro D.hcomp_iso_in_hom D.iso_in_homI) auto
        show "«D.inv (η1 ?g) D η'.map0 b :
                 (η0 b' D ?g) D η'.map0 b D (G (F ?g) D η0 b) D η'.map0 b»"
          using assms η.map1_in_hom [of ?g] η.iso_map1_ide η'.map0_in_hom g_in_hhom ide_g
          by (intro D.hcomp_iso_in_hom D.iso_in_homI) auto
        show "«𝖺D[G (F ?g), η0 b, η'.map0 b] :
                (G (F ?g) D η0 b) D η'.map0 b D G (F ?g) D η0 b D η'.map0 b»"
          using assms η.ide_map0_obj η'.ide_map0_obj η.map0_in_hom η'.map0_in_hom
                D.assoc_in_hom D.iso_assoc
          by (intro D.iso_in_homI) auto
        show "«G (F ?g) D η'.counit b :
                 G (F ?g) D η0 b D η'.map0 b D G (F ?g) D G.map0 (F.map0 b)»"
          using assms η'.counit_in_hom D.ide_in_hom(2) ide_g
          by (intro D.hcomp_iso_in_hom) auto
        show "«𝗋D[G (F ?g)] : G (F ?g) D G.map0 (F.map0 b) D G (F ?g)»"
          using assms ide_g
          by (intro D.iso_in_homI) auto
      qed
      thus "«Φ  f (F ?g) : f C F ?g»"
        using assms ide_g G_reflects_isomorphic [of f "F ?g" ]
        by (metis C.horizontal_homs_axioms D.in_hhomE F.preserves_ide F.preserves_src
            F.preserves_trg g_in_hhom horizontal_homs.in_hhomE)
    qed

    interpretation F: equivalence_pseudofunctor
                        VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC F ΦF
    proof
      show "b. C.obj b  a. D.obj a  C.equivalent_objects (F.map0 a) b"
        by (metis FG.biessentially_surjective_on_objects FG.map0_simp G.map0_simps(1))
      show "g g' μ. D.ide g; D.ide g'; srcD g = srcD g'; trgD g = trgD g';
                       «μ : F g C F g'»
                          ν. «ν : g D g'»  F ν = μ"
      proof -
        fix g g' μ
        assume g: "D.ide g" and g': "D.ide g'"
        assume eq_src: "srcD g = srcD g'" and eq_trg: "trgD g = trgD g'"
        assume μ: "«μ : F g C F g'»"
        obtain ν where ν: "«ν : g D g'»  G (F ν) = G μ"
          using g g' eq_src eq_trg μ GF.locally_full [of g g' "G μ"] by auto
        have "F ν = μ"
          using μ ν G.is_faithful
          by (metis (mono_tags, lifting) C.in_homE F.preserves_hom)
        thus "ν. «ν : g D g'»  F ν = μ"
          using ν by auto
      qed
      show "b b' f. D.obj b; D.obj b'; «f : F.map0 b C F.map0 b'»; C.ide f
                        g. «g : b D b'»  D.ide g  C.isomorphic (F g) f"
      proof -
        fix b b' f
        assume b: "D.obj b" and b': "D.obj b'" and f: "«f : F.map0 b C F.map0 b'»"
        assume ide_f: "C.ide f"
        let ?g = "η'.map0 b' D G f D η0 b"
        have g_in_hhom: "«?g : b D b'»"
          using b b' f η.ide_map0_obj η.map0_in_hhom by auto
        have ide_g: "D.ide ?g"
          using b b' f ide_f g_in_hhom η.ide_map0_obj η'.ide_map0_obj G.preserves_ide by blast
        have "f C F ?g"
          using b b' f ide_f F_is_locally_essentially_surjective C.isomorphic_symmetric
          by (meson C.isomorphicI')
        thus "g. «g : b D b'»  D.ide g  F g C f"
          using g_in_hhom ide_g C.isomorphic_symmetric C.isomorphic_def by blast
      qed
    qed

    lemma equivalence_pseudofunctor_left:
    shows "equivalence_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC F ΦF"
      ..

  end

  subsection "Equivalences Respect Pseudonatural Equivalence"

  text ‹
    In this section we prove that, in an equivalence of bicategories, either pseudofunctor
    may be replaced by a pseudonaturally equivalent one and still obtain an equivalence of
    bicategories.
  ›

  locale equivalence_of_bicategories_and_pseudonatural_equivalence_left =  (* 30 sec *)
    C: bicategory VC HC 𝖺C 𝗂C srcC trgC +
    D: bicategory VD HD 𝖺D 𝗂D srcD trgD +
    E: equivalence_of_bicategories +
    τ: pseudonatural_equivalence
         VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC F ΦF F' ΦF' τ0 τ1
  for F'
  and ΦF'
  and τ0
  and τ1
  (*
   * TODO: If I attempt to declare these free variables with the types they are ultimately
   * inferred to have, then a disjoint set of type variables gets used, resulting in an error.
   *)
  begin

    interpretation GF': composite_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD
                           VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F' ΦF' G ΦG
      ..
    interpretation F'G: composite_pseudofunctor VC HC 𝖺C 𝗂C srcC trgC
                           VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC G ΦG F' ΦF'
      ..
    interpretation τ': converse_pseudonatural_equivalence
                          VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC F ΦF F' ΦF' τ0 τ1
      ..
    interpretation τ'oG: pseudonatural_equivalence_whisker_right VC HC 𝖺C 𝗂C srcC trgC
                          VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                          F' ΦF' F ΦF G ΦG τ'.map0 τ'.map1
      ..
    interpretation Goτ: pseudonatural_equivalence_whisker_left VD HD 𝖺D 𝗂D srcD trgD
                           VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                           F ΦF F' ΦF' G ΦG τ0 τ1
      ..
    sublocale unit: composite_pseudonatural_equivalence
                           VD HD 𝖺D 𝗂D srcD trgD VD HD 𝖺D 𝗂D srcD trgD
                           E.ID.map E.ID.cmp E.GF.map E.GF.cmp GF'.map GF'.cmp
                           η0 η1 Goτ.map0 Goτ.map1
      ..
    sublocale counit: composite_pseudonatural_equivalence
                           VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                           F'G.map F'G.cmp E.FG.map E.FG.cmp E.IC.map E.IC.cmp
                           τ'oG.map0 τ'oG.map1 ε0 ε1
      ..

    definition unit0
    where "unit0  unit.map0"

    definition unit1
    where "unit1  unit.map1"

    definition counit0
    where "counit0  counit.map0"

    definition counit1
    where "counit1  counit.map1"

    sublocale equivalence_of_bicategories
                VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F' ΦF' G ΦG
                unit0 unit1 counit0 counit1
      unfolding unit0_def unit1_def counit0_def counit1_def
      ..

    lemma induces_equivalence_of_bicategories:
    shows "equivalence_of_bicategories
             VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F' ΦF' G ΦG
             unit0 unit1 counit0 counit1"
      ..

  end

  locale equivalence_of_bicategories_and_pseudonatural_equivalence_right =  (* 30 sec *)
    C: bicategory VC HC 𝖺C 𝗂C srcC trgC +
    D: bicategory VD HD 𝖺D 𝗂D srcD trgD +
    E: equivalence_of_bicategories +
    τ: pseudonatural_equivalence
         VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD G ΦG G' ΦG' τ0 τ1
  for G'
  and ΦG'
  and τ0
  and τ1
  begin

    interpretation G'F: composite_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD
                          VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F ΦF G' ΦG'
      ..
    interpretation FG': composite_pseudofunctor VC HC 𝖺C 𝗂C srcC trgC
                          VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC G' ΦG' F ΦF
      ..
    interpretation τ': converse_pseudonatural_equivalence
                          VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD G ΦG G' ΦG' τ0 τ1
      ..
    interpretation τoF: pseudonatural_equivalence_whisker_right VD HD 𝖺D 𝗂D srcD trgD
                          VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                          G ΦG G' ΦG' F ΦF τ0 τ1
      ..
    interpretation Foτ': pseudonatural_equivalence_whisker_left
                          VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                          G' ΦG' G ΦG F ΦF τ'.map0 τ'.map1
      ..
    sublocale unit: composite_pseudonatural_equivalence
                          VD HD 𝖺D 𝗂D srcD trgD VD HD 𝖺D 𝗂D srcD trgD
                          E.ID.map E.ID.cmp E.GF.map E.GF.cmp G'F.map G'F.cmp
                          η0 η1 τoF.map0 τoF.map1
      ..
    sublocale counit: composite_pseudonatural_equivalence
                          VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                          FG'.map FG'.cmp E.FG.map E.FG.cmp E.IC.map E.IC.cmp
                          Foτ'.map0 Foτ'.map1 ε0 ε1
      ..

    definition unit0
    where "unit0  unit.map0"

    definition unit1
    where "unit1  unit.map1"

    definition counit0
    where "counit0  counit.map0"

    definition counit1
    where "counit1  counit.map1"

    sublocale equivalence_of_bicategories
                     VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F ΦF G' ΦG'
                     unit0 unit1 counit0 counit1
      unfolding unit0_def unit1_def counit0_def counit1_def
      ..

    lemma induces_equivalence_of_bicategories:
    shows "equivalence_of_bicategories
             VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F ΦF G' ΦG'
             unit0 unit1 counit0 counit1"
      ..

  end

  subsection "Converse of an Equivalence"

  text ‹
    Equivalence of bicategories is a symmetric notion, in the sense that from an equivalence
    of bicategories from C› to D› we may obtain an equivalence of bicategories from
    D› to C›.  The converse equivalence is obtained by interchanging the pseudofunctors
    F› and G› and replacing the pseudonatural equivalences η› and ε› by converse
    equivalences.  Essentially all the work goes into proving that pseudonatural equivalences
    have pseudonatural converses, which we have already done.
  ›

  locale converse_equivalence_of_bicategories =  (* 25 sec *)
    C: bicategory VC HC 𝖺C 𝗂C srcC trgC +
    D: bicategory VD HD 𝖺D 𝗂D srcD trgD +
    IC: identity_pseudofunctor VC HC 𝖺C 𝗂C srcC trgC +
    ID: identity_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD +
    E: equivalence_of_bicategories
  begin

    sublocale counit: converse_pseudonatural_equivalence
                         VD HD 𝖺D 𝗂D srcD trgD VD HD 𝖺D 𝗂D srcD trgD
                         ID.map ID.cmp E.GF.map E.GF.cmp η0 η1
      ..
    sublocale unit: converse_pseudonatural_equivalence
                      VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                      E.FG.map E.FG.cmp IC.map IC.cmp ε0 ε1
      ..

    definition unit0
    where "unit0  unit.map0"

    definition unit1
    where "unit1  unit.map1"

    definition counit0
    where "counit0  counit.map0"

    definition counit1
    where "counit1  counit.map1"

    sublocale equivalence_of_bicategories
                 VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC G ΦG F ΦF
                 unit0 unit1 counit0 counit1
      unfolding unit0_def unit1_def counit0_def counit1_def
      ..

    lemma is_equivalence_of_bicategories:
    shows "equivalence_of_bicategories
             VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC G ΦG F ΦF
             unit0 unit1 counit0 counit1"
      ..

  end

  subsection "Composition of Equivalences"

  text ‹
    An equivalence of bicategories from B› to C› and an equivalence of bicategories
    from C› to D› may be composed to obtain an equivalence of bicategories
    from B› to D›.
  ›

  locale composite_equivalence_of_bicategories =  (* 60 sec *)
    B: bicategory VB HB 𝖺B 𝗂B srcB trgB +
    C: bicategory VC HC 𝖺C 𝗂C srcC trgC +
    D: bicategory VD HD 𝖺D 𝗂D srcD trgD +
    IB: identity_pseudofunctor VB HB 𝖺B 𝗂B srcB trgB +
    IC: identity_pseudofunctor VC HC 𝖺C 𝗂C srcC trgC +
    ID: identity_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD +
    F_: pseudofunctor VC HC 𝖺C 𝗂C srcC trgC VB HB 𝖺B 𝗂B srcB trgB F ΦF +
    G_: pseudofunctor VB HB 𝖺B 𝗂B srcB trgB VC HC 𝖺C 𝗂C srcC trgC G ΦG +
    H: pseudofunctor VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC H ΦH +
    K: pseudofunctor VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD K ΦK +
    FG: equivalence_of_bicategories
          VB HB 𝖺B 𝗂B srcB trgB VC HC 𝖺C 𝗂C srcC trgC F ΦF G ΦG ρ0 ρ1 σ0 σ1 +
    HK: equivalence_of_bicategories
          VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD H ΦH K ΦK ζ0 ζ1 ξ0 ξ1
  for VB :: "'b comp"                    (infixr "B" 55)
  and HB :: "'b comp"                   (infixr "B" 53)
  and 𝖺B :: "'b  'b  'b  'b"       ("𝖺B[_, _, _]")
  and 𝗂B :: "'b  'b"                   ("𝗂B[_]")
  and srcB :: "'b  'b"
  and trgB :: "'b  'b"
  and VC :: "'c comp"                    (infixr "C" 55)
  and HC :: "'c comp"                   (infixr "C" 53)
  and 𝖺C :: "'c  'c  'c  'c"       ("𝖺C[_, _, _]")
  and 𝗂C :: "'c  'c"                   ("𝗂C[_]")
  and srcC :: "'c  'c"
  and trgC :: "'c  'c"
  and VD :: "'d comp"                    (infixr "D" 55)
  and HD :: "'d comp"                   (infixr "D" 53)
  and 𝖺D :: "'d  'd  'd  'd"       ("𝖺D[_, _, _]")
  and 𝗂D :: "'d  'd"                   ("𝗂D[_]")
  and srcD :: "'d  'd"
  and trgD :: "'d  'd"
  and F :: "'c  'b"
  and ΦF :: "'c * 'c  'b"
  and G :: "'b  'c"
  and ΦG :: "'b * 'b  'c"
  and H :: "'d  'c"
  and ΦH :: "'d * 'd  'c"
  and K :: "'c  'd"
  and ΦK :: "'c * 'c  'd"
  and ρ0 :: "'c  'c"
  and ρ1 :: "'c  'c"
  and σ0 :: "'b  'b"
  and σ1 :: "'b  'b"
  and ζ0 :: "'d  'd"
  and ζ1 :: "'d  'd"
  and ξ0 :: "'c  'c"
  and ξ1 :: "'c  'c"
  begin

    notation B.𝖺'                         ("𝖺B-1[_, _, _]")

    text ‹
      At this point we could make the explicit definitions:
      \begin{itemize}
      \item  η0 = K (ρ0 (H.map0 a)) ⋆D ζ0 a›
      \item  η1 = 𝖺D-1[K (ρ0 (H.map0 (trgD f))), ζ0 (trgD f), f] ⋅D
                       (K (ρ0 (H.map0 (trgD f))) ⋆D ζ1 f) ⋅D
                       𝖺D[K (ρ0 (H.map0 (trgD f))), K (H f), ζ0 (srcD f)] ⋅D
                       (D.inv (ΦK0 (H.map0 (trgD f)), H f)) ⋅D
                          K (ρ1 (H f)) ⋅D
                          ΦK (G (F (H f)), ρ0 (H.map0 (srcD f))) ⋆D ζ0 (srcD f)) ⋅D
                       𝖺D-1[K (G (F (H f))), K (ρ0 (H.map0 (srcD f))), ζ0 (srcD f)]›
      \item  ε0 = σ0 a ⋆B F (ξ0 (G_.map0 a))›
      \item  ε1 = 𝖺B-10 (trgB f), F (ξ0 (G_.map0 (trgB f))), F (H (K (G f)))] ⋅B0 (trgB f) ⋆B
                          B.inv (ΦF0 (G_.map0 (trgB f)), H (K (G f)))) ⋅B
                            F (ξ1 (G f)) ⋅B ΦF (G f, ξ0 (G_.map0 (srcB f)))) ⋅B
                       𝖺B0 (trgB f), F (G f), F (ξ0 (G_.map0 (srcB f)))] ⋅B1 f ⋆B F (ξ0 (G_.map0 (srcB f)))) ⋅B
                       𝖺B-1[f, σ0 (srcB f), F (ξ0 (G_.map0 (srcB f)))]›
      \end{itemize}
      but then it is a daunting task to establish the necessary coherence conditions.
      It is easier (and more useful) to use general results about composite pseudonatural
      equivalences, which are somewhat easier to prove, though long calculations are
      still required for those.
    ›

    sublocale FH: composite_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD
                    VC HC 𝖺C 𝗂C srcC trgC VB HB 𝖺B 𝗂B srcB trgB H ΦH F ΦF
      ..
    sublocale KG: composite_pseudofunctor VB HB 𝖺B 𝗂B srcB trgB
                    VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD G ΦG K ΦK
      ..
    interpretation IG: composite_pseudofunctor VB HB 𝖺B 𝗂B srcB trgB
                         VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                         G ΦG IC.map IC.cmp
      ..
    interpretation IH: composite_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD
                          VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                          H ΦH IC.map IC.cmp
      ..
    interpretation HKG: composite_pseudofunctor VB HB 𝖺B 𝗂B srcB trgB
                          VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                          G ΦG HK.FG.map HK.FG.cmp
      ..
    interpretation GFH: composite_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD
                          VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                          H ΦH FG.GF.map FG.GF.cmp
      ..
    interpretation KGFH: composite_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD
                           VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                           GFH.map GFH.cmp K ΦK
      ..
    interpretation FHKG: composite_pseudofunctor VB HB 𝖺B 𝗂B srcB trgB
                           VC HC 𝖺C 𝗂C srcC trgC VB HB 𝖺B 𝗂B srcB trgB
                           HKG.map HKG.cmp F ΦF
      ..
    interpretation ρoH: pseudonatural_equivalence_whisker_right VD HD 𝖺D 𝗂D srcD trgD
                          VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                          IC.map IC.cmp FG.GF.map FG.GF.cmp H ΦH ρ0 ρ1
      ..
    interpretation KoρoH: pseudonatural_equivalence_whisker_left VD HD 𝖺D 𝗂D srcD trgD
                            VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                            H ΦH GFH.map GFH.cmp K ΦK ρoH.map0 ρoH.map1
    proof -
      interpret KoρoH: pseudonatural_equivalence_whisker_left VD HD 𝖺D 𝗂D srcD trgD
                         VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                          IH.map IH.cmp GFH.map GFH.cmp K ΦK ρoH.map0 ρoH.map1
        ..
      have "IH.map = H"
        using H.is_extensional IH.is_extensional H.functor_axioms by force
      moreover have "IH.cmp = ΦH"
      proof
        fix μν
        show "IH.cmp μν = ΦH μν"
          using IH.cmp_def D.VV.arr_charSbC D.VV.dom_simp D.VV.cod_simp H.FF_def
                C.comp_arr_dom C.comp_cod_arr H.Φ.is_natural_1 H.Φ.is_extensional
          by (cases "D.VV.arr μν") auto
      qed
      ultimately show "pseudonatural_equivalence_whisker_left VD HD 𝖺D 𝗂D srcD trgD
                         VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                         H ΦH GFH.map GFH.cmp K ΦK ρoH.map0 ρoH.map1"
        using KoρoH.pseudonatural_equivalence_whisker_left_axioms by simp
    qed
    interpretation ξoG: pseudonatural_equivalence_whisker_right VB HB 𝖺B 𝗂B srcB trgB
                          VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                          HK.FG.map HK.FG.cmp IC.map IC.cmp G ΦG ξ0 ξ1
      ..
    interpretation FoξoG: pseudonatural_equivalence_whisker_left VB HB 𝖺B 𝗂B srcB trgB
                            VC HC 𝖺C 𝗂C srcC trgC VB HB 𝖺B 𝗂B srcB trgB
                            HKG.map HKG.cmp G ΦG F ΦF ξoG.map0 ξoG.map1
    proof -
      interpret FoξoG: pseudonatural_equivalence_whisker_left VB HB 𝖺B 𝗂B srcB trgB
                         VC HC 𝖺C 𝗂C srcC trgC VB HB 𝖺B 𝗂B srcB trgB 
                         HKG.map HKG.cmp IG.map IG.cmp F ΦF ξoG.map0 ξoG.map1
        ..
      have "IG.map = G"
        using G_.is_extensional IG.is_extensional
        by (meson G_.functor_axioms comp_identity_functor)
      moreover have "IG.cmp = ΦG"
      proof
        fix μν
        show "IG.cmp μν = ΦG μν"
          using IG.cmp_def B.VV.arr_charSbC B.VV.dom_simp B.VV.cod_simp G_.FF_def
                C.comp_arr_dom C.comp_cod_arr G_.Φ.is_natural_1 G_.Φ.is_extensional
          by (cases "B.VV.arr μν") auto
      qed
      ultimately show "pseudonatural_equivalence_whisker_left VB HB 𝖺B 𝗂B srcB trgB
                         VC HC 𝖺C 𝗂C srcC trgC VB HB 𝖺B 𝗂B srcB trgB
                         HKG.map HKG.cmp G ΦG F ΦF ξoG.map0 ξoG.map1"
        using FoξoG.pseudonatural_equivalence_whisker_left_axioms by simp
    qed

    sublocale unit: composite_pseudonatural_equivalence
                      VD HD 𝖺D 𝗂D srcD trgD VD HD 𝖺D 𝗂D srcD trgD
                      ID.map ID.cmp HK.GF.map HK.GF.cmp KGFH.map KGFH.cmp
                      ζ0 ζ1 KoρoH.map0 KoρoH.map1
      ..
    sublocale counit: composite_pseudonatural_equivalence
                        VB HB 𝖺B 𝗂B srcB trgB VB HB 𝖺B 𝗂B srcB trgB
                        FHKG.map FHKG.cmp FG.FG.map FG.FG.cmp IB.map IB.cmp
                        FoξoG.map0 FoξoG.map1 σ0 σ1
      ..

    abbreviation left_map
    where "left_map  FH.map"

    abbreviation right_map
    where "right_map  KG.map"

    abbreviation left_cmp
    where "left_cmp  FH.cmp"

    abbreviation right_cmp
    where "right_cmp  KG.cmp"

    definition unit0
    where "unit0  unit.map0"

    definition unit1
    where "unit1  unit.map1"

    definition counit0
    where "counit0  counit.map0"

    definition counit1
    where "counit1 == counit.map1"

    lemma unit0_simp:
    assumes "D.obj a"
    shows "unit0 a = K (ρ0 (H.map0 a)) D ζ0 a"
      using assms unit0_def unit.map0_def KoρoH.map0_def ρoH.map0_def by simp

    lemma unit1_simp:
    assumes "D.ide f"
    shows "unit1 f = 𝖺D-1[K (ρ0 (H.map0 (trgD f))), ζ0 (trgD f), f] D
                     (K (ρ0 (H.map0 (trgD f))) D ζ1 f) D
                     𝖺D[K (ρ0 (H.map0 (trgD f))), K (H f), ζ0 (srcD f)] D
                     (D.inv (ΦK (ρ0 (H.map0 (trgD f)), H f)) D
                     K (ρ1 (H f)) D
                     ΦK (G (F (H f)), ρ0 (H.map0 (srcD f))) D ζ0 (srcD f)) D
                     𝖺D-1[K (G (F (H f))), K (ρ0 (H.map0 (srcD f))), ζ0 (srcD f)]"
      using assms unit1_def unit.map1_def KoρoH.map0_def ρoH.map0_def KoρoH.map1_def ρoH.map1_def
      by simp

    lemma counit0_simp:
    assumes "B.obj a"
    shows "counit0 a = σ0 a B F (ξ0 (G_.map0 a))"
      using assms counit0_def counit.map0_def FoξoG.map0_def ξoG.map0_def by simp

    lemma counit1_simp:
    assumes "B.ide f"
    shows "counit1 f = 𝖺B-1[σ0 (trgB f), F (ξ0 (G_.map0 (trgB f))), F (H (K (G f)))] B
                       (σ0 (trgB f) B
                          B.inv (ΦF (ξ0 (G_.map0 (trgB f)), H (K (G f)))) B
                          F (ξ1 (G f)) B
                          ΦF (G f, ξ0 (G_.map0 (srcB f)))) B
                       𝖺B[σ0 (trgB f), F (G f), F (ξ0 (G_.map0 (srcB f)))] B
                       (σ1 f B F (ξ0 (G_.map0 (srcB f)))) B
                       𝖺B-1[f, σ0 (srcB f), F (ξ0 (G_.map0 (srcB f)))]"
      using assms counit1_def counit.map1_def FoξoG.map0_def FoξoG.map1_def
            ξoG.map0_def ξoG.map1_def
      by simp

    sublocale equivalence_of_bicategories VB HB 𝖺B 𝗂B srcB trgB VD HD 𝖺D 𝗂D srcD trgD
                FH.map FH.cmp KG.map KG.cmp unit0 unit1 counit0 counit1
    proof -
      interpret FH_KG: composite_pseudofunctor
                          VB HB 𝖺B 𝗂B srcB trgB VD HD 𝖺D 𝗂D srcD trgD VB HB 𝖺B 𝗂B srcB trgB
                          KG.map KG.cmp FH.map FH.cmp
        ..
      interpret KG_FH: composite_pseudofunctor
                         VD HD 𝖺D 𝗂D srcD trgD VB HB 𝖺B 𝗂B srcB trgB VD HD 𝖺D 𝗂D srcD trgD
                         FH.map FH.cmp KG.map KG.cmp
        ..
      have "FH_KG.map = FHKG.map" by auto
      moreover have "FH_KG.cmp = FHKG.cmp"
      proof
        fix μν
        show "FH_KG.cmp μν = FHKG.cmp μν"
        proof (cases "B.VV.arr μν")
          case False
          thus ?thesis
            using FH_KG.Φ.is_extensional FHKG.Φ.is_extensional by simp
          next
          case True
          have "FH_KG.cmp μν =
                  F (H (K (G (IB.cmp μν)))) B
                  F (H (K (G (B.dom (fst μν) B B.dom (snd μν))))) B
                  F (H (K (ΦG (B.dom (fst μν), B.dom (snd μν))))) B
                  F (H (ΦK (G (B.dom (fst μν)), G (B.dom (snd μν))))) B
                  (F (H (K (G (B.dom (fst μν))) D K (G (B.dom (snd μν))))) B
                  F (ΦH (K (G (B.dom (fst μν))), K (G (B.dom (snd μν)))))) B
                  ΦF (H (K (G (B.dom (fst μν)))), H (K (G (B.dom (snd μν)))))"
            using True FH_KG.cmp_def FHKG.cmp_def HKG.cmp_def
                  FH.cmp_def HK.FG.cmp_def KG.cmp_def
                  B.VV.arr_charSbC B.VV.dom_simp B.VV.cod_simp
                  C.VV.arr_charSbC C.VV.dom_simp C.VV.cod_simp
                  D.VV.arr_charSbC D.VV.dom_simp D.VV.cod_simp
                  G_.FF_def H.FF_def K.FF_def B.comp_assoc
            by auto
          also have "... = F (H (K (G (IB.cmp μν)))) B
                           F (H (K (G (B.dom (fst μν) B B.dom (snd μν))))) B
                           F (H (K (ΦG (B.dom (fst μν), B.dom (snd μν))))) B
                           F (H (ΦK (G (B.dom (fst μν)), G (B.dom (snd μν))))) B
                           F (ΦH (K (G (B.dom (fst μν))), K (G (B.dom (snd μν))))) B
                           ΦF (H (K (G (B.dom (fst μν)))), H (K (G (B.dom (snd μν)))))"
          proof -
            have "F (H (K (G (B.dom (fst μν))) D K (G (B.dom (snd μν))))) B
                  F (ΦH (K (G (B.dom (fst μν))), K (G (B.dom (snd μν))))) =
                    F (ΦH (K (G (B.dom (fst μν))), K (G (B.dom (snd μν)))))"
              using True B.VV.arr_charSbC B.VV.dom_simp B.VV.cod_simp
                    C.VV.arr_charSbC C.VV.dom_simp C.VV.cod_simp
                    D.VV.arr_charSbC D.VV.dom_simp D.VV.cod_simp
                    B.comp_cod_arr
              by auto
            thus ?thesis
              using D.comp_assoc by simp
          qed
          also have "... = F (H (K (G (IB.cmp μν)))) B
                           F (H (K (G (B.dom (fst μν) B B.dom (snd μν))))) B
                           F (H (K (ΦG (B.dom (fst μν), B.dom (snd μν))))) B
                           (F (H (K (G (B.dom (fst μν)) C G (B.dom (snd μν))))) B
                           F (H (ΦK (G (B.dom (fst μν)), G (B.dom (snd μν)))))) B
                           F (ΦH (K (G (B.dom (fst μν))), K (G (B.dom (snd μν))))) B
                           ΦF (H (K (G (B.dom (fst μν)))), H (K (G (B.dom (snd μν)))))"
          proof -
            have "F (H (K (G (B.dom (fst μν)) C G (B.dom (snd μν))))) B
                  F (H (ΦK (G (B.dom (fst μν)), G (B.dom (snd μν))))) =
                    F (H (ΦK (G (B.dom (fst μν)), G (B.dom (snd μν)))))"
              using True B.VV.arr_charSbC B.VV.dom_simp B.VV.cod_simp
                    C.VV.arr_charSbC C.VV.dom_simp C.VV.cod_simp
                    D.VV.arr_charSbC D.VV.dom_simp D.VV.cod_simp
                    G_.FF_def H.FF_def K.FF_def B.comp_assoc
                    B.comp_cod_arr
              by auto
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = FHKG.cmp μν"
            using True FH_KG.cmp_def FHKG.cmp_def HKG.cmp_def
                  FH.cmp_def HK.FG.cmp_def KG.cmp_def
                  B.VV.arr_charSbC B.VV.dom_simp B.VV.cod_simp
                  C.VV.arr_charSbC C.VV.dom_simp C.VV.cod_simp
                  D.VV.arr_charSbC D.VV.dom_simp D.VV.cod_simp
                  G_.FF_def H.FF_def K.FF_def B.comp_assoc
            by simp
          finally show ?thesis by blast
        qed
      qed
      moreover have "KG_FH.map = KGFH.map" by auto
      moreover have "KG_FH.cmp = KGFH.cmp"
      proof
        fix μν
        show "KG_FH.cmp μν = KGFH.cmp μν"
        proof (cases "D.VV.arr μν")
          case False
          thus ?thesis
            using KG_FH.Φ.is_extensional KGFH.Φ.is_extensional by simp
          next
          case True
          have "KG_FH.cmp μν =
                  K (G (F (H (ID.cmp μν)))) D
                  K (G (F (H (D.dom (fst μν) D D.dom (snd μν))))) D
                  K (G (F (ΦH (D.dom (fst μν), D.dom (snd μν))))) D
                  K (G (ΦF (H (D.dom (fst μν)), H (D.dom (snd μν))))) D
                  (K (G (F (H (D.dom (fst μν))) B F (H (D.dom (snd μν))))) D
                  K (ΦG (F (H (D.dom (fst μν))), F (H (D.dom (snd μν)))))) D
                  ΦK (G (F (H (D.dom (fst μν)))), G (F (H (D.dom (snd μν)))))"
            using True KG_FH.cmp_def FH.cmp_def KG.cmp_def
                  B.VV.arr_charSbC B.VV.dom_simp B.VV.cod_simp
                  C.VV.arr_charSbC C.VV.dom_simp C.VV.cod_simp
                  D.VV.arr_charSbC D.VV.dom_simp D.VV.cod_simp
                  G_.FF_def H.FF_def K.FF_def
                  D.comp_assoc
            by auto
          also have "... = K (G (F (H (ID.cmp μν)))) D
                           K (G (F (H (D.dom (fst μν) D D.dom (snd μν))))) D
                           K (G (F (ΦH (D.dom (fst μν), D.dom (snd μν))))) D
                           K (G (ΦF (H (D.dom (fst μν)), H (D.dom (snd μν))))) D
                           K (ΦG (F (H (D.dom (fst μν))), F (H (D.dom (snd μν))))) D
                           ΦK (G (F (H (D.dom (fst μν)))), G (F (H (D.dom (snd μν)))))"
          proof -
            have "K (G (F (H (D.dom (fst μν))) B F (H (D.dom (snd μν))))) D
                  K (ΦG (F (H (D.dom (fst μν))), F (H (D.dom (snd μν))))) =
                    K (ΦG (F (H (D.dom (fst μν))), F (H (D.dom (snd μν)))))"
              using True KG_FH.cmp_def FH.cmp_def KG.cmp_def
                    B.VV.arr_charSbC B.VV.dom_simp B.VV.cod_simp
                    C.VV.arr_charSbC C.VV.dom_simp C.VV.cod_simp
                    D.VV.arr_charSbC D.VV.dom_simp D.VV.cod_simp
                    D.comp_cod_arr
              by simp
            thus ?thesis
              using D.comp_assoc by simp
          qed
          also have "... = K (G (F (H (ID.cmp μν)))) D
                           K (G (F (H (D.dom (fst μν) D D.dom (snd μν))))) D
                           K (G (F (ΦH (D.dom (fst μν), D.dom (snd μν))))) D
                           (K (G (F (H (D.dom (fst μν)) C H (D.dom (snd μν))))) D
                           K (G (ΦF (H (D.dom (fst μν)), H (D.dom (snd μν)))))) D
                           K (ΦG (F (H (D.dom (fst μν))), F (H (D.dom (snd μν))))) D
                           ΦK (G (F (H (D.dom (fst μν)))), G (F (H (D.dom (snd μν)))))"
          proof -
            have "K (G (F (H (D.dom (fst μν)) C H (D.dom (snd μν))))) D
                  K (G (ΦF (H (D.dom (fst μν)), H (D.dom (snd μν))))) =
                    K (G (ΦF (H (D.dom (fst μν)), H (D.dom (snd μν)))))"
              using True KGFH.cmp_def GFH.cmp_def FG.GF.cmp_def
                    B.VV.arr_charSbC B.VV.dom_simp B.VV.cod_simp
                    C.VV.arr_charSbC C.VV.dom_simp C.VV.cod_simp
                    D.VV.arr_charSbC D.VV.dom_simp D.VV.cod_simp
                    D.comp_cod_arr
              by simp
            thus ?thesis
              using D.comp_assoc by simp
          qed
          also have "... = KGFH.cmp μν"
            using True KGFH.cmp_def GFH.cmp_def FG.GF.cmp_def
                  B.VV.arr_charSbC B.VV.dom_simp B.VV.cod_simp
                  C.VV.arr_charSbC C.VV.dom_simp C.VV.cod_simp
                  D.VV.arr_charSbC D.VV.dom_simp D.VV.cod_simp
                  F_.FF_def G_.FF_def H.FF_def K.FF_def
                  D.comp_assoc
            by auto
          finally show ?thesis by blast
        qed
      qed
      ultimately show "equivalence_of_bicategories
                         VB HB 𝖺B 𝗂B srcB trgB VD HD 𝖺D 𝗂D srcD trgD
                         FH.map FH.cmp KG.map KG.cmp unit0 unit1 counit0 counit1"
        unfolding unit0_def unit1_def counit0_def counit1_def
        using B.bicategory_axioms D.bicategory_axioms FH.pseudofunctor_axioms
              KG.pseudofunctor_axioms unit.pseudonatural_equivalence_axioms
              counit.pseudonatural_equivalence_axioms IB.identity_pseudofunctor_axioms
              ID.identity_pseudofunctor_axioms FH_KG.composite_pseudofunctor_axioms
              KG_FH.composite_pseudofunctor_axioms
        unfolding equivalence_of_bicategories_def by simp
    qed

    lemma is_equivalence_of_bicategories:
    shows "equivalence_of_bicategories VB HB 𝖺B 𝗂B srcB trgB VD HD 𝖺D 𝗂D srcD trgD
             left_map left_cmp right_map right_cmp unit0 unit1 counit0 counit1"
      ..

  end

  subsection "Equivalence with a Dense Sub-bicategory"

  text ‹
    The purpose of this section is to show that, given a bicategory B› and a sub-bicategory
    defined by a ``dense'' set of objects of B›, the embedding of the sub-bicategory in B›
    extends to an equivalence of bicategories.  Here by ``dense'' we mean that every object
    of B› is equivalent to some object of the subbicategory.
  ›

  locale dense_subbicategory =
    B: bicategory VB HB 𝖺B 𝗂B srcB trgB +
    subbicategory VB HB 𝖺B 𝗂B srcB trgB λμ. B.arr μ  srcB μ  Obj  trgB μ  Obj
  for VB :: "'b comp"                    (infixr "B" 55)
  and HB :: "'b comp"                    (infixr "B" 53)
  and 𝖺B :: "'b  'b  'b  'b"        ("𝖺B[_, _, _]")
  and 𝗂B :: "'b  'b"                    ("𝗂B[_]")
  and srcB :: "'b  'b"
  and trgB :: "'b  'b"
  and Obj :: "'b set" +
  assumes dense: "a. B.obj a  a'. a'  Obj  B.equivalent_objects a' a"
  begin

    notation B.𝖺'         ("𝖺B-1[_, _, _]")
    notation B.lunit      ("𝗅B[_]")
    notation B.lunit'     ("𝗅B-1[_]")
    notation B.runit      ("𝗋B[_]")
    notation B.runit'     ("𝗋B-1[_]")

    notation comp         (infixr "" 55)
    notation hcomp        (infixr "" 53)
    notation in_hom       ("«_ : _  _»")
    notation in_hhom      ("«_ : _  _»")
    notation 𝖺            ("𝖺[_, _, _]")
    notation 𝖺'           ("𝖺-1[_, _, _]")
    notation lunit        ("𝗅[_]")
    notation lunit'       ("𝗅-1[_]")
    notation runit        ("𝗋[_]")
    notation runit'       ("𝗋-1[_]")

    abbreviation (input) Arr
    where "Arr  λμ. B.arr μ  srcB μ  Obj  trgB μ  Obj"

    sublocale emb: embedding_pseudofunctor VB HB 𝖺B 𝗂B srcB trgB Arr
      ..

    abbreviation E
    where "E  emb.map"

    abbreviation ΦE
    where "ΦE  emb.cmp"

    text ‹
      We define a projection P› by transporting arrows of B› across chosen
      equivalences between objects of B› and objects of the sub-bicategory.
    ›

    definition P0
    where "P0 a  SOME a'. obj a'  B.equivalent_objects a' a"

    lemma P0_props:
    assumes "B.obj a"
    shows "obj (P0 a)"
    and "B.equivalent_objects (P0 a) a"
    and "B.equivalent_objects a a'  P0 a = P0 a'"
    and "P0 (P0 a) = P0 a"
    and "B.obj (P0 a)"
    and "P0 a  Obj"
    proof -
      have "a'. obj a'  B.equivalent_objects a' a"
        using assms dense [of a] obj_char arr_charSbC
        by (metis (no_types, lifting) B.equivalent_objects_def B.in_hhomE B.obj_def B.obj_simps(3))
      hence 1: "obj (P0 a)  B.equivalent_objects (P0 a) a"
        unfolding P0_def
        using someI_ex [of "λa'. obj a'  B.equivalent_objects a' a"] by blast
      show "obj (P0 a)"
        using 1 by simp
      show 2: "B.equivalent_objects (P0 a) a"
        using 1 by simp
      show 3: "a'. B.equivalent_objects a a'  P0 a = P0 a'"
        unfolding P0_def
        using B.equivalent_objects_symmetric B.equivalent_objects_transitive by meson
      show "P0 (P0 a) = P0 a"
        using 2 3 [of "P0 a"] B.equivalent_objects_symmetric by auto
      show "B.obj (P0 a)"
        using 1 B.equivalent_objects_def by auto
      thus "P0 a  Obj"
        using 1 3 obj_char arr_charSbC by auto
    qed

    text ‹
      For each object a› of B›, we choose an adjoint equivalence from a› to P0 a›.
      The use of adjoint equivalences is necessary in order to establish the required
      coherence conditions.
    ›

    definition e
    where "e a = (SOME e. «e : a B P0 a» 
                   (d η ε. adjoint_equivalence_in_bicategory
                              VB HB 𝖺B 𝗂B srcB trgB e d η ε))"

    definition d
    where "d a = (SOME d. η ε. adjoint_equivalence_in_bicategory
                                  VB HB 𝖺B 𝗂B srcB trgB (e a) d η ε)"

    definition η
    where "η a = (SOME η. ε. adjoint_equivalence_in_bicategory
                                VB HB 𝖺B 𝗂B srcB trgB (e a) (d a) η ε)"

    definition ε
    where "ε a = (SOME ε. adjoint_equivalence_in_bicategory
                            VB HB 𝖺B 𝗂B srcB trgB (e a) (d a) (η a) ε)"

    lemma chosen_adjoint_equivalence:
    assumes "B.obj a"
    shows "adjoint_equivalence_in_bicategory
             VB HB 𝖺B 𝗂B srcB trgB (e a) (d a) (η a) (ε a)"
    and "«e a : a B P0 a»" and "B.ide (d a)" and "B.ide (e a)" and "B.iso (η a)" and "B.iso (ε a)"
    proof -
      have "e. «e : a B P0 a» 
                (d η ε. adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB e d η ε)"
      proof -
        obtain e where e: "«e : a B P0 a»  B.equivalence_map e"
          using assms P0_props(2) B.equivalent_objects_symmetric B.equivalent_objects_def
          by meson
        obtain d η ε where d: "adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB e d η ε"
          using e B.equivalence_map_extends_to_adjoint_equivalence by blast
        thus ?thesis
          using e d by auto
      qed
      hence 1: "«e a : a B P0 a» 
                adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB (e a) (d a) (η a) (ε a)"
        using d_def e_def η_def ε_def arr_charSbC
            someI_ex [of "λe. «e : a B P0 a» 
                          (d η ε. adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB e d η ε)"]
            someI_ex [of "λd. (η ε. adjoint_equivalence_in_bicategory
                                       VB HB 𝖺B 𝗂B srcB trgB (e a) d η ε)"]
            someI_ex [of "λη. ε. adjoint_equivalence_in_bicategory
                                    VB HB 𝖺B 𝗂B srcB trgB (e a) (d a) η ε"]
            someI_ex [of "λε. adjoint_equivalence_in_bicategory
                                VB HB 𝖺B 𝗂B srcB trgB (e a) (d a) (η a) ε"]
        by simp
      interpret adjoint_equivalence_in_bicategory
                  VB HB 𝖺B 𝗂B srcB trgB e a d a η a ε a
        using 1 by simp
      show "adjoint_equivalence_in_bicategory
              VB HB 𝖺B 𝗂B srcB trgB (e a) (d a) (η a) (ε a)" ..
      show "«e a : a B P0 a»" using 1 by simp
      show "B.ide (d a)" by simp
      show "B.ide (e a)" by simp
      show "B.iso (η a)" by simp
      show "B.iso (ε a)" by simp
    qed

    lemma equivalence_data_in_homB [intro]:
    assumes "B.obj a"
    shows "«e a : a B P0 a»" and "«d a : P0 a B a»"
    and "«e a : e a B e a»" and "«d a : d a B d a»"
    and "«η a : a B a»" and "«ε a : P0 a B P0 a»"
    and "«η a : a B d a B e a»" and "«ε a : e a B d a B P0 a»"
    proof -
      interpret adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB e a d a η a ε a
        using assms chosen_adjoint_equivalence by simp
      show e: "«e a : a B P0 a»"
        using assms chosen_adjoint_equivalence by simp
      show "«d a : P0 a B a»" using e antipar by auto
      show "«e a : e a B e a»" by auto
      show "«d a : d a B d a»" by auto
      show "«η a : a B a»" using unit_in_hom e by auto
      show "«ε a : P0 a B P0 a»" using counit_in_hom e by auto
      show "«η a : a B d a B e a»" using unit_in_hom e by auto
      show "«ε a : e a B d a B P0 a»" using counit_in_hom e by auto
    qed

    lemma equivalence_data_simpsB [simp]:
    assumes "B.obj a"
    shows "B.ide (d a)" and "B.ide (e a)" and "B.iso (η a)" and "B.iso (ε a)"
    and "srcB (e a) = a" and "trgB (e a) = P0 a" and "srcB (d a) = P0 a" and "trgB (d a) = a"
    and "B.dom (e a) = e a" and "B.cod (e a) = e a"
    and "B.dom (d a) = d a" and "B.cod (d a) = d a"
    and "srcB (η a) = a" and "trgB (η a) = a" and "srcB (ε a) = P0 a" and "trgB (ε a) = P0 a"
    and "B.dom (η a) = a" and "B.cod (η a) = d a B e a"
    and "B.dom (ε a) = e a B d a" and "B.cod (ε a) = P0 a"
      using assms chosen_adjoint_equivalence equivalence_data_in_homB B.in_hhom_def
                       apply auto
      by (meson B.in_homE)+

    lemma equivalence_data_in_hom [intro]:
    assumes "obj a"
    shows "«e a : a  P0 a»" and "«d a : P0 a  a»"
    and "«e a : e a  e a»" and "«d a : d a  d a»"
    and "«η a : a  a»" and "«ε a : P0 a  P0 a»"
    and "«η a : a  d a  e a»" and "«ε a : e a  d a  P0 a»"
    proof -
      have a: "B.obj a  a  Obj"
        using assms obj_char arr_charSbC by auto
      have P0a: "obj (P0 a)  P0 a  Obj"
        using assms a P0_props [of a] arr_charSbC obj_char by auto
      show ea: "«e a : a  P0 a»"
        using a P0a arr_charSbC chosen_adjoint_equivalence(2) src_def trg_def by auto
      show "«e a : e a  e a»"
        using a ea chosen_adjoint_equivalence(4) ide_charSbC in_hom_charSbC by auto
      show da: "«d a : P0 a  a»"
        using a P0a arr_charSbC src_def trg_def P0_props(1) equivalence_data_in_homB(2)
        by auto
      show "«d a : d a  d a»"
        using a da chosen_adjoint_equivalence(3) ide_charSbC in_hom_charSbC by auto
      show ηa: "«η a : a  d a  e a»"
      proof
        show 1: "arr (η a)"
          using assms a P0a da ea arr_charSbC equivalence_data_in_homB(7) hcomp_closed
          by (metis (no_types, lifting) equivalence_data_in_homB(5) B.in_hhom_def)
        show "dom (η a) = a"
          using a 1 dom_charSbC equivalence_data_in_homB(7) by auto
        show "cod (η a) = d a  e a"
          using a 1 da ea cod_charSbC in_hom_charSbC equivalence_data_in_homB(7) hcomp_def
          by (metis (no_types, lifting) hcomp_char in_hhomE B.in_homE)
      qed
      show εa: "«ε a : e a  d a  P0 a»"
      proof
        show 1: "arr (ε a)"
          using assms a P0a da ea arr_charSbC equivalence_data_in_homB(8) hcomp_closed
          by (metis (no_types, lifting) equivalence_data_in_homB(6) B.in_hhom_def)
        show "dom (ε a) = e a  d a"
          using a 1 da ea dom_charSbC in_hom_charSbC equivalence_data_in_homB(8) hcomp_def
          by (metis (no_types, lifting) hcomp_char in_hhomE B.in_homE)
        show "cod (ε a) = P0 a"
          using a 1 cod_charSbC equivalence_data_in_homB(8) by auto
      qed
      show "«η a : a  a»"
        using assms ηa src_def trg_def src_dom trg_dom
        by (metis (no_types, lifting) in_hhom_def in_hom_charSbC obj_simps(2-3)
            vconn_implies_hpar(1-2))
      show "«ε a : P0 a  P0 a»"
        using P0a εa src_def trg_def src_cod trg_cod
        by (metis (no_types, lifting) in_hhom_def in_hom_charSbC obj_simps(2-3)
            vconn_implies_hpar(1-4))
    qed

    lemma equivalence_data_simps [simp]:
    assumes "obj a"
    shows "ide (d a)" and "ide (e a)" and "iso (η a)" and "iso (ε a)"
    and "src (e a) = a" and "trg (e a) = P0 a" and "src (d a) = P0 a" and "trg (d a) = a"
    and "dom (e a) = e a" and "cod (e a) = e a"
    and "dom (d a) = d a" and "cod (d a) = d a"
    and "src (η a) = a" and "trg (η a) = a" and "src (ε a) = P0 a" and "trg (ε a) = P0 a"
    and "dom (η a) = a" and "cod (η a) = d a B e a"
    and "dom (ε a) = e a B d a" and "cod (ε a) = P0 a"
      (* These can be done in a one-liner, but it takes 20 sec. *)
      using assms equivalence_data_in_hom(2) ide_charSbC obj_char apply auto[1]
      using assms equivalence_data_in_hom(1) ide_charSbC obj_char apply auto[1]
      using assms obj_char arr_charSbC iso_charSbC B.iso_is_arr apply auto[1]
      using assms obj_char arr_charSbC iso_charSbC P0_props B.iso_is_arr apply auto[1]
      using assms obj_char arr_charSbC P0_props src_def apply auto[1]
      using assms obj_char arr_charSbC P0_props trg_def apply auto[1]
      using assms obj_char arr_charSbC P0_props src_def apply auto[1]
      using assms obj_char arr_charSbC P0_props trg_def apply auto[1]
      using assms obj_char arr_charSbC dom_charSbC P0_props apply auto[1]
      using assms obj_char arr_charSbC cod_charSbC P0_props apply auto[1]
      using assms obj_char arr_charSbC dom_charSbC P0_props apply auto[1]
      using assms obj_char arr_charSbC cod_charSbC P0_props apply auto[1]
      using assms obj_char arr_charSbC B.iso_is_arr src_def apply auto[1]
      using assms obj_char arr_charSbC B.iso_is_arr trg_def apply auto[1]
      using assms obj_char arr_charSbC dom_charSbC P0_props B.iso_is_arr src_def apply auto[1]
      using assms obj_char arr_charSbC cod_charSbC P0_props B.iso_is_arr trg_def apply auto[1]
      using assms obj_char arr_charSbC dom_charSbC B.iso_is_arr apply auto[1]
      using assms obj_char arr_charSbC cod_charSbC B.iso_is_arr apply auto[1]
      using assms obj_char arr_charSbC dom_charSbC P0_props B.iso_is_arr apply auto[1]
      using assms obj_char arr_charSbC cod_charSbC P0_props B.iso_is_arr by auto[1]

    definition P
    where "P μ = e (trgB μ) B μ B d (srcB μ)"

    lemma P_in_homB [intro]:
    assumes "B.arr μ"
    shows "«P μ : P0 (srcB μ) B P0 (trgB μ)»"
    and "«P μ : P (B.dom μ) B P (B.cod μ)»"
      unfolding P_def using assms by auto

    lemma P_simpsB [simp]:
    assumes "B.arr μ"
    shows "B.arr (P μ)"
    and "srcB (P μ) = P0 (srcB μ)" and "trgB (P μ) = P0 (trgB μ)"
    and "B.dom (P μ) = P (B.dom μ)" and "B.cod (P μ) = P (B.cod μ)"
      using assms P_in_homB by blast+

    lemma P_in_hom [intro]:
    assumes "B.arr μ"
    shows "«P μ : P0 (srcB μ)  P0 (trgB μ)»"
    and "«P μ : P (B.dom μ)  P (B.cod μ)»"
    proof -
      show 1: "«P μ : P0 (srcB μ)  P0 (trgB μ)»"
        using assms P_in_homB(1) arr_charSbC src_def trg_def P0_props(1)
        by (metis (no_types, lifting) in_hhom_def obj_char B.in_hhomE B.obj_simps(2)
            B.obj_src B.obj_trg)
      show "«P μ : P (B.dom μ)  P (B.cod μ)»"
        using assms 1 dom_charSbC cod_charSbC
        by (intro in_homI) auto
    qed

    lemma P_simps [simp]:
    assumes "B.arr μ"
    shows "arr (P μ)"
    and "src (P μ) = P0 (srcB μ)" and "trg (P μ) = P0 (trgB μ)"
    and "dom (P μ) = P (B.dom μ)" and "cod (P μ) = P (B.cod μ)"
      using assms P_in_hom by blast+

    interpretation P: "functor" VB comp P
    proof
      show "μ. ¬ B.arr μ  P μ = null"
        using P_def null_char B.hseq_char B.hseq_char' by auto
      have 0: "μ. B.arr μ  B.arr (P μ)"
        by simp
      show 1: "μ. B.arr μ  arr (P μ)"
        using P_simpsB(1) by simp
      show 2: "μ. B.arr μ  dom (P μ) = P (B.dom μ)"
        using 1 dom_simp by auto
      show 3: "μ. B.arr μ  cod (P μ) = P (B.cod μ)"
        using 1 cod_simp by auto
      show "μ ν. B.seq μ ν  P (μ B ν) = P μ  P ν"
      proof -
        fix μ ν
        assume seq: "B.seq μ ν"
        show "P (μ B ν) = P μ  P ν"
        proof -
          have 4: "P (μ B ν) = e (trgB μ) B μ B ν B d (srcB μ)"
            unfolding P_def using seq B.src_vcomp B.trg_vcomp by simp
          also have "... = P μ B P ν"
            unfolding P_def
            using seq 0 4 B.whisker_left B.whisker_right
            by (metis equivalence_data_simpsB(1-2) B.hseqE B.obj_trg B.seqE B.src_vcomp
                B.vseq_implies_hpar(2))
          also have "... = P μ  P ν"
            using seq 1 seq_charSbC comp_char by auto
          finally show ?thesis by blast
        qed
      qed
    qed

    interpretation faithful_functor VB comp P
    proof
      fix μ μ'
      assume par: "B.par μ μ'"
      have 1: "srcB μ = srcB μ'  trgB μ = trgB μ'"
        using par by (metis B.src_dom B.trg_dom)
      assume eq: "P μ = P μ'"
      interpret src: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                       e (srcB μ) d (srcB μ) η (srcB μ) ε (srcB μ)
        using par chosen_adjoint_equivalence by simp
      interpret trg: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                       e (trgB μ) d (trgB μ) η (trgB μ) ε (trgB μ)
        using par chosen_adjoint_equivalence by simp
      show "μ = μ'"
        using eq par 1
        unfolding P_def
        using B.equivalence_cancel_left [of "e (trgB μ)" "μ B d (srcB μ)" "μ' B d (srcB μ')"]
              B.equivalence_cancel_right [of "d (srcB μ)" μ μ']
              B.equivalence_map_def src.dual_equivalence trg.equivalence_in_bicategory_axioms
        by auto
    qed

    interpretation P: weak_arrow_of_homs VB srcB trgB comp src trg P
    proof
      fix μ
      assume μ: "B.arr μ"
      interpret src: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                       e (srcB μ) d (srcB μ) η (srcB μ) ε (srcB μ)
        using μ chosen_adjoint_equivalence by simp
      show "isomorphic (P (srcB μ)) (src (P μ))"
      proof (unfold isomorphic_def)
        show "f. «f : P (srcB μ)  src (P μ)»  iso f"
        proof -
          let  = "ε (srcB μ) B (e (srcB μ) B 𝗅B[d (srcB μ)])"
          have "« : P (srcB μ)  src (P μ)»  iso "
          proof -
            have 1: "« : P (srcB μ) B P0 (srcB μ)»  B.iso "
              unfolding P_def using μ by auto
            moreover have "arr   arr (B.inv )"
              using μ 1 arr_charSbC P0_props(1)
              by (metis (no_types, lifting) P0_props(6) B.arrI B.arr_inv equivalence_data_simpsB(16)
                  B.obj_src src.counit_simps(4-5) B.src_inv B.src_vcomp
                  B.trg_inv B.trg_vcomp)
            moreover have "P0 (srcB μ) = src (P μ)"
              using μ by simp
            ultimately show ?thesis
              using μ in_hom_charSbC arr_charSbC iso_charSbC
              by (metis (no_types, lifting) src_closed equivalence_data_simpsB(6)
                  B.obj_src P.preserves_arr src.counit_simps(4) B.src.preserves_arr B.src_vcomp)
          qed
          thus ?thesis by blast
        qed
      qed
      interpret trg: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                       e (trgB μ) d (trgB μ) η (trgB μ) ε (trgB μ)
        using μ chosen_adjoint_equivalence by simp
      show "isomorphic (P (trgB μ)) (trg (P μ))"
      proof (unfold isomorphic_def)
        show "f. «f : P (trgB μ)  trg (P μ)»  iso f"
        proof -
          let  = "ε (trgB μ) B (e (trgB μ) B 𝗅B[d (trgB μ)])"
          have "« : P (trgB μ)  trg (P μ)»  iso "
          proof -
            have 1: "« : P (trgB μ) B P0 (trgB μ)»  B.iso "
              unfolding P_def using μ by auto
            moreover have "arr   arr (B.inv )"
              using μ 1 arr_charSbC P0_props(1)
              by (metis (no_types, lifting) obj_char B.arrI B.arr_inv B.obj_trg B.src_inv B.trg_inv
                  B.vconn_implies_hpar(1-4))
            moreover have "arr (P (trgB μ))  arr (P0 (trgB μ))"
              using μ arr_charSbC P0_props(1) P_simps(1) obj_char by auto
            moreover have "P0 (trgB μ) = trg (P μ)"
              using μ by simp
            ultimately show ?thesis
              using in_hom_charSbC iso_charSbC by force
          qed
          thus ?thesis by blast
        qed
      qed
    qed

    text ‹
      The following seems to be needed to avoid non-confluent simplifications,
      \emph{e.g.}~of S.src (P μ)› to P.map0 a› and to P0 a›.
    ›

    lemma P_map0_simp [simp]:
    assumes "B.obj a"
    shows "P.map0 a = P0 a"
      using assms P.map0_def B.obj_simps(1-2) by simp

    interpretation HoPP: composite_functor B.VV.comp VV.comp comp
                           P.FF λμν. hcomp (fst μν) (snd μν)
      ..
    interpretation PoH: composite_functor B.VV.comp VB comp (λμν. fst μν B snd μν) P
      ..

    no_notation B.in_hom  ("«_ : _ B _»")

    definition CMP
    where "CMP f g  (e (trgB f) B 𝖺B-1[f, g, d (srcB g)]) B
                      (e (trgB f) B f B
                          𝗅B[g B d (srcB g)] B (B.inv (η (trgB g)) B g B d (srcB g))) B
                      (e (trgB f) B f B 𝖺B-1[d (srcB f), e (trgB g), g B d (srcB g)]) B
                      𝖺B[e (trgB f), f, d (srcB f) B P g] B
                      𝖺B[e (trgB f) B f, d (srcB f), P g] B
                      (𝖺B-1[e (trgB f), f, d (srcB f)] B P g)"

    text ‹
      The 2-cell CMP f g› has the right type to be a compositor for a pseudofunctor
      whose underlying mapping is P›.
    ›

    lemma CMP_in_hom [intro]:
    assumes "B.ide f" and "B.ide g" and "srcB f = trgB g"
    shows "«CMP f g : P0 (srcB g)  P0 (trgB f)»"
    and "«CMP f g : P f  P g  P (f B g)»"
    and "«CMP f g : P0 (srcB g) B P0 (trgB f)»"
    and "«CMP f g : P f B P g B P (f B g)»"
    proof -
      show 1: "«CMP f g : P f B P g B P (f B g)»"
        using assms
        by (unfold P_def CMP_def, intro B.comp_in_homI' B.seqI B.hseqI') auto
      show 2: "«CMP f g : P0 (srcB g) B P0 (trgB f)»"
        using assms 1 B.src_cod B.trg_cod B.vconn_implies_hpar(1-2) by auto
      show 3: "«CMP f g : P f  P g  P (f B g)»"
        using assms 1 B.arrI B.vconn_implies_hpar(1-2) P_simps(1) arr_charSbC hcomp_eqI
              hseqI' in_hom_charSbC
        by force
      show "«CMP f g : P0 (srcB g)  P0 (trgB f)»"
        using 2 3 arr_charSbC src_def trg_def by fastforce
    qed

    lemma CMP_simps [simp]:
    assumes "B.ide f" and "B.ide g" and "srcB f = trgB g"
    shows "arr (CMP f g)"
    and "src (CMP f g) = P0 (srcB g)" and "trg (CMP f g) = P0 (trgB f)"
    and "dom (CMP f g) = P f  P g" and "cod (CMP f g) = P (f B g)"
    and "B.arr (CMP f g)"
    and "srcB (CMP f g) = P0 (srcB g)" and "trgB (CMP f g) = P0 (trgB f)"
    and "B.dom (CMP f g) = P f B P g" and "B.cod (CMP f g) = P (f B g)"
      using assms CMP_in_hom [of f g] by auto

    lemma iso_CMP:
    assumes "B.ide f" and "B.ide g" and "srcB f = trgB g"
    shows "iso (CMP f g)"
    and "B.iso (CMP f g)"
    proof -
      show "B.iso (CMP f g)"
        unfolding CMP_def P_def
        using assms B.VV.arr_charSbC P.as_nat_iso.components_are_iso
        by (intro B.isos_compose B.iso_hcomp) auto
      thus "iso (CMP f g)"
        using assms iso_charSbC arr_charSbC CMP_simps(1) by auto
    qed

    abbreviation (input) SRC
    where "SRC μ  d (srcB μ) B e (srcB μ)"

    abbreviation (input) TRG
    where "TRG μ  d (trgB μ) B e (trgB μ)"

    definition LUNIT
    where "LUNIT f  𝗅B[f] B (B.inv (η (trgB f)) B f)"

    definition RUNIT
    where "RUNIT f  𝗋B[f] B (f B B.inv (η (srcB f)))"

    text ‹
      Here we prove a series of results that would be automatic if we had some notion of
      ``bicategory with SRC› and TRG› as alternative source and target''.
      Perhaps this idea can be developed in future work and used to simplify the overall
      development.
    ›

    lemma LUNIT_in_hom [intro]:
    assumes "B.ide f"
    shows "«LUNIT f : srcB f B trgB f»"
    and "«LUNIT f : TRG f B f B f»"
    proof -
      interpret e_trg_f: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                            e (trgB f) d (trgB f) η (trgB f) ε (trgB f)
        using assms chosen_adjoint_equivalence by simp
      show "«LUNIT f : srcB f B trgB f»"
        unfolding LUNIT_def
        using assms e_trg_f.unit_is_iso by auto
      show "«LUNIT f : TRG f B f B f»"
        unfolding LUNIT_def
        using assms e_trg_f.unit_is_iso by auto
    qed

    lemma LUNIT_simps [simp]:
    assumes "B.ide f"
    shows "B.arr (LUNIT f)"
    and "srcB (LUNIT f) = srcB f" and "trgB (LUNIT f) = trgB f"
    and "B.dom (LUNIT f) = TRG f B f"
    and "B.cod (LUNIT f) = f"
      using assms LUNIT_in_hom by auto

    lemma RUNIT_in_hom [intro]:
    assumes "B.ide f"
    shows "«RUNIT f : srcB f B trgB f»"
    and "«RUNIT f : f B SRC f B f»"
    proof -
      interpret e_src_f: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                            e (srcB f) d (srcB f) η (srcB f) ε (srcB f)
        using assms chosen_adjoint_equivalence by simp
      show "«RUNIT f : srcB f B trgB f»"
        unfolding RUNIT_def
        using assms e_src_f.unit_is_iso by auto
      show "«RUNIT f : f B SRC f B f»"
        unfolding RUNIT_def
        using assms e_src_f.unit_is_iso by auto
    qed

    lemma RUNIT_simps [simp]:
    assumes "B.ide f"
    shows "B.arr (RUNIT f)"
    and "srcB (RUNIT f) = srcB f" and "trgB (RUNIT f) = trgB f"
    and "B.dom (RUNIT f) = f B SRC f"
    and "B.cod (RUNIT f) = f"
      using assms RUNIT_in_hom by auto

    lemma iso_LUNIT:
    assumes "B.ide f"
    shows "B.iso (LUNIT f)"
      by (simp add: assms B.isos_compose LUNIT_def)

    lemma iso_RUNIT:
    assumes "B.ide f"
    shows "B.iso (RUNIT f)"
      by (simp add: assms B.isos_compose RUNIT_def)

    lemma LUNIT_naturality:
    assumes "B.arr μ"
    shows "μ B LUNIT (B.dom μ) = LUNIT (B.cod μ) B (TRG μ B μ)"
    proof -
      interpret e_trg_μ: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                            e (trgB μ) d (trgB μ) η (trgB μ) ε (trgB μ)
        using assms chosen_adjoint_equivalence by simp
      show ?thesis
      proof -
        have "μ B LUNIT (B.dom μ) = (μ B 𝗅B[B.dom μ]) B (B.inv (η (trgB μ)) B B.dom μ)"
          unfolding LUNIT_def
          using assms B.comp_assoc by simp
        also have "... = 𝗅B[B.cod μ] B (trgB μ B μ) B (B.inv (η (trgB μ)) B B.dom μ)"
          using assms B.lunit_naturality B.comp_assoc by simp
        also have "... = 𝗅B[B.cod μ] B (B.inv (η (trgB μ)) B μ)"
          using assms B.interchange [of "trgB μ" "B.inv (η (trgB μ))" μ "B.dom μ"]
                e_trg_μ.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp
        also have "... = 𝗅B[B.cod μ] B (B.inv (η (trgB μ)) B B.cod μ) B (TRG μ B μ)"
          using assms B.interchange [of "B.inv (η (trgB μ))" "TRG μ" "B.cod μ" μ]
                e_trg_μ.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp
        also have "... = LUNIT (B.cod μ) B (TRG μ B μ)"
          unfolding LUNIT_def
          using assms B.comp_assoc by simp
        finally show ?thesis by simp
      qed
    qed

    lemma RUNIT_naturality:
    assumes "B.arr μ"
    shows "μ B RUNIT (B.dom μ) = RUNIT (B.cod μ) B (μ B SRC μ)"
    proof -
      interpret e_src_μ: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                            e (srcB μ) d (srcB μ) η (srcB μ) ε (srcB μ)
        using assms chosen_adjoint_equivalence by simp
      show ?thesis
      proof -
        have "μ B RUNIT (B.dom μ) =
              (μ B 𝗋B[B.dom μ]) B (B.dom μ B B.inv (η (srcB μ)))"
          unfolding RUNIT_def
          using assms B.comp_assoc by simp
        also have "... = 𝗋B[B.cod μ] B (μ B srcB μ) B (B.dom μ B B.inv (η (srcB μ)))"
          using assms B.runit_naturality B.comp_assoc by simp
        also have "... = 𝗋B[B.cod μ] B (μ B B.inv (η (srcB μ)))"
          using assms B.interchange [of μ "B.dom μ" "srcB μ" "B.inv (η (srcB μ))"]
                e_src_μ.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp
        also have "... = 𝗋B[B.cod μ] B (B.cod μ B B.inv (η (srcB μ))) B (μ B SRC μ)"
          using assms B.interchange [of "B.cod μ" μ "B.inv (η (srcB μ))" "SRC μ"]
                e_src_μ.unit_is_iso B.comp_arr_dom B.comp_cod_arr
          by simp
        also have "... = RUNIT (B.cod μ) B (μ B SRC μ)"
          unfolding RUNIT_def
          using assms B.comp_assoc by simp
        finally show ?thesis by simp
      qed
    qed

    lemma LUNIT_hcomp:
    assumes "B.ide f" and "B.ide g" and "srcB f = trgB g"
    shows "LUNIT (f B g) B 𝖺B[d (trgB f) B e (trgB f), f, g] = LUNIT f B g"
    proof -
      interpret e_trg_f: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                            e (trgB f) d (trgB f) η (trgB f) ε (trgB f)
        using assms chosen_adjoint_equivalence by simp
      have "LUNIT (f B g) B 𝖺B[TRG f, f, g] =
            𝗅B[f B g] B (B.inv (η (trgB f)) B f B g) B 𝖺B[TRG f, f, g]"
        unfolding LUNIT_def
        using assms B.comp_assoc by simp
      also have "... = (𝗅B[f B g] B 𝖺B[trgB f, f, g]) B ((B.inv (η (trgB f)) B f) B g)"
        using assms B.assoc_naturality [of "B.inv (η (trgB f))" f g] e_trg_f.unit_is_iso
              B.comp_assoc
        by simp
      also have "... = (𝗅B[f] B g) B ((B.inv (η (trgB f)) B f) B g)"
        using assms B.lunit_hcomp by simp
      also have "... = LUNIT f B g"
        using assms LUNIT_def LUNIT_simps(1) B.whisker_right by auto
      finally show ?thesis by simp
    qed

    lemma RUNIT_hcomp:
    assumes "B.ide f" and "B.ide g" and "srcB f = trgB g"
    shows "RUNIT (f B g) = (f B RUNIT g) B 𝖺B[f, g, SRC g]"
    proof -
      interpret e_src_g: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                            e (srcB g) d (srcB g) η (srcB g) ε (srcB g)
        using assms chosen_adjoint_equivalence by simp
      have "(f B RUNIT g) B 𝖺B[f, g, SRC g] =
            (f B 𝗋B[g]) B (f B g B B.inv (η (srcB g))) B 𝖺B[f, g, SRC g]"
        unfolding RUNIT_def
        using assms B.whisker_left e_src_g.unit_is_iso B.comp_assoc by simp
      also have "... = ((f B 𝗋B[g]) B 𝖺B[f, g, srcB g]) B ((f B g) B B.inv (η (srcB g)))"
        using assms B.assoc_naturality [of f g "B.inv (η (srcB g))"] e_src_g.unit_is_iso
              B.comp_assoc
        by simp
      also have "... = 𝗋B[f B g] B ((f B g) B B.inv (η (srcB g)))"
        using assms B.runit_hcomp by simp
      also have "... = RUNIT (f B g)"
        using assms RUNIT_def by simp
      finally show ?thesis by simp
    qed

    lemma TRIANGLE:
    assumes "B.ide f" and "B.ide g" and "srcB f = trgB g"
    shows "(f B LUNIT g) B 𝖺B[f, SRC f, g] = RUNIT f B g"
    proof -
      interpret e_trg_g: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                            e (trgB g) d (trgB g) η (trgB g) ε (trgB g)
        using assms chosen_adjoint_equivalence by simp
      show ?thesis
      proof -
        have "(f B LUNIT g) B 𝖺B[f, SRC f, g] =
              (f B 𝗅B[g]) B (f B B.inv (η (trgB g)) B g) B 𝖺B[f, SRC f, g]"
          using assms B.whisker_left e_trg_g.unit_is_iso LUNIT_def LUNIT_simps(1) B.comp_assoc
          by auto
        also have "... = ((f B 𝗅B[g]) B 𝖺B[f, srcB f, g]) B ((f B B.inv (η (trgB g))) B g)"
          using assms B.assoc_naturality [of f "B.inv (η (trgB g))" g] e_trg_g.unit_is_iso
                B.comp_assoc
          by auto
        also have "... = (𝗋B[f] B g) B ((f B B.inv (η (trgB g))) B g)"
          using assms B.triangle by simp
        also have "... = RUNIT f B g"
          using assms B.whisker_right e_trg_g.unit_is_iso RUNIT_def RUNIT_simps
          by metis
        finally show ?thesis by simp
      qed
    qed

    text ‹
      The CMP f g› also satisfy the naturality conditions required of compositors.
    ›

    lemma CMP_naturality:
    assumes "B.arr μ" and "B.arr ν" and "srcB μ = trgB ν"
    shows "CMP (B.cod μ) (B.cod ν) B (P μ B P ν)
             = P (μ B ν) B CMP (B.dom μ) (B.dom ν)"
    proof -
      let ?a = "srcB ν"
      let ?b = "srcB μ"
      let ?c = "trgB μ"
      let ?f = "B.dom μ"
      let ?g = "B.cod μ"
      let ?h = "B.dom ν"
      let ?k = "B.cod ν"
      have "CMP ?g ?k B (P μ B P ν)
              = (e ?c B 𝖺B-1[?g, ?k, d ?a]) B
                (e ?c B ?g B LUNIT (?k B d ?a)) B
                (e ?c B ?g B 𝖺B-1[d ?b, e ?b, ?k B d ?a]) B
                𝖺B[e ?c, ?g, d ?b B P ?k] B
                𝖺B[e ?c B ?g, d ?b, P ?k] B
                (𝖺B-1[e ?c, ?g, d ?b] B P ?k) B
                (P μ B P ν)"
        unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp
      also have "... = (e ?c B 𝖺B-1[?g, ?k, d ?a]) B
                       (e ?c B ?g B LUNIT (?k B d ?a)) B
                       (e ?c B ?g B 𝖺B-1[d ?b, e ?b, ?k B d ?a]) B
                       𝖺B[e ?c, ?g, d ?b B P ?k] B
                       (𝖺B[e ?c B ?g, d ?b, P ?k] B
                       (((e ?c B μ) B d ?b) B P ν)) B
                       (𝖺B-1[e ?c, ?f, d ?b] B P ?h)"
      proof -
        have "(𝖺B-1[e ?c, ?g, d ?b] B P ?k) B (P μ B P ν)
                = 𝖺B-1[e ?c, ?g, d ?b] B P μ B P ?k B P ν"
          using assms P_def B.interchange by fastforce
        also have
          "... = ((e ?c B μ) B d ?b) B 𝖺B-1[e ?c, ?f, d ?b] B P ?k B P ν"
          using assms P_def B.assoc'_naturality [of "e ?c" μ "d ?b"] by simp
        also have
          "... = ((e ?c B μ) B d ?b) B 𝖺B-1[e ?c, ?f, d ?b] B P ν B P ?h"
          using assms B.comp_arr_dom B.comp_cod_arr B.src_cod B.src_dom
                B.trg_cod B.trg_dom P.as_nat_trans.naturality
          by simp
        also have "... = (((e ?c B μ) B d ?b) B P ν) B (𝖺B-1[e ?c, ?f, d ?b] B P ?h)"
          using assms B.interchange by auto
        finally have "(𝖺B-1[e ?c, ?g, d ?b] B P ?k) B (P μ B P ν)
                        = (((e ?c B μ) B d ?b) B P ν) B (𝖺B-1[e ?c, ?f, d ?b] B P ?h)"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?c B 𝖺B-1[?g, ?k, d ?a]) B
                       (e ?c B ?g B LUNIT (?k B d ?a)) B
                       (e ?c B ?g B 𝖺B-1[d ?b, e ?b, ?k B d ?a]) B
                       (𝖺B[e ?c, ?g, d ?b B P ?k] B
                       ((e ?c B μ) B d ?b B P ν)) B
                       𝖺B[e ?c B ?f, d ?b, P ?h] B
                       (𝖺B-1[e ?c, ?f, d ?b] B P ?h)"
        using assms B.assoc_naturality [of "e ?c B μ" "d ?b" "P ν"] B.comp_assoc by simp
      also have "... = (e ?c B 𝖺B-1[?g, ?k, d ?a]) B
                       (e ?c B ?g B LUNIT (?k B d ?a)) B
                       ((e ?c B ?g B 𝖺B-1[d ?b, e ?b, ?k B d ?a]) B
                       (e ?c B μ B d ?b B P ν)) B
                       𝖺B[e ?c, ?f, d ?b B P ?h] B
                       𝖺B[e ?c B ?f, d ?b, P ?h] B
                       (𝖺B-1[e ?c, ?f, d ?b] B P ?h)"
        using assms B.assoc_naturality [of "e ?c" μ "d ?b B P ν"] B.comp_assoc by simp
      also have "... = (e ?c B 𝖺B-1[?g, ?k, d ?a]) B
                       ((e ?c B ?g B LUNIT (?k B d ?a)) B
                       (e ?c B μ B SRC μ B ν B d ?a)) B
                       (e ?c B ?f B 𝖺B-1[d ?b, e ?b, ?h B d ?a]) B
                       𝖺B[e ?c, ?f, d ?b B P ?h] B
                       𝖺B[e ?c B ?f, d ?b, P ?h] B
                       (𝖺B-1[e ?c, ?f, d ?b] B P ?h)"
      proof -
        have
          "(e ?c B ?g B 𝖺B-1[d ?b, e ?b, ?k B d ?a]) B (e ?c B μ B d ?b B P ν)
             = (e ?c B μ B TRG ν B ν B d ?a) B
               (e ?c B ?f B 𝖺B-1[d ?b, e ?b, ?h B d ?a])"
        proof -
          have "(e ?c B ?g B 𝖺B-1[d ?b, e ?b, ?k B d ?a]) B (e ?c B μ B d ?b B P ν)
                  = e ?c B (?g B 𝖺B-1[d ?b, e ?b, ?k B d ?a]) B (μ B d ?b B P ν)"
            using assms P_def B.whisker_left by simp
          also have "... = e ?c B μ B 𝖺B-1[d ?b, e ?b, ?k B d ?a] B (d ?b B P ν)"
            using assms P_def B.comp_cod_arr
                  B.interchange [of "?g" μ "𝖺B-1[d ?b, e ?b, ?k B d ?a]" "d ?b B P ν"]
            by simp
          also have "... = e ?c B μ B
                             (TRG ν B ν B d ?a) B 𝖺B-1[d ?b, e ?b, ?h B d ?a]"
            using assms P_def B.assoc'_naturality [of "d ?b" "e ?b" "ν B d ?a"]
            by simp
          also have "... = e ?c B
                             (μ B TRG ν B ν B d ?a) B (?f B 𝖺B-1[d ?b, e ?b, ?h B d ?a])"
            using assms B.comp_arr_dom
                  B.interchange [of μ "?f" "TRG ν B ν B d ?a""𝖺B-1[d ?b, e ?b, ?h B d ?a]"]
            by fastforce
          also have
            "... = (e ?c B μ B TRG ν B ν B d ?a) B
                   (e ?c B ?f B 𝖺B-1[d ?b, e ?b, ?h B d ?a])"
            using assms B.whisker_left by simp
          finally show ?thesis by simp
        qed
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = ((e ?c B 𝖺B-1[?g, ?k, d ?a]) B
                       (e ?c B μ B ν B d ?a)) B
                       (e ?c B ?f B LUNIT (?h B d ?a)) B
                       (e ?c B ?f B 𝖺B-1[d ?b, e ?b, ?h B d ?a]) B
                       𝖺B[e ?c, ?f, d ?b B P ?h] B
                       𝖺B[e ?c B ?f, d ?b, P ?h] B
                       (𝖺B-1[e ?c, ?f, d ?b] B P ?h)"
      proof -
        have "((e ?c B ?g B LUNIT (?k B d ?a)) B (e ?c B μ B TRG ν B ν B d ?a))
                = e ?c B μ B LUNIT (B.cod (ν B d ?a)) B ((d ?b B e ?b) B ν B d ?a)"
          using assms comp_arr_dom B.comp_cod_arr B.whisker_left
                B.interchange [of "?g" μ "LUNIT (?k B d ?a)" "(d ?b B e ?b) B ν B d ?a"]
          by fastforce
        also have "... = e ?c B μ B (ν B d ?a) B LUNIT (?h B d ?a)"
          using assms LUNIT_naturality [of "ν B d ?a"] by simp
        also have "... = (e ?c B μ B ν B d ?a) B (e ?c B ?f B LUNIT (?h B d ?a))"
          using assms B.comp_arr_dom B.comp_cod_arr B.whisker_left
                B.interchange [of μ "?f" "ν B d ?a" "LUNIT (?h B d ?a)"]
          by simp
        finally have "((e ?c B ?g B LUNIT (?k B d ?a)) B
                      (e ?c B μ B TRG ν B ν B d ?a))
                        = (e ?c B μ B ν B d ?a) B (e ?c B ?f B LUNIT (?h B d ?a))"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = P (μ B ν) B
                       (e ?c B 𝖺B-1[?f, ?h, d ?a]) B
                       (e ?c B ?f B LUNIT (?h B d ?a)) B
                       (e ?c B ?f B 𝖺B-1[d ?b, e ?b, ?h B d ?a]) B
                       𝖺B[e ?c, ?f, d ?b B P ?h] B
                       𝖺B[e ?c B ?f, d ?b, P ?h] B
                       (𝖺B-1[e ?c, ?f, d ?b] B P ?h)"
      proof -
        have "(e ?c B 𝖺B-1[?g, ?k, d ?a]) B (e ?c B μ B ν B d ?a)
                = P (μ B ν) B (e ?c B 𝖺B-1[?f, ?h, d ?a])"
        proof -
          have "(e ?c B 𝖺B-1[?g, ?k, d ?a]) B (e ?c B μ B ν B d ?a)
                  = e ?c B 𝖺B-1[?g, ?k, d ?a] B (μ B ν B d ?a)"
            using assms B.whisker_left by simp
          also have "... = e ?c B ((μ B ν) B d ?a) B 𝖺B-1[?f, ?h, d ?a]"
            using assms B.assoc'_naturality [of μ ν "d ?a"] by simp
          also have "... = P (μ B ν) B (e ?c B 𝖺B-1[?f, ?h, d ?a])"
            using assms P_def B.whisker_left by simp
          finally show ?thesis by simp
        qed
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = P (μ B ν) B CMP ?f ?h"
        unfolding CMP_def LUNIT_def using assms by simp
      finally show ?thesis by simp
    qed

    interpretation EV: self_evaluation_map VB HB 𝖺B 𝗂B srcB trgB ..
    notation EV.eval ("_")

    abbreviation (input) SRCt  ("SRC")
    where "SRC μ  d (srcB μ)  e (srcB μ)"

    abbreviation (input) TRGt  ("TRG")
    where "TRG μ  d (trgB μ)  e (trgB μ)"

    abbreviation (input) PRJt  ("PRJ")
    where "PRJ μ  e (trgB μ)  μ  d (srcB μ)"

    text ‹
      The CMP f g› satisfy the coherence conditions with respect to associativity that are
      required of compositors.
    ›

    lemma CMP_coherence:
    assumes "B.ide f" and "B.ide g" and "B.ide h" and "srcB f = trgB g" and "srcB g = trgB h"
    shows "P 𝖺B[f, g, h] B CMP (f B g) h B (CMP f g B P h)
             = CMP f (g B h) B (P f B CMP g h) B 𝖺B[P f, P g, P h]"
    proof -
      text ‹
        The overall strategy of the proof is to expand the definition of CMP› on the
        left and right-hand sides, then permute the occurrences of LUNIT› and
        RUNIT› to the left ends of both the left-hand side and right-hand side of the
        equation to be proved, so that the initial portions of these expressions become
        identical and the remaining parts to the right consist only of canonical isomorphisms.
        Then the Coherence Theorem is applied to prove syntactically (and automatically) that the
        canonical parts are equal, which implies equality of the complete expressions.
        The rest is just grinding through the calculations.
      ›
      let ?a = "trgB f"
      let ?b = "trgB g"
      let ?c = "trgB h"
      let ?d = "srcB h"
      have "P 𝖺B[f, g, h] B CMP (f B g) h B (CMP f g B P h)
              = P 𝖺B[f, g, h] B
                (e ?a B 𝖺B-1[f B g, h, d ?d]) B
                (e ?a B (f B g) B LUNIT (h B d ?d)) B
                (e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                𝖺B[e ?a, f B g, d ?c B P h] B
                𝖺B[e ?a B f B g, d ?c, P h] B
                (𝖺B-1[e ?a, f B g, d ?c] B P h) B
                ((e ?a B 𝖺B-1[f, g, d ?c]) B
                 (e ?a B f B LUNIT (g B d ?c)) B
                 (e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B
                 𝖺B[e ?a, f, d ?b B P g] B
                 𝖺B[e ?a B f, d ?b, P g] B
                 (𝖺B-1[e ?a, f, d ?b] B P g)
                    B P h)"
        unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp
      also have "... = P 𝖺B[f, g, h] B
                       (e ?a B 𝖺B-1[f B g, h, d ?d]) B
                       (e ?a B (f B g) B LUNIT (h B d ?d)) B
                       (e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, f B g, d ?c B P h] B
                       𝖺B[e ?a B f B g, d ?c, P h] B
                       (𝖺B-1[e ?a, f B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f, g, d ?c]) B P h) B
                       ((e ?a B f B LUNIT (g B d ?c)) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
        using assms B.whisker_right P_def by simp (* 6 sec *)
      also have "... = P 𝖺B[f, g, h] B
                       ((e ?a B 𝖺B-1[f B g, h, d ?d]) B
                       (e ?a B (f B g) B LUNIT h B d ?d)) B
                       (e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, f B g, d ?c B P h] B
                       𝖺B[e ?a B f B g, d ?c, P h] B
                       (𝖺B-1[e ?a, f B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f, g, d ?c]) B P h) B
                       ((e ?a B f B LUNIT (g B d ?c)) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "e ?a B (f B g) B LUNIT (h B d ?d)
                = e ?a B (f B g) B (LUNIT h B d ?d) B 𝖺B-1[TRG h, h, d ?d]"
          using assms LUNIT_hcomp [of h "d ?d"] B.invert_side_of_triangle by simp
        also have "... = (e ?a B (f B g) B LUNIT h B d ?d) B
                         (e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d])"
          using assms B.whisker_left by simp
        finally have "e ?a B (f B g) B LUNIT (h B d ?d)
                        = (e ?a B (f B g) B LUNIT h B d ?d) B
                          (e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d])"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (P 𝖺B[f, g, h] B
                       (e ?a B ((f B g) B LUNIT h) B d ?d)) B
                       (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, f B g, d ?c B P h] B
                       𝖺B[e ?a B f B g, d ?c, P h] B
                       (𝖺B-1[e ?a, f B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f, g, d ?c]) B P h) B
                       ((e ?a B f B LUNIT (g B d ?c)) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "(e ?a B 𝖺B-1[f B g, h, d ?d]) B (e ?a B (f B g) B LUNIT h B d ?d)
                = e ?a B 𝖺B-1[f B g, h, d ?d] B ((f B g) B LUNIT h B d ?d)"
          using assms B.whisker_left by simp
        also have "... = e ?a B (((f B g) B LUNIT h) B d ?d) B
                                 𝖺B-1[f B g, TRG h B h, d ?d]"
          using assms B.assoc'_naturality [of "f B g" "LUNIT h" "d ?d"] by simp
        also have "... = (e ?a B ((f B g) B LUNIT h) B d ?d) B
                         (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d])"
          using assms B.whisker_left by simp
        finally have "(e ?a B 𝖺B-1[f B g, h, d ?d]) B
                      (e ?a B (f B g) B LUNIT h B d ?d)
                        = (e ?a B ((f B g) B LUNIT h) B d ?d) B
                          (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d])"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B[f, g, TRG h B h] B d ?d) B
                       (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, f B g, d ?c B P h] B
                       𝖺B[e ?a B f B g, d ?c, P h] B
                       (𝖺B-1[e ?a, f B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f, g, d ?c]) B P h) B
                       ((e ?a B f B LUNIT (g B d ?c)) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "P 𝖺B[f, g, h] B (e ?a B ((f B g) B LUNIT h) B d ?d)
                = e ?a B 𝖺B[f, g, h] B ((f B g) B LUNIT h) B d ?d"
          using assms B.whisker_left B.whisker_right P_def by simp
        also have "... = e ?a B (f B g B LUNIT h) B 𝖺B[f, g, TRG h B h] B d ?d"
          using assms B.assoc_naturality [of f g "LUNIT h"] by simp
        also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                         (e ?a B 𝖺B[f, g, TRG h B h] B d ?d)"
          using assms B.whisker_left B.whisker_right by simp
        finally have "P 𝖺B[f, g, h] B (e ?a B ((f B g) B LUNIT h) B d ?d)
                        = (e ?a B (f B g B LUNIT h) B d ?d) B
                          (e ?a B 𝖺B[f, g, TRG h B h] B d ?d)"
          by simp
        thus ?thesis using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B[f, g, TRG h B h] B d ?d) B
                       (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, f B g, d ?c B P h] B
                       𝖺B[e ?a B f B g, d ?c, P h] B
                       (𝖺B-1[e ?a, f B g, d ?c] B P h) B
                       (((e ?a B 𝖺B-1[f, g, d ?c]) B P h) B
                       ((e ?a B RUNIT f B g B d ?c) B P h)) B
                       ((e ?a B 𝖺B-1[f, TRG g, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "((e ?a B f B LUNIT (g B d ?c)) B P h)
                = (e ?a B (RUNIT f B g B d ?c) B 𝖺B-1[f, TRG g, g B d ?c]) B P h"
          using assms TRIANGLE [of f "g B d ?c"] B.invert_side_of_triangle by simp
        also have "... = ((e ?a B RUNIT f B g B d ?c) B P h) B
                         ((e ?a B 𝖺B-1[f, TRG g, g B d ?c]) B P h)"
          using assms B.whisker_left B.whisker_right P_def by simp
        finally have "((e ?a B f B LUNIT (g B d ?c)) B P h)
                        = ((e ?a B RUNIT f B g B d ?c) B P h) B
                          ((e ?a B 𝖺B-1[f, TRG g, g B d ?c]) B P h)"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B[f, g, TRG h B h] B d ?d) B
                       (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, f B g, d ?c B P h] B
                       𝖺B[e ?a B f B g, d ?c, P h] B
                       ((𝖺B-1[e ?a, f B g, d ?c] B P h) B
                       ((e ?a B (RUNIT f B g) B d ?c) B P h)) B
                       ((e ?a B 𝖺B-1[f B TRG g, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, TRG g, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "((e ?a B 𝖺B-1[f, g, d ?c]) B P h) B ((e ?a B RUNIT f B g B d ?c) B P h)
                = (e ?a B 𝖺B-1[f, g, d ?c] B (RUNIT f B g B d ?c)) B P h"
          using assms B.whisker_left B.whisker_right P_def by simp
        also have "... = (e ?a B
                           ((RUNIT f B g) B d ?c) B 𝖺B-1[f B TRG g, g, d ?c])
                              B P h"
          using assms B.assoc'_naturality [of "RUNIT f" g "d ?c"] by simp
        also have "... = ((e ?a B (RUNIT f B g) B d ?c) B P h) B
                         ((e ?a B 𝖺B-1[f B TRG g, g, d ?c]) B P h)"
          using assms B.whisker_left B.whisker_right P_def by simp
        finally have "((e ?a B 𝖺B-1[f, g, d ?c]) B P h) B
                      ((e ?a B RUNIT f B g B d ?c) B P h)
                        = ((e ?a B (RUNIT f B g) B d ?c) B P h) B
                          ((e ?a B 𝖺B-1[f B TRG g, g, d ?c]) B P h)"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B[f, g, TRG h B h] B d ?d) B
                       (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, f B g, d ?c B P h] B
                       (𝖺B[e ?a B f B g, d ?c, P h] B
                       (((e ?a B (RUNIT f B g)) B d ?c) B P h)) B
                       (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "(𝖺B-1[e ?a, f B g, d ?c] B P h) B ((e ?a B (RUNIT f B g) B d ?c) B P h)
                = 𝖺B-1[e ?a, f B g, d ?c] B (e ?a B (RUNIT f B g) B d ?c) B P h"
          using assms B.whisker_left B.whisker_right P_def by simp
        also have "... = ((e ?a B (RUNIT f B g)) B d ?c) B
                         𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h"
          using assms B.assoc'_naturality [of "e ?a" "RUNIT f B g" "d ?c"] by simp
        also have "... = (((e ?a B (RUNIT f B g)) B d ?c) B P h) B
                         (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h)"
          using assms B.whisker_left B.whisker_right P_def by simp
        finally have "(𝖺B-1[e ?a, f B g, d ?c] B P h) B
                      ((e ?a B (RUNIT f B g) B d ?c) B P h)
                        = (((e ?a B (RUNIT f B g)) B d ?c) B P h) B
                          (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h)"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B[f, g, TRG h B h] B d ?d) B
                       (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (𝖺B[e ?a, f B g, d ?c B P h] B
                       ((e ?a B (RUNIT f B g)) B d ?c B P h)) B
                       𝖺B[e ?a B (f B d ?b B e ?b) B g, d ?c, P h] B
                       (𝖺B-1[e ?a, (f B d ?b B e ?b) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B d ?b B e ?b, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, d ?b B e ?b, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
        using assms B.assoc_naturality [of "e ?a B (RUNIT f B g)" "d ?c" "P h"] B.comp_assoc
        by simp
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B[f, g, TRG h B h] B d ?d) B
                       (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d]) B
                       (e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                       ((e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (e ?a B (RUNIT f B g) B d ?c B P h)) B
                       𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                       𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                       (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
        using assms B.assoc_naturality [of "e ?a" "RUNIT f B g" "d ?c B P h"] B.comp_assoc
        by simp
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B[f, g, TRG h B h] B d ?d) B
                       (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d]) B
                       ((e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B (RUNIT f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d])) B
                       𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                       𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                       (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "((e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
              (e ?a B (RUNIT f B g) B d ?c B P h))
                = e ?a B (RUNIT f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d] B (d ?c B P h)"
          using assms B.comp_cod_arr B.whisker_left P_def
                B.interchange [of "f B g" "RUNIT f B g"]
          by simp
        also have "... = e ?a B (RUNIT f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]"
          using assms B.comp_arr_dom P_def by simp
        finally have "((e ?a B (f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                      (e ?a B (RUNIT f B g) B d ?c B P h))
                        = e ?a B (RUNIT f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       ((e ?a B 𝖺B[f, g, TRG h B h] B d ?d) B
                       (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d])) B
                       (e ?a B (RUNIT f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B ((f B SRC f) B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                       𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                       (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "(e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d]) B
              (e ?a B (RUNIT f B g) B
                 𝖺B-1[d ?c, e ?c, h B d ?d]) =
                e ?a B (RUNIT f B g) B
                  𝖺B-1[TRG h, h, d ?d] B
                  𝖺B-1[d ?c, e ?c, h B d ?d]"
          using assms B.comp_arr_dom B.comp_cod_arr B.whisker_left
                B.interchange
                  [of "f B g" "RUNIT f B g" "𝖺B-1[TRG h, h, d ?d]" "𝖺B-1[d ?c, e ?c, h B d ?d]"]
          by simp (* 3 sec *)
        also have "... = (e ?a B (RUNIT f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                         (e ?a B ((f B SRC f) B g) B 𝖺B-1[d ?c, e ?c, h B d ?d])"
          using assms B.comp_arr_dom B.whisker_left
                B.interchange [of "RUNIT f B g" "(f B SRC f) B g" "𝖺B-1[TRG h, h, d ?d]"
                                  "𝖺B-1[d ?c, e ?c, h B d ?d]"]
          by simp
        finally have "(e ?a B (f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                      (e ?a B (RUNIT f B g) B 𝖺B-1[d ?c, e ?c, h B d ?d])
                        = (e ?a B (RUNIT f B g) B 𝖺B-1[TRG h, h, d ?d]) B
                          (e ?a B ((f B SRC f) B g) B
                             𝖺B-1[d ?c, e ?c, h B d ?d])"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B-1[f, g B TRG h B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[g, TRG h B h, d ?d]) B
                       ((e ?a B 𝖺B[f, g, (TRG h B h) B d ?d]) B
                       (e ?a B (RUNIT f B g) B 𝖺B-1[TRG h, h, d ?d])) B
                       (e ?a B ((f B SRC f) B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                       𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                       (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "(e ?a B 𝖺B[f, g, TRG h B h] B d ?d) B
              (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d])
                = e ?a B (𝖺B[f, g, TRG h B h] B d ?d) B
                          𝖺B-1[f B g, TRG h B h, d ?d]"
          using assms B.whisker_left by simp
        also have "... = e ?a B
                           𝖺B-1[f, g B (d ?c B e ?c) B h, d ?d] B
                           (f B 𝖺B-1[g, (d ?c B e ?c) B h, d ?d]) B
                           𝖺B[f, g, ((d ?c B e ?c) B h) B d ?d]"
        proof -
          have "(𝖺B-1[f, g, (d ?c B e ?c) B h] B d ?d) B
                𝖺B-1[f, g B (d ?c B e ?c) B h, d ?d] B
                (f B 𝖺B-1[g, (d ?c B e ?c) B h, d ?d])
                  = 𝖺B-1[f B g, (d ?c B e ?c) B h, d ?d] B
                    𝖺B-1[f, g, ((d ?c B e ?c) B h) B d ?d]"
            using assms B.pentagon' B.comp_assoc by simp
          moreover have "B.inv (𝖺B-1[f, g, (d ?c B e ?c) B h] B d ?d)
                           = 𝖺B[f, g, (d ?c B e ?c) B h] B d ?d"
            using assms by simp
          ultimately show ?thesis
            using assms B.comp_assoc
                  B.invert_opposite_sides_of_square
                    [of "𝖺B-1[f, g, (d ?c B e ?c) B h] B d ?d"
                        "𝖺B-1[f, g B (d ?c B e ?c) B h, d ?d] B
                           (f B 𝖺B-1[g, (d ?c B e ?c) B h, d ?d])"
                        "𝖺B-1[f B g, (d ?c B e ?c) B h, d ?d]"
                        "𝖺B-1[f, g, ((d ?c B e ?c) B h) B d ?d]"]
            by simp (* 3 sec *)
        qed
        also have "... = (e ?a B 𝖺B-1[f, g B TRG h B h, d ?d]) B
                         (e ?a B f B 𝖺B-1[g, TRG h B h, d ?d]) B
                         (e ?a B 𝖺B[f, g, (TRG h B h) B d ?d])"
          using assms B.whisker_left by simp
        finally have "(e ?a B 𝖺B[f, g, TRG h B h] B d ?d) B
                      (e ?a B 𝖺B-1[f B g, TRG h B h, d ?d])
                        = (e ?a B 𝖺B-1[f, g B TRG h B h, d ?d]) B
                          (e ?a B f B 𝖺B-1[g, TRG h B h, d ?d]) B
                          (e ?a B 𝖺B[f, g, (TRG h B h) B d ?d])"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B-1[f, g B TRG h B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[g, TRG h B h, d ?d]) B
                       (e ?a B RUNIT f B g B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B 𝖺B[f B SRC f, g, (d ?c B e ?c) B h B d ?d]) B
                       (e ?a B ((f B SRC f) B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                       𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                       (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "(e ?a B 𝖺B[f, g, (TRG h B h) B d ?d]) B
              (e ?a B (RUNIT f B g) B 𝖺B-1[TRG h, h, d ?d])
                = e ?a B 𝖺B[f, g, (TRG h B h) B d ?d] B ((RUNIT f B g) B 𝖺B-1[TRG h, h, d ?d])"
          using assms B.whisker_left by simp
        also have "... = e ?a B
                           (RUNIT f B g B 𝖺B-1[TRG h, h, d ?d]) B
                           𝖺B[f B SRC f, g, TRG h B h B d ?d]"
          using assms B.hseqI' B.assoc_naturality [of "RUNIT f" g "𝖺B-1[TRG h, h, d ?d]"]
          by simp
        also have "... = (e ?a B RUNIT f B g B 𝖺B-1[TRG h, h, d ?d]) B
                         (e ?a B 𝖺B[f B SRC f, g, TRG h B h B d ?d])"
          using assms B.whisker_left by simp
        finally have "(e ?a B 𝖺B[f, g, (TRG h B h) B d ?d]) B
                      (e ?a B (RUNIT f B g) B 𝖺B-1[TRG h, h, d ?d])
                        = (e ?a B RUNIT f B g B 𝖺B-1[TRG h, h, d ?d]) B
                          (e ?a B 𝖺B[f B SRC f, g, TRG h B h B d ?d])"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B-1[f, g B TRG h B h, d ?d]) B
                       ((e ?a B f B 𝖺B-1[g, TRG h B h, d ?d]) B
                       (e ?a B f B LUNIT (g B (TRG h B h) B d ?d))) B
                       (e ?a B 𝖺B[f, d ?b B e ?b, g B (TRG h B h) B d ?d]) B
                       (e ?a B (f B SRC f) B g B 𝖺B-1[d ?c B e ?c, h, d ?d]) B
                       (e ?a B 𝖺B[f B SRC f, g, TRG h B h B d ?d]) B
                       (e ?a B ((f B SRC f) B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                       𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                       (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, TRG g, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "e ?a B RUNIT f B g B 𝖺B-1[d ?c B e ?c, h, d ?d]
                = (e ?a B RUNIT f B g B (TRG h B h) B d ?d) B
                  (e ?a B (f B SRC f) B g B 𝖺B-1[TRG h, h, d ?d])"
          using assms B.whisker_left B.comp_arr_dom B.comp_cod_arr
                B.interchange [of "RUNIT f" "f B SRC f" "g B ((TRG h B h) B d ?d)"
                                  "g B 𝖺B-1[TRG h, h, d ?d]"]
          by simp (* 5 sec *)
        also have "... = (e ?a B
                           (f B LUNIT (g B (TRG h B h) B d ?d)) B
                           𝖺B[f, d ?b B e ?b, g B (TRG h B h) B d ?d]) B
                         (e ?a B (f B SRC f) B g B 𝖺B-1[TRG h, h, d ?d])"
          using assms TRIANGLE [of f "g B (TRG h B h) B d ?d"] by simp
        also have "... = (e ?a B f B LUNIT (g B (TRG h B h) B d ?d)) B
                         (e ?a B 𝖺B[f, d ?b B e ?b, g B (TRG h B h) B d ?d]) B
                         (e ?a B (f B SRC f) B g B 𝖺B-1[TRG h, h, d ?d])"
          using assms B.whisker_left B.comp_assoc by simp
        finally have "e ?a B RUNIT f B g B 𝖺B-1[TRG h, h, d ?d]
                        = (e ?a B f B LUNIT (g B (TRG h B h) B d ?d)) B
                          (e ?a B 𝖺B[f, d ?b B e ?b, g B (TRG h B h) B d ?d]) B
                          (e ?a B (f B SRC f) B g B 𝖺B-1[TRG h, h, d ?d])"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B-1[f, g B TRG h B h, d ?d]) B
                       ((e ?a B f B 𝖺B-1[g, TRG h B h, d ?d]) B
                       (e ?a B f B LUNIT g B (TRG h B h) B d ?d)) B
                       (e ?a B f B 𝖺B-1[TRG g, g, (TRG h B h) B d ?d]) B
                       (e ?a B 𝖺B[f, d ?b B e ?b, g B (TRG h B h) B d ?d]) B
                       (e ?a B (f B SRC f) B g B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B 𝖺B[f B SRC f, g, TRG h B h B d ?d]) B
                       (e ?a B ((f B SRC f) B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                       𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                       (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "e ?a B f B LUNIT (g B (TRG h B h) B d ?d)
                = e ?a B f B
                    (LUNIT g B (TRG h B h) B d ?d) B
                    𝖺B-1[d ?b B e ?b, g, (TRG h B h) B d ?d]"
          using assms LUNIT_hcomp [of g "(TRG h B h) B d ?d"]
                B.invert_side_of_triangle
          by simp
        also have "... = (e ?a B f B LUNIT g B (TRG h B h) B d ?d) B
                         (e ?a B f B 𝖺B-1[TRG g, g, (TRG h B h) B d ?d])"
          using assms B.whisker_left by simp
        finally have "e ?a B f B LUNIT (g B (TRG h B h) B d ?d)
                        = (e ?a B f B LUNIT g B (TRG h B h) B d ?d) B
                          (e ?a B f B 𝖺B-1[TRG g, g, (TRG h B h) B d ?d])"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B g B LUNIT h) B d ?d) B
                       ((e ?a B 𝖺B-1[f, g B TRG h B h, d ?d]) B
                       (e ?a B f B (LUNIT g B TRG h B h) B d ?d)) B
                       (e ?a B f B 𝖺B-1[TRG g B g, TRG h B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, g, (TRG h B h) B d ?d]) B
                       (e ?a B 𝖺B[f, d ?b B e ?b, g B (TRG h B h) B d ?d]) B
                       (e ?a B (f B SRC f) B g B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B 𝖺B[f B SRC f, g, TRG h B h B d ?d]) B
                       (e ?a B ((f B SRC f) B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                       𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                       (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "(e ?a B f B 𝖺B-1[g, TRG h B h, d ?d]) B
              (e ?a B f B LUNIT g B (TRG h B h) B d ?d)
                = e ?a B f B
                    𝖺B-1[g, TRG h B h, d ?d] B (LUNIT g B (TRG h B h) B d ?d)"
          using assms B.whisker_left by simp
        also have "... = e ?a B f B
                           ((LUNIT g B TRG h B h) B d ?d) B
                           𝖺B-1[TRG g B g, TRG h B h, d ?d]"
          using assms B.assoc'_naturality [of "LUNIT g" "TRG h B h" "d ?d"]
          by simp
        also have "... = (e ?a B f B (LUNIT g B TRG h B h) B d ?d) B
                         (e ?a B f B 𝖺B-1[TRG g B g, TRG h B h, d ?d])"
          using assms B.whisker_left by simp
        finally have "(e ?a B f B 𝖺B-1[g, TRG h B h, d ?d]) B
                      (e ?a B f B LUNIT g B (TRG h B h) B d ?d)
                        = (e ?a B f B (LUNIT g B TRG h B h) B d ?d) B
                          (e ?a B f B 𝖺B-1[TRG g B g, TRG h B h, d ?d])"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = ((e ?a B (f B g B LUNIT h) B d ?d) B
                       (e ?a B (f B LUNIT g B SRC g B h) B d ?d)) B
                       (e ?a B 𝖺B-1[f, (TRG g B g) B TRG h B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g B g, TRG h B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, g, (TRG h B h) B d ?d]) B
                       (e ?a B 𝖺B[f, d ?b B e ?b, g B (TRG h B h) B d ?d]) B
                       (e ?a B (f B SRC f) B g B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B 𝖺B[f B SRC f, g, TRG h B h B d ?d]) B
                       (e ?a B ((f B SRC f) B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                       𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                       (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "(e ?a B 𝖺B-1[f, g B TRG h B h, d ?d]) B
              (e ?a B f B (LUNIT g B TRG h B h) B d ?d)
                = e ?a B 𝖺B-1[f, g B TRG h B h, d ?d] B
                         (f B (LUNIT g B TRG h B h) B d ?d)"
          using assms B.whisker_left by simp
        also have "... = e ?a B
                           ((f B LUNIT g B TRG h B h) B d ?d) B
                           𝖺B-1[f, (TRG g B g) B TRG h B h, d ?d]"
          using assms B.assoc'_naturality [of f "LUNIT g B TRG h B h" "d ?d"] by simp
        also have "... = (e ?a B (f B LUNIT g B TRG h B h) B d ?d) B
                         (e ?a B 𝖺B-1[f, (TRG g B g) B TRG h B h, d ?d])"
          using assms B.whisker_left by simp
        finally have "(e ?a B 𝖺B-1[f, g B TRG h B h, d ?d]) B
                      (e ?a B f B (LUNIT g B TRG h B h) B d ?d)
                        = (e ?a B (f B LUNIT g B TRG h B h) B d ?d) B
                          (e ?a B 𝖺B-1[f, (TRG g B g) B TRG h B h, d ?d])"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B LUNIT g B LUNIT h) B d ?d) B
                       (e ?a B 𝖺B-1[f, (TRG g B g) B TRG h B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g B g, TRG h B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, g, (TRG h B h) B d ?d]) B
                       (e ?a B 𝖺B[f, d ?b B e ?b, g B (TRG h B h) B d ?d]) B
                       (e ?a B (f B SRC f) B g B 𝖺B-1[TRG h, h, d ?d]) B
                       (e ?a B 𝖺B[f B SRC f, g, TRG h B h B d ?d]) B
                       (e ?a B ((f B SRC f) B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                       𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                       (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                       ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                       ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                       (𝖺B[e ?a, f, d ?b B P g] B P h) B
                       (𝖺B[e ?a B f, d ?b, P g] B P h) B
                       ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
      proof -
        have "(e ?a B (f B g B LUNIT h) B d ?d) B
              (e ?a B (f B LUNIT g B TRG h B h) B d ?d)
                = e ?a B (f B LUNIT g B LUNIT h) B d ?d"
          using assms B.whisker_left B.whisker_right B.comp_arr_dom B.comp_cod_arr
                B.interchange [of g "LUNIT g" "LUNIT h" "(d ?c B e ?c) B h"]
          by simp (* 5 sec *)
        thus ?thesis
          using assms by simp
      qed
      finally have L: "P 𝖺B[f, g, h] B CMP (f B g) h B (CMP f g B P h)
                         = (e ?a B (f B LUNIT g B LUNIT h) B d ?d) B
                           (e ?a B 𝖺B-1[f, (TRG g B g) B TRG h B h, d ?d]) B
                           (e ?a B f B 𝖺B-1[TRG g B g, TRG h B h, d ?d]) B
                           (e ?a B f B 𝖺B-1[TRG g, g, (TRG h B h) B d ?d]) B
                           (e ?a B 𝖺B[f, d ?b B e ?b, g B (TRG h B h) B d ?d]) B
                           (e ?a B (f B SRC f) B g B 𝖺B-1[TRG h, h, d ?d]) B
                           (e ?a B 𝖺B[f B SRC f, g, TRG h B h B d ?d]) B
                           (e ?a B ((f B SRC f) B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                           𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                           𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                           (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                           ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                           ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                           ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                           (𝖺B[e ?a, f, d ?b B P g] B P h) B
                           (𝖺B[e ?a B f, d ?b, P g] B P h) B
                           ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"
        by simp

      have "CMP f (g B h) B (P f B CMP g h) B 𝖺B[P f, P g, P h]
              = (e ?a B 𝖺B-1[f, g B h, d ?d]) B
                (e ?a B f B LUNIT ((g B h) B d ?d)) B
                (e ?a B f B 𝖺B-1[d ?b, e ?b, (g B h) B d ?d]) B
                𝖺B[e ?a, f, d ?b B P (g B h)] B
                𝖺B[e ?a B f, d ?b, P (g B h)] B
                (𝖺B-1[e ?a, f, d ?b] B P (g B h)) B
                (P f B
                  (e ?b B 𝖺B-1[g, h, d ?d]) B
                  (e ?b B g B LUNIT (h B d ?d)) B
                  (e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                  𝖺B[e ?b, g, d ?c B P h] B
                  𝖺B[e ?b B g, d ?c, P h] B
                  (𝖺B-1[e ?b, g, d ?c] B P h)) B
                𝖺B[P f, P g, P h]"
        unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp
      also have "... = (e ?a B 𝖺B-1[f, g B h, d ?d]) B
                       (e ?a B f B LUNIT ((g B h) B d ?d)) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, (g B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P (g B h)] B
                       𝖺B[e ?a B f, d ?b, P (g B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B P (g B h)) B
                       (P f B e ?b B 𝖺B-1[g, h, d ?d]) B
                       (P f B e ?b B g B LUNIT (h B d ?d)) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
        using assms B.whisker_left P_def B.comp_assoc by auto (* 5 sec *)
      also have "... = ((e ?a B 𝖺B-1[f, g B h, d ?d]) B
                       (e ?a B f B LUNIT (g B h) B d ?d)) B
                       (e ?a B f B 𝖺B-1[d ?b B e ?b, g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, (g B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P (g B h)] B
                       𝖺B[e ?a B f, d ?b, P (g B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B P (g B h)) B
                       (P f B e ?b B 𝖺B-1[g, h, d ?d]) B
                       (P f B e ?b B g B LUNIT (h B d ?d)) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have "e ?a B f B LUNIT ((g B h) B d ?d)
                = e ?a B f B (LUNIT (g B h) B d ?d) B 𝖺B-1[TRG g, g B h, d ?d]"
          using assms LUNIT_hcomp [of "g B h" "d ?d"] B.invert_side_of_triangle
          by simp
        also have "... = (e ?a B f B LUNIT (g B h) B d ?d) B
                         (e ?a B f B 𝖺B-1[TRG g, g B h, d ?d])"
          using assms B.whisker_left by simp
        finally have "e ?a B f B LUNIT ((g B h) B d ?d)
                        = (e ?a B f B LUNIT (g B h) B d ?d) B
                          (e ?a B f B 𝖺B-1[TRG g, g B h, d ?d])"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                       (e ?a B 𝖺B-1[f, TRG g B g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, (g B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P (g B h)] B
                       𝖺B[e ?a B f, d ?b, P (g B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B P (g B h)) B
                       (P f B e ?b B 𝖺B-1[g, h, d ?d]) B
                       (P f B e ?b B g B LUNIT (h B d ?d)) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have "(e ?a B 𝖺B-1[f, g B h, d ?d]) B (e ?a B f B LUNIT (g B h) B d ?d)
                = e ?a B 𝖺B-1[f, g B h, d ?d] B (f B LUNIT (g B h) B d ?d)"
          using assms B.whisker_left by simp
        also have "... = e ?a B ((f B LUNIT (g B h)) B d ?d) B
                                 𝖺B-1[f, TRG g B g B h, d ?d]"
          using assms B.assoc'_naturality [of f "LUNIT (g B h)" "d ?d"]
                LUNIT_in_hom [of "g B h"]
          by auto
        also have "... = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                         (e ?a B 𝖺B-1[f, TRG g B g B h, d ?d])"
          using assms B.whisker_left by simp
        finally have "(e ?a B 𝖺B-1[f, g B h, d ?d]) B
                      (e ?a B f B LUNIT (g B h) B d ?d)
                        = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                          (e ?a B 𝖺B-1[f, TRG g B g B h, d ?d])"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                       (e ?a B 𝖺B-1[f, TRG g B g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, (g B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P (g B h)] B
                       𝖺B[e ?a B f, d ?b, P (g B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B P (g B h)) B
                       ((P f B e ?b B 𝖺B-1[g, h, d ?d]) B
                       (P f B e ?b B RUNIT g B h B d ?d)) B
                       (P f B e ?b B 𝖺B-1[g, TRG h, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have "P f B e ?b B g B LUNIT (h B d ?d)
                = P f B e ?b B (RUNIT g B h B d ?d) B 𝖺B-1[g, TRG h, h B d ?d]"
          using assms TRIANGLE [of g "h B d ?d"] B.invert_side_of_triangle by simp
        also have "... = (P f B e ?b B RUNIT g B h B d ?d) B
                         (P f B e ?b B 𝖺B-1[g, TRG h, h B d ?d])"
          using assms B.whisker_left P_def by simp
        finally have "P f B e ?b B g B LUNIT (h B d ?d) =
                        (P f B e ?b B RUNIT g B h B d ?d) B
                        (P f B e ?b B 𝖺B-1[g, TRG h, h B d ?d])"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                       (e ?a B 𝖺B-1[f, TRG g B g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, (g B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P (g B h)] B
                       𝖺B[e ?a B f, d ?b, P (g B h)] B
                       ((𝖺B-1[e ?a, f, d ?b] B P (g B h)) B
                       (P f B e ?b B (RUNIT g B h) B d ?d)) B
                       (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                       (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have "(P f B e ?b B 𝖺B-1[g, h, d ?d]) B (P f B e ?b B RUNIT g B h B d ?d)
                = P f B e ?b B 𝖺B-1[g, h, d ?d] B (RUNIT g B h B d ?d)"
          using assms B.whisker_left P_def by simp
        also have "... = P f B e ?b B
                           ((RUNIT g B h) B d ?d) B 𝖺B-1[g B SRC g, h, d ?d]"
          using assms B.assoc'_naturality [of "RUNIT g" h "d ?d"] by auto
        also have "... = (P f B e ?b B (RUNIT g B h) B d ?d) B
                         (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d])"
          using assms B.whisker_left P_def by simp
        finally have "(P f B e ?b B 𝖺B-1[g, h, d ?d]) B
                      (P f B e ?b B RUNIT g B h B d ?d)
                        = (P f B e ?b B (RUNIT g B h) B d ?d) B
                          (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d])"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                       (e ?a B 𝖺B-1[f, TRG g B g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, (g B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P (g B h)] B
                       (𝖺B[e ?a B f, d ?b, P (g B h)] B
                       (((e ?a B f) B d ?b) B e ?b B (RUNIT g B h) B d ?d)) B
                       (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d) B
                       (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                       (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have "(𝖺B-1[e ?a, f, d ?b] B P (g B h)) B (P f B e ?b B (RUNIT g B h) B d ?d)
                = 𝖺B-1[e ?a, f, d ?b] B e ?b B (RUNIT g B h) B d ?d"
          using assms B.comp_arr_dom B.comp_cod_arr P_def
                B.interchange
                  [of "𝖺B-1[e ?a, f, d ?b]" "P f" "P (g B h)" "e ?b B (RUNIT g B h) B d ?d"]
          by simp
        also have "... = (((e ?a B f) B d ?b) B e ?b B (RUNIT g B h) B d ?d) B
                         (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d)"
          using assms B.comp_arr_dom B.comp_cod_arr
                B.interchange
                  [of "(e ?a B f) B d ?b" "𝖺B-1[e ?a, f, d ?b]"
                      "e ?b B (RUNIT g B h) B d ?d"
                      "e ?b B ((g B SRC g) B h) B d ?d"]
          by simp (* 4 sec *)
        finally have "(𝖺B-1[e ?a, f, d ?b] B P (g B h)) B
                      (P f B e ?b B (RUNIT g B h) B d ?d)
                        = (((e ?a B f) B d ?b) B e ?b B (RUNIT g B h) B d ?d) B
                          (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d)"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                       (e ?a B 𝖺B-1[f, TRG g B g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b B e ?b, g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, (g B h) B d ?d]) B
                       (𝖺B[e ?a, f, d ?b B P (g B h)] B
                       ((e ?a B f) B d ?b B e ?b B (RUNIT g B h) B d ?d)) B
                       𝖺B[e ?a B f, d ?b, P ((g B SRC g) B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d) B
                       (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                       (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
        using assms P_def B.comp_assoc
              B.assoc_naturality [of "e ?a B f" "d ?b" "e ?b B (RUNIT g B h) B d ?d"]
        by simp
      also have "... = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                       (e ?a B 𝖺B-1[f, (d ?b B e ?b) B g B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b B e ?b, g B h, d ?d]) B
                       ((e ?a B f B 𝖺B-1[d ?b, e ?b, (g B h) B d ?d]) B
                       (e ?a B f B d ?b B e ?b B (RUNIT g B h) B d ?d)) B
                       𝖺B[e ?a, f, d ?b B P ((g B d ?c B e ?c) B h)] B
                       𝖺B[e ?a B f, d ?b, P ((g B d ?c B e ?c) B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B d ?c B e ?c) B h) B d ?d) B
                       (P f B e ?b B 𝖺B-1[g B d ?c B e ?c, h, d ?d]) B
                       (P f B e ?b B 𝖺B-1[g, d ?c B e ?c, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
        using assms P_def B.comp_assoc
              B.assoc_naturality [of "e ?a" f "d ?b B e ?b B (RUNIT g B h) B d ?d"]
        by simp
      also have "... = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                       (e ?a B 𝖺B-1[f, TRG g B g B h, d ?d]) B
                       ((e ?a B f B 𝖺B-1[TRG g, g B h, d ?d]) B
                       (e ?a B f B TRG g B (RUNIT g B h) B d ?d)) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P ((g B d ?c B e ?c) B h)] B
                       𝖺B[e ?a B f, d ?b, P ((g B d ?c B e ?c) B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d) B
                       (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                       (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have
          "(e ?a B f B 𝖺B-1[d ?b, e ?b, (g B h) B d ?d]) B
           (e ?a B f B d ?b B e ?b B (RUNIT g B h) B d ?d)
             = e ?a B f B
                 𝖺B-1[d ?b, e ?b, (g B h) B d ?d] B
                 (d ?b B e ?b B (RUNIT g B h) B d ?d)"
          using assms B.whisker_left by simp
        also have "... = e ?a B f B
                           (TRG g B (RUNIT g B h) B d ?d) B
                           𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d]"
          using assms B.assoc'_naturality [of "d ?b" "e ?b" "(RUNIT g B h) B d ?d"] by auto
        also have "... = (e ?a B f B SRC f B (RUNIT g B h) B d ?d) B
                         (e ?a B f B 𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d])"
          using assms B.whisker_left by simp
        finally have
          "(e ?a B f B 𝖺B-1[d ?b, e ?b, (g B h) B d ?d]) B
           (e ?a B f B d ?b B e ?b B (RUNIT g B h) B d ?d)
             = (e ?a B f B SRC f B (RUNIT g B h) B d ?d) B
               (e ?a B f B 𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d])"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                       ((e ?a B 𝖺B-1[f, TRG g B g B h, d ?d]) B
                       (e ?a B f B (TRG g B (RUNIT g B h)) B d ?d)) B
                       (e ?a B f B 𝖺B-1[TRG g, (g B SRC g) B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P ((g B SRC g) B h)] B
                       𝖺B[e ?a B f, d ?b, P ((g B SRC g) B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d) B
                       (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                       (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have "(e ?a B f B 𝖺B-1[TRG g, g B h, d ?d]) B
              (e ?a B f B TRG g B (RUNIT g B h) B d ?d)
                = e ?a B f B
                    𝖺B-1[TRG g, g B h, d ?d] B (TRG g B (RUNIT g B h) B d ?d)"
          using assms B.whisker_left by simp
        also have "... = e ?a B f B
                           ((TRG g B (RUNIT g B h)) B d ?d) B
                           𝖺B-1[TRG g, (g B SRC g) B h, d ?d]"
          using assms B.assoc'_naturality [of "TRG g" "RUNIT g B h" "d ?d"] by simp
        also have "... = (e ?a B f B (TRG g B (RUNIT g B h)) B d ?d) B
                         (e ?a B f B 𝖺B-1[TRG g, (g B SRC g) B h, d ?d])"
          using assms B.whisker_left by simp
        finally have "(e ?a B f B 𝖺B-1[TRG g, g B h, d ?d]) B
                      (e ?a B f B TRG g B (RUNIT g B h) B d ?d)
                        = (e ?a B f B (TRG g B (RUNIT g B h)) B d ?d) B
                          (e ?a B f B 𝖺B-1[TRG g, (g B SRC g) B h, d ?d])"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                       (e ?a B (f B TRG g B (RUNIT g B h)) B d ?d) B
                       (e ?a B 𝖺B-1[f, TRG g B (g B SRC g) B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, (g B SRC g) B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P ((g B SRC g) B h)] B
                       𝖺B[e ?a B f, d ?b, P ((g B SRC g) B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d) B
                       (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                       (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have "(e ?a B 𝖺B-1[f, TRG g B g B h, d ?d]) B
              (e ?a B f B (TRG g B (RUNIT g B h)) B d ?d)
                = e ?a B 𝖺B-1[f, TRG g B g B h, d ?d] B
                          (f B (TRG g B (RUNIT g B h)) B d ?d)"
          using assms B.whisker_left by simp
        also have "... = e ?a B
                           ((f B SRC f B (RUNIT g B h)) B d ?d) B
                           𝖺B-1[f, TRG g B (g B SRC g) B h, d ?d]"
          using assms
                B.assoc'_naturality [of f "(d ?b B e ?b) B (RUNIT g B h)" "d ?d"]
          by simp
        also have "... = (e ?a B (f B SRC f B (RUNIT g B h)) B d ?d) B
                         (e ?a B 𝖺B-1[f, TRG g B (g B SRC g) B h, d ?d])"
          using assms B.whisker_left by simp
        finally have "(e ?a B 𝖺B-1[f, TRG g B g B h, d ?d]) B
                      (e ?a B f B (TRG g B (RUNIT g B h)) B d ?d)
                        = (e ?a B (f B TRG g B (RUNIT g B h)) B d ?d) B
                          (e ?a B 𝖺B-1[f, TRG g B (g B SRC g) B h, d ?d])"
          using assms by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B LUNIT (g B h)) B d ?d) B
                       (e ?a B (f B SRC f B g B LUNIT h) B d ?d) B
                       (e ?a B (f B SRC f B 𝖺B[g, d ?c B e ?c, h]) B d ?d) B
                       (e ?a B 𝖺B-1[f, TRG g B (g B SRC g) B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, (g B SRC g) B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P ((g B SRC g) B h)] B
                       𝖺B[e ?a B f, d ?b, P ((g B SRC g) B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d) B
                       (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                       (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have "e ?a B (f B SRC f B (RUNIT g B h)) B d ?d
                = e ?a B (f B SRC f B(g B LUNIT h) B 𝖺B[g, SRC g, h]) B d ?d"
          using assms TRIANGLE [of g h] by simp
        also have "... = (e ?a B (f B SRC f B g B LUNIT h) B d ?d) B
                         (e ?a B (f B SRC f B 𝖺B[g, TRG h, h]) B d ?d)"
          using assms B.whisker_left B.whisker_right by simp
        finally have "e ?a B (f B SRC f B (RUNIT g B h)) B d ?d
                        = (e ?a B (f B SRC f B g B LUNIT h) B d ?d) B
                          (e ?a B (f B SRC f B 𝖺B[g, TRG h, h]) B d ?d)"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B LUNIT g B h) B d ?d) B
                       ((e ?a B (f B 𝖺B-1[TRG g, g, h]) B d ?d) B
                       (e ?a B (f B TRG g B g B LUNIT h) B d ?d)) B
                       (e ?a B (f B TRG g B 𝖺B[g, SRC g, h]) B d ?d) B
                       (e ?a B 𝖺B-1[f, TRG g B (g B SRC g) B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, (g B SRC g) B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P ((g B SRC g) B h)] B
                       𝖺B[e ?a B f, d ?b, P ((g B SRC g) B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d) B
                       (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                       (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have "e ?a B (f B LUNIT (g B h)) B d ?d
                = e ?a B (f B (LUNIT g B h) B 𝖺B-1[TRG g, g, h]) B d ?d"
          using assms LUNIT_hcomp [of g h] B.invert_side_of_triangle by simp
        also have "... = (e ?a B (f B LUNIT g B h) B d ?d) B
                         (e ?a B (f B 𝖺B-1[TRG g, g, h]) B d ?d)"
          using assms B.whisker_left B.whisker_right by simp
        finally have "e ?a B (f B LUNIT (g B h)) B d ?d
                        = (e ?a B (f B LUNIT g B h) B d ?d) B
                          (e ?a B (f B 𝖺B-1[TRG g, g, h]) B d ?d)"
          by simp
        thus ?thesis
          using assms B.comp_assoc by simp
      qed
      also have "... = ((e ?a B (f B LUNIT g B h) B d ?d) B
                       (e ?a B (f B (TRG g B g) B LUNIT h) B d ?d)) B
                       (e ?a B (f B 𝖺B-1[TRG g, g, TRG h B h]) B d ?d) B
                       (e ?a B (f B TRG g B 𝖺B[g, SRC g, h]) B d ?d) B
                       (e ?a B 𝖺B-1[f, TRG g B (g B SRC g) B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, (g B SRC g) B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P ((g B SRC g) B h)] B
                       𝖺B[e ?a B f, d ?b, P ((g B SRC g) B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d) B
                       (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                       (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have
          "(e ?a B (f B 𝖺B-1[TRG g, g, h]) B d ?d) B
           (e ?a B (f B TRG g B g B LUNIT h) B d ?d)
             = e ?a B (f B 𝖺B-1[TRG g, g, h] B (TRG g B g B LUNIT h)) B d ?d"
          using assms B.whisker_left B.whisker_right by auto
        also have "... = e ?a B
                           (f B ((TRG g B g) B LUNIT h) B
                           𝖺B-1[TRG g, g, SRC g B h]) B d ?d"
          using assms B.assoc'_naturality [of "TRG g" g "LUNIT h"] by auto
        also have "... = (e ?a B (f B (TRG g B g) B LUNIT h) B d ?d) B
                         (e ?a B (f B 𝖺B-1[TRG g, g, TRG h B h]) B d ?d)"
          using assms B.whisker_left B.whisker_right by auto
        finally have "(e ?a B (f B 𝖺B-1[TRG g, g, h]) B d ?d) B
                      (e ?a B (f B TRG g B g B LUNIT h) B d ?d)
                        = (e ?a B (f B (TRG g B g) B LUNIT h) B d ?d) B
                          (e ?a B (f B 𝖺B-1[TRG g, g, TRG h B h]) B d ?d)"
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (e ?a B (f B LUNIT g B LUNIT h) B d ?d) B
                       (e ?a B (f B 𝖺B-1[TRG g, g, TRG h B h]) B d ?d) B
                       (e ?a B (f B TRG g B 𝖺B[g, d ?c B e ?c, h]) B d ?d) B
                       (e ?a B 𝖺B-1[f, TRG g B (g B d ?c B e ?c) B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[TRG g, (g B SRC g) B h, d ?d]) B
                       (e ?a B f B 𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d]) B
                       𝖺B[e ?a, f, d ?b B P ((g B SRC g) B h)] B
                       𝖺B[e ?a B f, d ?b, P ((g B SRC g) B h)] B
                       (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B d ?c B e ?c) B h) B d ?d) B
                       (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                       (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                       (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                       (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                       (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                       (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                       𝖺B[P f, P g, P h]"
      proof -
        have "(e ?a B (f B LUNIT g B h) B d ?d) B
              (e ?a B (f B (TRG g B g) B LUNIT h) B d ?d)
                = e ?a B (f B LUNIT g B LUNIT h) B d ?d"
          using assms B.whisker_left B.whisker_right B.comp_arr_dom B.comp_cod_arr
                B.interchange [of "LUNIT g" "TRG g B g" h "LUNIT h"]
          by simp (* 4 sec *)
        thus ?thesis
          using assms by simp
      qed
      finally have R: "CMP f (g B h) B (P f B CMP g h) B 𝖺B[P f, P g, P h]
                         = (e ?a B (f B LUNIT g B LUNIT h) B d ?d) B
                           (e ?a B (f B 𝖺B-1[TRG g, g, TRG h B h]) B d ?d) B
                           (e ?a B (f B SRC f B 𝖺B[g, TRG h, h]) B d ?d) B
                           (e ?a B 𝖺B-1[f, TRG g B (g B SRC g) B h, d ?d]) B
                           (e ?a B f B 𝖺B-1[TRG g, (g B SRC g) B h, d ?d]) B
                           (e ?a B f B 𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d]) B
                           𝖺B[e ?a, f, d ?b B P ((g B SRC g) B h)] B
                           𝖺B[e ?a B f, d ?b, P ((g B SRC g) B h)] B
                           (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d) B
                           (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                           (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                           (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                           (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                           (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                           (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                           𝖺B[P f, P g, P h]"
         using assms by simp

      text ‹
        The portions of the expressions on the right-hand sides of assertions L› and R›
        that are not identical only involve canonical isomorphisms, and thus they can be proved
        equal automatically by the simplifier, once we have expressed them in the formal
        language of B›.
      ›

      let ?LHS = "(e ?a B 𝖺B-1[f, (TRG g B g) B TRG h B h, d ?d]) B
                  (e ?a B f B 𝖺B-1[TRG g B g, TRG h B h, d ?d]) B
                  (e ?a B f B 𝖺B-1[TRG g, g, (TRG h B h) B d ?d]) B
                  (e ?a B 𝖺B[f, d ?b B e ?b, g B (TRG h B h) B d ?d]) B
                  (e ?a B (f B SRC f) B g B 𝖺B-1[TRG h, h, d ?d]) B
                  (e ?a B 𝖺B[f B SRC f, g, TRG h B h B d ?d]) B
                  (e ?a B ((f B SRC f) B g) B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                  𝖺B[e ?a, (f B SRC f) B g, d ?c B P h] B
                  𝖺B[e ?a B (f B SRC f) B g, d ?c, P h] B
                  (𝖺B-1[e ?a, (f B SRC f) B g, d ?c] B P h) B
                  ((e ?a B 𝖺B-1[f B SRC f, g, d ?c]) B P h) B
                  ((e ?a B 𝖺B-1[f, SRC f, g B d ?c]) B P h) B
                  ((e ?a B f B 𝖺B-1[d ?b, e ?b, g B d ?c]) B P h) B
                  (𝖺B[e ?a, f, d ?b B P g] B P h) B
                  (𝖺B[e ?a B f, d ?b, P g] B P h) B
                  ((𝖺B-1[e ?a, f, d ?b] B P g) B P h)"

      let ?LHSt = "(e ?a  𝖺-1[f, (TRG g  g)  TRG h  h, d ?d]) 
                   (e ?a  f  𝖺-1[TRG g  g, TRG h  h, d ?d]) 
                   (e ?a  f  𝖺-1[TRG g, g, (TRG h  h)  d ?d]) 
                   (e ?a  𝖺[f, d ?b  e ?b, g  (TRG h  h)  d ?d]) 
                   (e ?a  (f  SRC f)  g  𝖺-1[TRG h, h, d ?d]) 
                   (e ?a  𝖺[f  SRC f, g, TRG h  h  d ?d]) 
                   (e ?a  ((f  SRC f)  g)  𝖺-1[d ?c, e ?c, h  d ?d]) 
                   𝖺[e ?a, (f  SRC f)  g, d ?c  PRJ h] 
                   𝖺[e ?a  (f  SRC f)  g, d ?c, PRJ h] 
                   (𝖺-1[e ?a, (f  SRC f)  g, d ?c]  PRJ h) 
                   ((e ?a  𝖺-1[f  SRC f, g, d ?c])  PRJ h) 
                   ((e ?a  𝖺-1[f, SRC f, g  d ?c])  PRJ h) 
                   ((e ?a  f  𝖺-1[d ?b, e ?b, g  d ?c])  PRJ h) 
                   (𝖺[e ?a, f, d ?b  PRJ g]  PRJ h) 
                   (𝖺[e ?a  f, d ?b, PRJ g]  PRJ h) 
                   ((𝖺-1[e ?a, f, d ?b]  PRJ g)  PRJ h)"

      let ?RHS = "(e ?a B (f B 𝖺B-1[TRG g, g, TRG h B h]) B d ?d) B
                  (e ?a B (f B SRC f B 𝖺B[g, TRG h, h]) B d ?d) B
                  (e ?a B 𝖺B-1[f, TRG g B (g B SRC g) B h, d ?d]) B
                  (e ?a B f B 𝖺B-1[TRG g, (g B SRC g) B h, d ?d]) B
                  (e ?a B f B 𝖺B-1[d ?b, e ?b, ((g B SRC g) B h) B d ?d]) B
                  𝖺B[e ?a, f, d ?b B P ((g B SRC g) B h)] B
                  𝖺B[e ?a B f, d ?b, P ((g B SRC g) B h)] B
                  (𝖺B-1[e ?a, f, d ?b] B e ?b B ((g B SRC g) B h) B d ?d) B
                  (P f B e ?b B 𝖺B-1[g B SRC g, h, d ?d]) B
                  (P f B e ?b B 𝖺B-1[g, SRC g, h B d ?d]) B
                  (P f B e ?b B g B 𝖺B-1[d ?c, e ?c, h B d ?d]) B
                  (P f B 𝖺B[e ?b, g, d ?c B P h]) B
                  (P f B 𝖺B[e ?b B g, d ?c, P h]) B
                  (P f B 𝖺B-1[e ?b, g, d ?c] B P h) B
                  𝖺B[P f, P g, P h]"

      let ?RHSt = "(e ?a  (f  𝖺-1[TRG g, g, TRG h  h])  d ?d) 
                   (e ?a  (f  SRC f  𝖺[g, TRG h, h])  d ?d) 
                   (e ?a  𝖺-1[f, TRG g  (g  SRC g)  h, d ?d]) 
                   (e ?a  f  𝖺-1[TRG g, (g  SRC g)  h, d ?d]) 
                   (e ?a  f  𝖺-1[d ?b, e ?b, ((g  SRC g)  h)  d ?d]) 
                   𝖺[e ?a, f, d ?b  (e ?b  ((g  SRC g)  h)  d ?d)] 
                   𝖺[e ?a  f, d ?b, (e ?b  ((g  SRC g)  h)  d ?d)] 
                   (𝖺-1[e ?a, f, d ?b]  e ?b  ((g  SRC g)  h)  d ?d) 
                   (PRJ f  e ?b  𝖺-1[g  SRC g, h, d ?d]) 
                   (PRJ f  e ?b  𝖺-1[g, SRC g, h  d ?d]) 
                   (PRJ f  e ?b  g  𝖺-1[d ?c, e ?c, h  d ?d]) 
                   (PRJ f  𝖺[e ?b, g, d ?c  PRJ h]) 
                   (PRJ f  𝖺[e ?b  g, d ?c, PRJ h]) 
                   (PRJ f  𝖺-1[e ?b, g, d ?c]  PRJ h) 
                   𝖺[PRJ f, PRJ g, PRJ h]"

      have E: "?LHS = ?RHS"
      proof -
        have "?LHS = ?LHSt"
          using assms B.α_def B.α'.map_ide_simp B.VVV.ide_charSbC B.VVV.arr_charSbC
                B.VV.ide_charSbC B.VV.arr_charSbC P_def
          by simp
        also have "... = ?RHSt"
          using assms by (intro EV.eval_eqI, auto)
        also have "... = ?RHS"
          using assms B.α_def B.α'.map_ide_simp B.VVV.ide_charSbC B.VVV.arr_charSbC
                B.VV.ide_charSbC B.VV.arr_charSbC P_def
          by simp
        finally show ?thesis by blast
      qed
      show ?thesis
        using L R E by argo
    qed

    interpretation CMP: transformation_by_components B.VV.comp comp HoPP.map PoH.map
                          λμν. CMP (fst μν) (snd μν)
    proof
      show 1: "fg. B.VV.ide fg  «CMP (fst fg) (snd fg) : HoPP.map fg  PoH.map fg»"
        using CMP_in_hom(2) B.VV.ide_charSbC B.VV.arr_charSbC P.FF_def hcomp_def arr_charSbC
              P0_props(1) P.preserves_arr
        by auto
      show "fg. B.VV.arr fg 
                   CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg))  HoPP.map fg =
                   PoH.map fg  CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
      proof -
        fix fg
        assume fg: "B.VV.arr fg"
        have "CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg))  HoPP.map fg =
              CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg)) B HoPP.map fg"
          using fg 1 B.VV.ide_charSbC B.VV.arr_charSbC B.VV.cod_charSbC HoPP.preserves_arr
                P.FF_def hcomp_char comp_char
          by auto
        also have "... = PoH.map fg B CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
          using fg CMP_naturality [of "fst fg" "snd fg"]
                B.VV.ide_charSbC B.VV.arr_charSbC B.VV.dom_charSbC B.VV.cod_charSbC P.FF_def
                hcomp_def arr_charSbC P0_props(1) P.preserves_arr
          by auto
        also have "... = PoH.map fg  CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
        proof -
          have "P0 (srcB (snd fg))  Obj  P0 (trgB (fst fg))  Obj"
            using fg P0_props(6) B.VV.arrE B.obj_src B.obj_trg by meson
          moreover have "B.seq (P (fst fg B snd fg))
                               (CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg)))"
            using fg 1 B.VV.ide_charSbC B.VV.arr_charSbC B.VV.dom_charSbC by simp
          ultimately show ?thesis
            using fg 1 comp_char arr_charSbC in_hom_charSbC by simp
        qed
        finally show "CMP (fst (B.VV.cod fg)) (snd (B.VV.cod fg))  HoPP.map fg =
                      PoH.map fg  CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))"
          by blast
      qed
    qed

    interpretation CMP: natural_isomorphism B.VV.comp comp HoPP.map PoH.map CMP.map
      using iso_CMP B.VV.ide_charSbC B.VV.arr_charSbC CMP.map_simp_ide
      by unfold_locales simp

    definition ΦP
    where "ΦP = CMP.map"

    interpretation ΦP: natural_isomorphism B.VV.comp comp HoPP.map PoH.map ΦP
      unfolding ΦP_def
      using CMP.natural_isomorphism_axioms by simp

    no_notation B.in_hom  ("«_ : _ B _»")

    lemma ΦP_in_hom [intro]:
    assumes "B.ide f" and "B.ide g" and "srcB f = trgB g"
    shows "«ΦP (f, g) : src (P g)  trg (P f)»"
    and "«ΦP (f, g) : P f  P g  P (f B g)»"
    proof -
      show 1: "«ΦP (f, g) : P f  P g  P (f B g)»"
        using assms B.VV.ide_charSbC B.VV.arr_charSbC P.FF_def by auto
      show "«ΦP (f, g) : src (P g)  trg (P f)»"
        using 1
        by (metis (no_types, lifting) hcomp_simps(2) in_hhom_def in_hom_charSbC
            src_hcomp vconn_implies_hpar(1-2))
    qed

    lemma ΦP_simps [simp]:
    assumes "B.ide f" and "B.ide g" and "srcB f = trgB g"
    shows "arr (ΦP (f, g))"
    and "src (ΦP (f, g)) = src (P g)"
    and "trg (ΦP (f, g)) = trg (P f)"
    and "dom (ΦP (f, g)) = P f  P g"
    and "cod (ΦP (f, g)) = P (f B g)"
      using assms ΦP_in_hom by blast+

    sublocale prj: pseudofunctor VB HB 𝖺B 𝗂B srcB trgB comp hcomp 𝖺 𝗂B src trg P ΦP
    proof
      fix f g h
      assume f: "B.ide f" and g: "B.ide g" and h: "B.ide h"
      and fg: "srcB f = trgB g" and gh: "srcB g = trgB h"
      have 1: "ide (P f)  ide (P g)  ide (P h)"
        using f g h P.preserves_ide by simp
      have "P 𝖺B[f, g, h]  ΦP (f B g, h)  (ΦP (f, g)  P h) =
            P 𝖺B[f, g, h] B ΦP (f B g, h) B (ΦP (f, g) B P h)"
        using f g h fg gh B.VV.arr_charSbC B.VV.dom_charSbC B.VV.cod_charSbC P.FF_def
        by (intro comp_eqI hcomp_eqI seqI hseqI') auto
      also have "... = CMP f (g B h) B (P f B CMP g h) B 𝖺B[P f, P g, P h]"
        using f g h fg gh CMP_coherence CMP.map_simp_ide B.VV.ide_charSbC B.VV.arr_charSbC
              ΦP_def
        by simp
      also have "... = ΦP (f, g B h) B (P f B ΦP (g, h)) B 𝖺B[P f, P g, P h]"
        using f g h fg gh B.VV.ide_charSbC B.VV.arr_charSbC ΦP_def by simp
      also have "... = ΦP (f, g B h)  (P f B ΦP (g, h))  𝖺B[P f, P g, P h]"
      proof -
        have 2: "arr ((P f  ΦP (g, h))  𝖺[P f, P g, P h])"
          using f g h fg gh B.VV.arr_charSbC B.VV.dom_charSbC P.FF_def by auto
        moreover have "(P f  ΦP (g, h))  𝖺[P f, P g, P h] =
                       (P f B ΦP (g, h)) B 𝖺B[P f, P g, P h]"
          using f g h fg gh 2 comp_eqI hcomp_eqI assoc_simp by auto
        moreover have 3: "arr (P f B ΦP (g, h))"
          using f g h fg gh arr_hcompI
          by (intro arr_hcompI hseqI') auto
        moreover have "B.dom (P f B ΦP (g, h)) = P f B P g B P h"
        proof -
          have "B.dom (P f B ΦP (g, h)) = P f B B.dom (ΦP (g, h))"
            using f g h fg gh 3 B.hcomp_simps(3)
            by (metis (no_types, lifting) 1 arrE ideD(1) ideD(2) dom_charSbC)
          also have "... = P f B P g B P h"
            using g h gh dom_charSbC hcomp_char ΦP_simps(1-4) by auto
          finally show ?thesis by blast
        qed
        moreover have "B.dom (ΦP (f, g B h)) =
                       B.cod ((P f B ΦP (g, h)) B 𝖺B[P f, P g, P h])"
        proof -
          have "B.dom (ΦP (f, g B h)) = dom (ΦP (f, g B h))"
            using f g h fg gh B.VV.arr_charSbC dom_charSbC by simp
          also have "... = cod ((P f  ΦP (g, h))  𝖺[P f, P g, P h])"
            using f g h fg gh 2 VV.arr_charSbC ΦP_simps by simp
          also have "... = B.cod ((P f  ΦP (g, h))  𝖺[P f, P g, P h])"
            using 2 cod_charSbC by presburger
          also have "... = B.cod ((P f B ΦP (g, h)) B 𝖺B[P f, P g, P h])"
          proof -
            have "(P f  ΦP (g, h))  𝖺[P f, P g, P h] = (P f B ΦP (g, h)) B 𝖺B[P f, P g, P h]"
             using f g h fg gh 2 comp_eqI hcomp_eqI assoc_simp by auto
            thus ?thesis by presburger
          qed
          finally show ?thesis by blast
        qed
        moreover have "B.cod 𝖺B[P f, P g, P h] = P f B P g B P h"
          using f g h fg gh 1 ide_charSbC by auto
        ultimately show ?thesis
          using f g h fg gh B.VV.arr_charSbC assoc_simp assoc_simps(1) dom_charSbC cod_charSbC
          by (intro comp_eqI' seqI arr_compI arr_hcompI) auto
      qed
      also have "... = ΦP (f, g B h)  (P f  ΦP (g, h))  𝖺[P f, P g, P h]"
        using f g h fg gh assoc_simp hcomp_eqI ΦP_simps(1) by auto
      finally show "P 𝖺B[f, g, h]  ΦP (f B g, h)  (ΦP (f, g)  P h) =
                    ΦP (f, g B h)  (P f  ΦP (g, h))  𝖺[P f, P g, P h]"
        by blast
    qed

    lemma pseudofunctor_prj:
    shows "pseudofunctor VB HB 𝖺B 𝗂B srcB trgB comp hcomp 𝖺 𝗂B src trg P ΦP"
      ..

    text ‹
      We need an explicit formula for the unit constraints for P›.
    ›

    lemma prj_unit_char:
    assumes "B.obj a"
    shows "prj.unit a = (e a B 𝗅B-1[d a]) B B.inv (ε a)"
    proof -
      text ‹
        We eventually need one of the triangle identities from the following interpretation.
        However in the meantime, it contains a lot of simps that make an otherwise arduous
        calculation much easier.  So interpret it up front.
      ›
      interpret adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB e a d a η a ε a
        using assms chosen_adjoint_equivalence(1) by simp

      let ?x = "(e a B 𝗅B-1[d a]) B B.inv (ε a)"
      have x: "«?x : P.map0 a B P a»"
        using assms P_def P.map0_def P_map0_simp
        by (intro B.comp_in_homI') auto
      have "?x = prj.unit a"
      proof (intro prj.unit_eqI)
        show "B.obj a" by fact
        have a_da: "«a B d a : P0 a B a»  B.ide (a B d a)"
          using assms B.obj_simps by auto
        show xS: "«?x : P.map0 a  P a»"
          using assms x P_map0_simp P_simpsB(1) arr_charSbC B.arrI
                equivalence_data_simpsB(6) counit_simps(1,4) B.obj_simps(1)
                B.src.preserves_arr B.vconn_implies_hpar(1-4)
          by (metis (no_types, lifting) P_simps(1) in_hom_charSbC)
        show "iso ?x"
          using assms xS B.isos_compose iso_charSbC arr_charSbC by auto
        have *: "«?x : P0 a  P0 a»"
          using assms xS P0_props vconn_implies_hpar(1-2)
          by (intro in_hhomI) auto
        show "?x  𝗂B[P.map0 a] = (P 𝗂B[a]  ΦP (a, a))  (?x  ?x)"
        proof -
          have 0: "«d a B e a B a B d a : P0 a B a»"
            using assms by force
          have 1: "B.seq (B.inv (η a) B a B d a) 𝖺B-1[d a, e a, a B d a]"
            using assms by (elim B.objE, intro B.seqI) auto
          have "(P 𝗂B[a]  ΦP (a, a))  (?x  ?x) = P 𝗂B[a]  ΦP (a, a)  (?x  ?x)"
            using comp_assoc by simp
          also have "... = P 𝗂B[a]  (CMP a a  (P a  P a))  (?x  ?x)"
            unfolding ΦP_def CMP.map_def
            using assms B.VV.arr_charSbC B.VV.cod_charSbC P.FF_def by auto
          also have "... = P 𝗂B[a]  (CMP a a  (P a B P a))  (?x B ?x)"
            using assms xS hcomp_char src_def trg_def by auto
          also have "... = P 𝗂B[a] B (CMP a a B (P a B P a)) B (?x B ?x)"
          proof -
            have "«P 𝗂B[a]  (CMP a a  (P a B P a))  (?x B ?x) : P0 a  P0 a  P a»"
            proof (intro comp_in_homI)
              show "«?x B ?x : P0 a  P0 a  P a  P a»"
              proof -
                have "«?x  ?x : P0 a  P0 a  P a  P a»"
                  using assms xS * P0_props P.map0_simps(3)
                  by (intro hcomp_in_vhom) auto
                moreover have "?x  ?x = ?x B ?x"
                  using xS * by (intro hcomp_eqI) auto
                ultimately show ?thesis by simp
              qed
              show "«P a B P a : P a  P a  P a  P a»"
                using assms hcomp_eqI by fastforce
              show "«CMP a a : P a  P a  P (a B a)»"
                using assms CMP_in_hom(2) by auto
              show "«P 𝗂B[a] : P (a B a)  P a»"
                using assms by auto
            qed
            thus ?thesis
              by (intro comp_eqI hcomp_eqI) auto
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a] B (B.inv (η a) B a B d a)) B
                           (e a B a B 𝖺B-1[d a, e a, a B d a]) B
                           𝖺B[e a, a, d a B P a] B
                           𝖺B[e a B a, d a, P a] B
                           (𝖺B-1[e a, a, d a] B P a) B
                           (P a B P a) B
                           (?x B ?x)"
            using assms B.comp_assoc CMP_def by auto
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           (e a B a B B.inv (η a) B a B d a) B
                           (e a B a B 𝖺B-1[d a, e a, a B d a]) B
                           𝖺B[e a, a, d a B P a] B
                           𝖺B[e a B a, d a, P a] B
                           (𝖺B-1[e a, a, d a] B P a) B
                           (P a B P a) B
                           (?x B ?x)"
            using assms 1 B.whisker_left B.comp_assoc by fastforce
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           (e a B a B B.inv (η a) B a B d a) B
                           (e a B a B 𝖺B-1[d a, e a, a B d a]) B
                           𝖺B[e a, a, d a B P a] B
                           𝖺B[e a B a, d a, P a] B
                           (𝖺B-1[e a, a, d a] B P a) B
                           ((P a B P a) B
                           ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a]))) B
                           (B.inv (ε a) B B.inv (ε a))"
            using assms B.interchange B.comp_assoc by simp
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           ((e a B a B B.inv (η a) B a B d a) B
                           (e a B a B 𝖺B-1[d a, e a, a B d a])) B
                           𝖺B[e a, a, d a B P a] B
                           𝖺B[e a B a, d a, P a] B
                           (𝖺B-1[e a, a, d a] B P a) B
                           ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a])) B
                           (B.inv (ε a) B B.inv (ε a))"
          proof -
            have "(P a B P a) B ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a]))
                    = ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a]))"
              unfolding P_def
              using assms B.comp_cod_arr [of "(e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a])"
                                             "(e a B a B d a) B (e a B a B d a)"]
              by auto
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           ((e a B a B (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]) B
                           𝖺B[e a, a, d a B P a]) B
                           𝖺B[e a B a, d a, P a] B
                           (𝖺B-1[e a, a, d a] B P a) B
                           ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a])) B
                           (B.inv (ε a) B B.inv (ε a))"
          proof -
            have "(e a B a B B.inv (η a) B a B d a) B
                  (e a B a B 𝖺B-1[d a, e a, a B d a])
                    = e a B (a B B.inv (η a) B a B d a) B (a B 𝖺B-1[d a, e a, a B d a])"
            proof -
              have "B.seq (a B B.inv (η a) B a B d a) (a B 𝖺B-1[d a, e a, a B d a])"
                using assms by (elim B.objE, intro B.seqI) auto
              thus ?thesis
                using assms B.whisker_left by simp
            qed
            also have "... = e a B a B (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]"
              using assms 1 B.whisker_left by fastforce
            finally have "(e a B a B B.inv (η a) B a B d a) B
                          (e a B a B 𝖺B-1[d a, e a, a B d a])
                            = e a B a B (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]"
              by blast
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           ((e a B a) B (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]) B
                           (𝖺B[e a B a, d a, P a] B
                           (𝖺B-1[e a, a, d a] B P a)) B
                           ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a])) B
                           (B.inv (ε a) B B.inv (ε a))"
          proof -
            have "(e a B a B
                    (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]) B 𝖺B[e a, a, d a B P a]
                    = 𝖺B[e a, a, a B a B d a] B
                      ((e a B a) B (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a])"
              using assms 1 a_da P_def
                    B.assoc_naturality
                      [of "e a" a "(B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]"]
              by fastforce
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           (((e a B a) B (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]) B
                           𝖺B-1[e a, a, d a B P a]) B
                           (e a B 𝖺B[a, d a, P a]) B
                           𝖺B[e a, a B d a, P a] B
                           ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a])) B
                           (B.inv (ε a) B B.inv (ε a))"
          proof -
            have 1: "B.ide (e a)  B.ide a  B.ide (d a)  B.ide (P a)"
              using assms ide_charSbC P.preserves_ide by auto
            have 2: "srcB (e a) = trgB a  srcB a = trgB (d a)  srcB (d a) = trgB (P a)"
              using assms by auto
            have "((e a B 𝖺B[a, d a, P a]) B 𝖺B[e a, a B d a, P a]) B (𝖺B[e a, a, d a] B P a)
                    = 𝖺B[e a, a, d a B P a] B 𝖺B[e a B a, d a, P a]"
              using assms 1 2 B.pentagon B.comp_assoc by force
            hence "(e a B 𝖺B[a, d a, P a]) B 𝖺B[e a, a B d a, P a]
                     = (𝖺B[e a, a, d a B P a] B 𝖺B[e a B a, d a, P a]) B
                       (𝖺B-1[e a, a, d a] B P a)"
              using assms 1 2
                    B.invert_side_of_triangle(2)
                      [of "𝖺B[e a, a, d a B P a] B 𝖺B[e a B a, d a, P a]"
                          "(e a B 𝖺B[a, d a, P a]) B 𝖺B[e a, a B d a, P a]"
                          "𝖺B[e a, a, d a] B P a"]
                by fastforce
            hence "𝖺B[e a B a, d a, P a] B (𝖺B-1[e a, a, d a] B P a)
                     = 𝖺B-1[e a, a, d a B P a] B (e a B 𝖺B[a, d a, P a]) B
                       𝖺B[e a, a B d a, P a]"
              using assms 1 2 B.invert_side_of_triangle(1) B.comp_assoc by fastforce
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           𝖺B-1[e a, a, a B a B d a] B
                           (e a B a B (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]) B
                           ((e a B 𝖺B[a, d a, P a]) B
                           𝖺B[e a, a B d a, P a]) B
                           ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a])) B
                           (B.inv (ε a) B B.inv (ε a))"
          proof -
            have "B.arr (e a)  B.arr a"
              using assms by auto
            moreover have "B.dom (e a) = e a  B.dom a = a  B.cod a = a 
                           B.cod (e a) = e a"
              using assms by auto
            moreover have "B.seq (B.inv (η a) B a B d a) 𝖺B-1[d a, e a, a B d a]"
              using assms by (elim B.objE, intro B.seqI) auto
            moreover have "srcB a =
                           trgB ((B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a])"
              using assms a_da trg_vcomp by fastforce
            moreover have "srcB a = a  trgB a = a"
              using assms by auto
            moreover have "B.dom ((B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a])
                             = d a B e a B a B d a"
              using assms a_da by fastforce
            moreover have "B.cod ((B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a])
                             = a B a B d a"
              using assms a_da by fastforce
            ultimately have "((e a B a) B
                                (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]) B
                                 𝖺B-1[e a, a, d a B P a]
                               = 𝖺B-1[e a, a, a B a B d a] B
                                 (e a B a B (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a])"
              unfolding P_def
              using assms
                    B.assoc'_naturality
                      [of "e a" a "(B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]"]
              by auto
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           𝖺B-1[e a, a, a B a B d a] B
                           ((e a B a B (B.inv (η a) B a B d a) B
                           𝖺B-1[d a, e a, a B d a]) B
                           𝖺B[e a, a, d a B P a]) B
                           𝖺B[e a B a, d a, P a] B
                           (𝖺B-1[e a, a, d a] B P a) B
                           ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a])) B
                           (B.inv (ε a) B B.inv (ε a))"
          proof -
            have 1: "B.ide (e a)  B.ide a  B.ide (d a)  B.ide (P a)"
              using assms ide_charSbC P.preserves_ide by auto
            have 2: "srcB (e a) = trgB a  srcB a = trgB (d a)  srcB (d a) = trgB (P a)"
              using assms by auto
            have "((e a B 𝖺B[a, d a, P a]) B 𝖺B[e a, a B d a, P a]) B (𝖺B[e a, a, d a] B P a)
                    = 𝖺B[e a, a, d a B P a] B 𝖺B[e a B a, d a, P a]"
              using assms 1 2 B.pentagon B.comp_assoc by fastforce
            hence "(e a B 𝖺B[a, d a, P a]) B 𝖺B[e a, a B d a, P a]
                     = 𝖺B[e a, a, d a B P a] B 𝖺B[e a B a, d a, P a] B
                       (𝖺B-1[e a, a, d a] B P a)"
              using assms 1 2 P.preserves_ide B.comp_assoc
                    B.invert_side_of_triangle(2)
                      [of "𝖺B[e a, a, d a B P a] B 𝖺B[e a B a, d a, P a]"
                          "(e a B 𝖺B[a, d a, P a]) B 𝖺B[e a, a B d a, P a]"
                          "𝖺B[e a, a, d a] B P a"]
              by force
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           𝖺B-1[e a, a, a B a B d a] B
                           𝖺B[e a, a, a B a B d a] B
                           ((e a B a) B (B.inv (η a) B a B d a) B
                           𝖺B-1[d a, e a, a B d a]) B
                           𝖺B[e a B a, d a, P a] B
                           (𝖺B-1[e a, a, d a] B P a) B
                           ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a])) B
                           (B.inv (ε a) B B.inv (ε a))"
          proof -
            have "(e a B a B (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]) B
                  𝖺B[e a, a, d a B P a]
                    = 𝖺B[e a, a, a B a B d a] B ((e a B a) B (B.inv (η a) B a B d a) B
                      𝖺B-1[d a, e a, a B d a])"
            proof -
              have 1: "B.seq (B.inv (η a) B a B d a) 𝖺B-1[d a, e a, a B d a]"
                using assms by (elim B.objE, intro B.seqI) auto
              moreover have "srcB a =
                             trgB ((B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a])"
                using a_da by force
              moreover have "B.arr a  B.dom a = a  B.cod a = a  srcB a = a  trgB a = a"
                using assms by auto
              moreover have "B.dom ((B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a])
                               = d a B e a B a B d a"
                using a_da by auto
              moreover have "B.cod ((B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a])
                               = a B a B d a"
                using a_da by auto
              ultimately show ?thesis
                unfolding P_def
                using assms
                      B.assoc_naturality
                        [of "e a" a "(B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]"]
                by auto
            qed
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           𝖺B-1[e a, a, a B a B d a] B
                           𝖺B[e a, a, a B a B d a] B
                           ((e a B a) B (B.inv (η a) B a B d a) B
                           𝖺B-1[d a, e a, a B d a]) B
                           𝖺B[e a B a, d a, P a] B
                           ((𝖺B-1[e a, a, d a] B P a) B
                           ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a]) B
                           B.inv (ε a))) B
                           (B.inv (ε a) B P0 a)"
          proof -
            have "((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a])) B
                  (B.inv (ε a) B B.inv (ε a))
                    = ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a]) B B.inv (ε a)) B
                      (B.inv (ε a) B P0 a)"
            proof -
              have "((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a])) B
                    (B.inv (ε a) B B.inv (ε a))
                      = (((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a])) B
                        ((e a B d a) B B.inv (ε a))) B
                        (B.inv (ε a) B P0 a)"
                using assms B.comp_arr_dom B.comp_cod_arr B.comp_assoc
                      B.interchange [of "e a B d a" "B.inv (ε a)" "B.inv (ε a)" "P0 a"]
                by simp
            also have "... = ((e a B 𝗅B-1[d a]) B (e a B d a) B (e a B 𝗅B-1[d a]) B
                             B.inv (ε a)) B (B.inv (ε a) B P0 a)"
                using assms B.interchange by simp
              also have "... = ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a]) B B.inv (ε a)) B
                               (B.inv (ε a) B P0 a)"
                using assms B.comp_arr_dom by simp
              finally show ?thesis by simp
            qed
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           𝖺B-1[e a, a, a B a B d a] B
                           𝖺B[e a, a, a B a B d a] B
                           ((e a B a) B (B.inv (η a) B a B d a) B
                           𝖺B-1[d a, e a, a B d a]) B
                           (𝖺B[e a B a, d a, P a] B
                           (((e a B a) B d a) B (e a B 𝗅B-1[d a]) B B.inv (ε a))) B
                           (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B P0 a) B
                           (B.inv (ε a) B P0 a)"
          proof -
            have "(𝖺B-1[e a, a, d a] B P a) B ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a]) B
                  B.inv (ε a))
                    = (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B P a B (e a B 𝗅B-1[d a]) B
                      B.inv (ε a))"
            proof -
              have "B.seq 𝖺B-1[e a, a, d a] (e a B 𝗅B-1[d a])"
                using assms a_da by (elim B.objE, intro B.seqI) auto
              moreover have "B.seq (P a) ((e a B 𝗅B-1[d a]) B B.inv (ε a))"
                using assms a_da P_def by auto
              ultimately show ?thesis
                using assms B.interchange by simp
            qed
            also have "... = (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a]) B
                             B.inv (ε a))"
            proof -
              have "P a B (e a B 𝗅B-1[d a]) B B.inv (ε a) = (e a B 𝗅B-1[d a]) B B.inv (ε a)"
                using assms x P_def B.comp_cod_arr by blast
              moreover have "𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) = 𝗋B-1[e a] B d a"
              proof -
                have "𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a])
                        = B.inv ((e a B 𝗅B[d a]) B 𝖺B[e a, a, d a])"
                proof -
                  have "B.inv ((e a B 𝗅B[d a]) B 𝖺B[e a, a, d a])
                          = B.inv 𝖺B[e a, a, d a] B B.inv (e a B 𝗅B[d a])"
                  proof -
                    have "B.seq (e a B 𝗅B[d a]) 𝖺B[e a, a, d a]"
                    proof -
                      have "B.iso ((e a B 𝗅B[d a]) B 𝖺B[e a, a, d a])"
                        using assms by (elim B.objE, intro B.isos_compose) auto
                      thus ?thesis
                        by blast
                    qed
                    moreover have "B.iso 𝖺B[e a, a, d a]"
                      using assms by force
                    ultimately show ?thesis
                      using assms B.inv_comp by auto
                  qed
                  also have "... = 𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a])"
                    using assms
                    by (elim B.objE) (simp add: assms)
                  finally show ?thesis by simp
                qed
                also have "... = B.inv (𝗋B[e a] B d a)"
                  using assms B.triangle [of "d a" "e a"] by simp
                also have "... = 𝗋B-1[e a] B d a"
                  using assms by simp
                finally show ?thesis by blast
              qed
              ultimately show ?thesis by simp
            qed
            also have "... = (((e a B a) B d a) B 𝖺B-1[e a, a, d a]) B (e a B 𝗅B-1[d a]) B
                             (e a B 𝗅B-1[d a]) B B.inv (ε a) B P0 a"
              using assms 0 B.comp_cod_arr B.comp_arr_dom
              by (elim B.objE) auto
            also have "... = ((e a B a) B d a) B 𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B
                             ((e a B 𝗅B-1[d a]) B B.inv (ε a)) B P0 a"
              using B.comp_assoc by simp
            also have "... = (((e a B a) B d a) B (e a B 𝗅B-1[d a]) B B.inv (ε a)) B
                             (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B P0 a)"
            proof -
              have "B.seq ((e a B a) B d a) (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]))"
                using assms 0 a_da
                by (elim B.objE, intro B.seqI) auto
              moreover have "B.seq ((e a B 𝗅B-1[d a]) B B.inv (ε a)) (P0 a)"
                using assms 0
                apply (intro B.seqI, auto simp add: B.obj_simps(5) P0_props(5))
                by (meson B.in_hhomE B.obj_simps(1) a_da)
              ultimately show ?thesis
                using assms B.interchange by blast
            qed
            finally have "(𝖺B-1[e a, a, d a] B P a) B
                          ((e a B 𝗅B-1[d a]) B (e a B 𝗅B-1[d a]) B B.inv (ε a))
                            = (((e a B a) B d a) B (e a B 𝗅B-1[d a]) B B.inv (ε a)) B
                              (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B P0 a)"
              by blast
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           𝖺B-1[e a, a, a B a B d a] B
                           𝖺B[e a, a, a B a B d a] B
                           (((e a B a) B (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]) B
                           ((e a B a) B d a B (e a B 𝗅B-1[d a]) B B.inv (ε a))) B
                           𝖺B[e a B a, d a, P0 a] B
                           (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B P0 a) B
                           (B.inv (ε a) B P0 a)"
          proof -
            have "𝖺B[e a B a, d a, P a] B
                  (((e a B a) B d a) B (e a B 𝗅B-1[d a]) B B.inv (ε a))
                    = ((e a B a) B d a B (e a B 𝗅B-1[d a]) B B.inv (ε a)) B
                      𝖺B[e a B a, d a, P0 a]"
            proof -
              have "B.hseq (e a) a"
                using assms by force
              moreover have "srcB (d a) = trgB ?x"
                using assms B.trg_vcomp [of "e a B 𝗅B-1[d a]" "B.inv (ε a)"] by simp
              moreover have "B.cod ((e a B 𝗅B-1[d a]) B B.inv (ε a)) = P a"
                using assms x by blast
              ultimately show ?thesis
                using assms B.assoc_naturality [of "e a B a" "d a" ?x] by auto
            qed
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           𝖺B-1[e a, a, a B a B d a] B
                           𝖺B[e a, a, a B a B d a] B
                           ((e a B a) B
                              (B.inv (η a) B a B d a) B
                              𝖺B-1[d a, e a, a B d a] B
                              (d a B (e a B 𝗅B-1[d a]) B B.inv (ε a))) B
                           𝖺B[e a B a, d a, P0 a] B
                           (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B P0 a) B
                           (B.inv (ε a) B P0 a)"
          proof -
            have "((e a B a) B (B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a]) B
                  ((e a B a) B d a B (e a B 𝗅B-1[d a]) B B.inv (ε a))
                    = (e a B a) B (B.inv (η a) B a B d a) B
                      𝖺B-1[d a, e a, a B d a] B
                      (d a B (e a B 𝗅B-1[d a]) B B.inv (ε a))"
            proof -
              have "B.ide (e a B a)"
                using assms by force
              moreover have "B.seq ((B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a])
                                   (d a B (e a B 𝗅B-1[d a]) B B.inv (ε a))"
                using a_da B.whisker_left by auto
              ultimately show ?thesis
                using assms B.whisker_left B.comp_assoc by fastforce
            qed
            thus ?thesis by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           𝖺B-1[e a, a, a B a B d a] B
                           𝖺B[e a, a, a B a B d a] B
                           ((e a B a) B (a B 𝗅B-1[d a]) B
                                        (B.inv (η a) B d a) B
                                        𝖺B-1[d a, e a, d a] B
                                        (d a B B.inv (ε a))) B
                           𝖺B[e a B a, d a, P0 a] B
                           (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B P0 a) B
                           (B.inv (ε a) B P0 a)"
          proof -
            have "(B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a] B
                  (d a B (e a B 𝗅B-1[d a]) B B.inv (ε a))
                    = (a B 𝗅B-1[d a]) B (B.inv (η a) B d a) B 𝖺B-1[d a, e a, d a] B
                      (d a B B.inv (ε a))"
            proof -
              have "(B.inv (η a) B a B d a) B 𝖺B-1[d a, e a, a B d a] B
                    (d a B (e a B 𝗅B-1[d a]) B B.inv (ε a))
                      = (B.inv (η a) B a B d a) B (𝖺B-1[d a, e a, a B d a] B
                        (d a B e a B 𝗅B-1[d a])) B (d a B B.inv (ε a))"
                using assms B.whisker_left B.comp_assoc by simp
              also have "... = ((B.inv (η a) B a B d a) B ((d a B e a) B 𝗅B-1[d a])) B
                               𝖺B-1[d a, e a, d a] B (d a B B.inv (ε a))"
               using assms B.assoc'_naturality [of "d a" "e a" "𝗅B-1[d a]"] B.comp_assoc by simp
              also have "... = (B.inv (η a) B 𝗅B-1[d a]) B 𝖺B-1[d a, e a, d a] B
                               (d a B B.inv (ε a))"
                using assms B.interchange [of "B.inv (η a)" "d a B e a" "a B d a" "𝗅B-1[d a]"]
                      B.comp_arr_dom B.comp_cod_arr
                by simp
              also have "... = (a B 𝗅B-1[d a]) B (B.inv (η a) B d a) B 𝖺B-1[d a, e a, d a] B
                               (d a B B.inv (ε a))"
                using assms B.interchange [of a "B.inv (η a)" "𝗅B-1[d a]" "d a"]
                      B.comp_arr_dom B.comp_cod_arr B.comp_assoc
                by simp
              finally show ?thesis by blast
            qed
            thus ?thesis
              using comp_assoc by simp
          qed
          also have "... = P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           𝖺B-1[e a, a, a B a B d a] B
                           𝖺B[e a, a, a B a B d a] B
                           ((e a B a) B (a B 𝗅B-1[d a]) B 𝗅B-1[d a] B 𝗋B[d a]) B
                           𝖺B[e a B a, d a, P0 a] B
                           (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B P0 a) B
                           (B.inv (ε a) B P0 a)"
          proof -
            have "(B.inv (η a) B d a) B 𝖺B-1[d a, e a, d a] B (d a B B.inv (ε a))
                    = 𝗅B-1[d a] B 𝗋B[d a]"
            proof -
              have "(B.inv (η a) B d a) B 𝖺B-1[d a, e a, d a] B (d a B B.inv (ε a))
                      = B.inv (η a B d a) B B.inv 𝖺B[d a, e a, d a] B B.inv (d a B ε a)"
                using assms by simp
              also have "... = B.inv ((d a B ε a) B 𝖺B[d a, e a, d a] B (η a B d a))"
              proof -
                have "B.iso ((d a B ε a) B 𝖺B[d a, e a, d a] B (η a B d a))"
                  using assms by (intro B.isos_compose) auto
               thus ?thesis
                  using assms B.inv_comp_left B.comp_assoc by auto
              qed
              also have "... = B.inv (𝗋B-1[d a] B 𝗅B[d a])"
                using assms triangle_right by simp
              also have "... = 𝗅B-1[d a] B 𝗋B[d a]"
                using assms B.inv_comp by simp
              finally show ?thesis by blast
            qed
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (P 𝗂B[a] B
                           (e a B 𝖺B-1[a, a, d a]) B
                           (e a B a B 𝗅B[a B d a]) B
                           𝖺B[e a, a, a B a B d a] B
                           𝖺B-1[e a, a, a B a B d a] B
                           𝖺B[e a, a, a B a B d a] B
                           ((e a B a) B (a B 𝗅B-1[d a]) B 𝗅B-1[d a] B 𝗋B[d a]) B
                           𝖺B[e a B a, d a, P0 a] B
                           (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B P0 a) B
                           𝗋B-1[e a B d a]) B
                           B.inv (ε a) B
                           𝗂B[P0 a]"
          proof -
            have "𝗋B[e a B d a] B (B.inv (ε a) B P0 a) = B.inv (ε a) B 𝗂B[P0 a]"
              using assms B.runit_naturality [of "B.inv (ε a)"] B.unitor_coincidence P0_props
              by simp
            hence "B.inv (ε a) B P0 a = 𝗋B-1[e a B d a] B B.inv (ε a) B 𝗂B[P0 a]"
               using assms P0_props(5) B.invert_side_of_triangle(1) by simp
            thus ?thesis
               using B.comp_assoc by simp
          qed
          also have "... = (e a B 𝗅B-1[d a]) B B.inv (ε a) B 𝗂B[P0 a]"
          proof -
            have P0_a: "B.obj (P0 a)  B.arr (P0 a)"
              using assms a_da by fastforce
            have i_a: "B.𝔯 a = 𝗂B[a]"
              using assms B.unitor_coincidence B.𝔯_ide_simp B.obj_simps
              by (metis (no_types, lifting) B.objE)
            have "P 𝗂B[a] B
                  (e a B 𝖺B-1[a, a, d a]) B
                  (e a B a B 𝗅B[a B d a]) B
                   𝖺B[e a, a, a B a B d a] B
                   𝖺B-1[e a, a, a B a B d a] B
                   𝖺B[e a, a, a B a B d a] B
                   ((e a B a) B (a B 𝗅B-1[d a]) B 𝗅B-1[d a] B 𝗋B[d a]) B
                   𝖺B[e a B a, d a, P0 a] B
                   (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B P0 a) B
                   𝗋B-1[e a B d a]
                     = e a B 𝗅B-1[d a]"
            proof -
              have "P 𝗂B[a] B
                    (e a B 𝖺B-1[a, a, d a]) B
                    (e a B a B 𝗅B[a B d a]) B
                    𝖺B[e a, a, a B a B d a] B
                    𝖺B-1[e a, a, a B a B d a] B
                    𝖺B[e a, a, a B a B d a] B
                    ((e a B a) B (a B 𝗅B-1[d a]) B 𝗅B-1[d a] B 𝗋B[d a]) B
                    𝖺B[e a B a, d a, P0 a] B
                    (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B P0 a) B
                    𝗋B-1[e a B d a]
                      = (e a  𝗋[a0]  d a) 
                         (e a  𝖺-1[a0, a0, d a]) 
                         (e a  a0  𝗅[a0  d a]) 
                         𝖺[e a, a0, a0  a0  d a] 
                         𝖺-1[e a, a0, a0  a0  d a] 
                         𝖺[e a, a0, a0  a0  d a] 
                         ((e a  a0)  (a0  𝗅-1[d a])  𝗅-1[d a]  𝗋[d a]) 
                         𝖺[e a  a0, d a, P0 a0] 
                         (𝖺-1[e a, a0, d a]  (e a  𝗅-1[d a])  P0 a0) 
                         𝗋-1[e a  d a]"
                unfolding P_def
                using assms B.obj_def [of a] P0_a i_a B.α_def B.α'.map_ide_simp
                      B.VVV.ide_charSbC B.VVV.arr_charSbC B.VV.ide_charSbC B.VV.arr_charSbC
                      B.𝔩_ide_simp B.𝔯_ide_simp
                by (elim B.objE) simp
              also have "... = e a  𝗅-1[d a]"
                using assms P0_a by (intro EV.eval_eqI) auto
              also have "... = e a B 𝗅B-1[d a]"
                using assms P_def B.𝔩_ide_simp by simp
              finally show ?thesis by blast
            qed
            thus ?thesis by simp
          qed
          also have "... = ?x  𝗂B[P.map0 a]"
          proof -
            have "B.arr 𝗂B[P.map0 a]  srcB 𝗂B[P.map0 a]  Obj  trgB 𝗂B[P.map0 a]  Obj 
                  P0 a B P0 a  Obj  B.cod 𝗂B[P.map0 a] = P0 a"
              using assms P0_props arr_charSbC unit_simps(1)
              apply auto
              using obj_char
              by (metis (no_types, lifting) B.comp_ide_self B.objE)
            thus ?thesis
              using assms comp_def B.comp_assoc P0_props [of a] by simp
          qed
          finally show ?thesis by argo
        qed
      qed
      thus ?thesis by simp
    qed

    sublocale PoE: composite_pseudofunctor
                     comp hcomp 𝖺 𝗂B src trg VB HB 𝖺B 𝗂B srcB trgB
                     comp hcomp 𝖺 𝗂B src trg
                     E ΦE P ΦP
      ..
    sublocale EoP: composite_pseudofunctor
                     VB HB 𝖺B 𝗂B srcB trgB comp hcomp 𝖺 𝗂B src trg VB HB 𝖺B 𝗂B srcB trgB
                     P ΦP E ΦE
      ..
    sublocale IC: identity_pseudofunctor VB HB 𝖺B 𝗂B srcB trgB ..
    sublocale IS: identity_pseudofunctor comp hcomp 𝖺 𝗂B src trg ..

    no_notation B.in_hom  ("«_ : _ B _»")

    abbreviation (input) unit0
    where "unit0  e"

    definition unit1
    where "unit1 f = (e (trgB f) B 𝗋B[f]) B
                     𝖺B[e (trgB f), f, srcB f] B
                     ((e (trgB f) B f) B B.inv (η (srcB f))) B
                     𝖺B[e (trgB f) B f, d (srcB f), e (srcB f)] B
                     (𝖺B-1[e (trgB f), f, d (srcB f)] B e (srcB f))"

    abbreviation (input) η0
    where "η0  unit0"

    abbreviation (input) η1
    where "η1  unit1"

    lemma unit1_in_hom [intro]:
    assumes "B.ide f"
    shows "«η1 f : srcB f B P0 (trgB f)»"
    and "«η1 f : (e (trgB f) B f B d (srcB f)) B e (srcB f) B e (trgB f) B f»"
    proof -
      show "«η1 f : (e (trgB f) B f B d (srcB f)) B e (srcB f) B e (trgB f) B f»"
        using assms unit1_def by force
      thus "«η1 f : srcB f B P0 (trgB f)»"
        using assms B.vconn_implies_hpar(1-2) by auto
    qed

    lemma unit1_simps [simp]:
    assumes "B.ide f"
    shows "B.arr (η1 f)"
    and "srcB (η1 f) = srcB f" and "trgB (η1 f) = P0 (trgB f)"
    and "B.dom (η1 f) = (e (trgB f) B f B d (srcB f)) B e (srcB f)"
    and "B.cod (η1 f) = e (trgB f) B f"
    and "B.iso (η1 f)"
      using assms unit1_in_hom
           apply auto
      unfolding unit1_def
      by (intro B.isos_compose) auto (* 10 sec *)

    lemma unit1_in_homS [intro]:
    assumes "ide f"
    shows "«η1 f : src f  P0 (trg f)»"
    and "«η1 f : PoE.map f  e (src f)  e (trg f)  IS.map f»"
      using assms ide_charSbC arr_charSbC P0_props(1,6) P.preserves_ide hcomp_def
            src_def trg_def P_def emb.map_def equivalence_data_simps(2)
            in_hom_charSbC
      by auto

    lemma unit1_simpsS [simp]:
    assumes "ide f"
    shows "arr (η1 f)"
    and "src (η1 f) = src f" and "trg (η1 f) = P0 (trg f)"
    and "dom (η1 f) = PoE.map f  e (src f)"
    and "cod (η1 f) = e (trg f)  IS.map f"
    and "iso (η1 f)"
      using assms unit1_in_homS iso_charSbC ide_charSbC arr_charSbC P.as_nat_iso.components_are_iso
      by auto

    sublocale unit: pseudonatural_equivalence comp hcomp 𝖺 𝗂B src trg
                      comp hcomp 𝖺 𝗂B src trg
                      IS.map IS.cmp P  E PoE.cmp unit0 unit1
    proof
      show "a. obj a  ide (e a)"
        using obj_char ide_charSbC arr_charSbC P0_props(1) by force
      show "a. obj a  «e a : src (IS.map a)  src ((P  E) a)»"
        using emb.map_def
        apply (intro in_hhomI)
          apply auto
        using obj_char by auto
      show "a. obj a  equivalence_map (e a)"
      proof -
        fix a
        assume a: "obj a"
        interpret Adj': equivalence_in_bicategory comp hcomp 𝖺 𝗂B src trg
                          e a d a η a ε a
          using a by unfold_locales auto
        show "equivalence_map (e a)"
          using Adj'.equivalence_in_bicategory_axioms equivalence_map_def by auto
      qed
      show "f. ide f  «η1 f : PoE.map f  e (src f)  e (trg f)  IS.map f»"
        using unit1_in_homS(2) by simp
      show "f. ide f  iso (η1 f)"
        by simp
      show *: "μ. arr μ 
                    η1 (cod μ)  (PoE.map μ  e (src μ)) = (e (trg μ)  IS.map μ)  η1 (dom μ)"
      proof -
        fix μ
        assume μ: "arr μ"
        have 1: "B.arr μ"
          using μ arr_charSbC by simp
        let ?f = "B.dom μ"
        let ?g = "B.cod μ"
        let ?a = "srcB μ"
        let ?b = "trgB μ"
        have "η1 (cod μ)  (PoE.map μ  e (src μ))
                = ((e ?b B 𝗋B[?g]) B
                  𝖺B[e ?b, ?g, ?a] B
                  ((e ?b B ?g) B B.inv (η ?a)) B
                  𝖺B[e ?b B ?g, d ?a, e ?a] B
                  (𝖺B-1[e ?b, ?g, d ?a] B e ?a)) B
                  ((e ?b B μ B d ?a) B e ?a)"
          unfolding unit1_def P_def emb.map_def hcomp_def src_def
          using μ comp_char cod_simp arr_charSbC P0_props [of ?a] P0_props [of ?b]
          by auto (* 10 sec *)
        also have "... = (e ?b B 𝗋B[?g]) B
                         𝖺B[e ?b, ?g, ?a] B
                         ((e ?b B ?g) B B.inv (η ?a)) B
                         𝖺B[e ?b B ?g, d ?a, e ?a] B
                         (𝖺B-1[e ?b, ?g, d ?a] B e ?a) B
                         ((e ?b B μ B d ?a) B e ?a)"
          using 1 B.comp_assoc by simp
        also have "... = (e ?b B 𝗋B[?g]) B
                         𝖺B[e ?b, ?g, ?a] B
                         ((e ?b B ?g) B B.inv (η ?a)) B
                         𝖺B[e ?b B ?g, d ?a, e ?a] B
                         (𝖺B-1[e ?b, ?g, d ?a] B (e ?b B μ B d ?a) B e ?a B e ?a)"
          using 1 B.interchange [of "𝖺B-1[e ?b, ?g, d ?a]" "e ?b B μ B d ?a" "e ?a" "e ?a"]
          by simp
        also have "... = (e ?b B 𝗋B[?g]) B
                         𝖺B[e ?b, ?g, ?a] B
                         ((e ?b B ?g) B B.inv (η ?a)) B
                         𝖺B[e ?b B ?g, d ?a, e ?a] B
                         (((e ?b B μ) B d ?a) B 𝖺B-1[e ?b, ?f, d ?a] B e ?a B e ?a)"
          using 1 B.assoc'_naturality [of "e ?b" μ "d ?a"] by simp
        also have "... = (e ?b B 𝗋B[?g]) B
                         𝖺B[e ?b, ?g, ?a] B
                         ((e ?b B ?g) B B.inv (η ?a)) B
                         (𝖺B[e ?b B ?g, d ?a, e ?a] B
                         (((e ?b B μ) B d ?a) B e ?a)) B
                         (𝖺B-1[e ?b, ?f, d ?a] B e ?a)"
          using 1 B.comp_assoc
                B.interchange [of "(e ?b B μ) B d ?a" "𝖺B-1[e ?b, ?f, d ?a]" "e ?a" "e ?a" ]
          by simp
        also have "... = (e ?b B 𝗋B[?g]) B
                         𝖺B[e ?b, ?g, ?a] B
                         (((e ?b B ?g) B B.inv (η ?a)) B
                         ((e ?b B μ) B d ?a B e ?a)) B
                         𝖺B[e ?b B ?f, d ?a, e ?a] B
                         (𝖺B-1[e ?b, ?f, d ?a] B e ?a)"
          using 1 B.assoc_naturality [of "e ?b B μ" "d ?a" "e ?a"] B.comp_assoc by simp
        also have "... = (e ?b B 𝗋B[?g]) B
                         𝖺B[e ?b, ?g, ?a] B
                         ((e ?b B μ) B B.inv (η ?a)) B
                         𝖺B[e ?b B ?f, d ?a, e ?a] B
                         (𝖺B-1[e ?b, ?f, d ?a] B e ?a)"
        proof -
          have "((e ?b B ?g) B B.inv (η ?a)) B ((e ?b B μ) B d ?a B e ?a)
                  = (e ?b B ?g) B (e ?b B μ) B B.inv (η ?a) B (d ?a B e ?a)"
            using 1 B.interchange by simp
          also have "... = (e ?b B μ) B B.inv (η ?a)"
            using 1 B.comp_arr_dom B.comp_cod_arr by simp
          finally have "((e ?b B ?g) B B.inv (η ?a)) B ((e ?b B μ) B d ?a B e ?a)
                          = (e ?b B μ) B B.inv (η ?a)"
            by simp
          thus ?thesis by simp
        qed
        also have "... = (e ?b B 𝗋B[?g]) B
                         (𝖺B[e ?b, ?g, ?a] B
                         ((e ?b B μ) B ?a)) B
                         ((e ?b B ?f) B B.inv (η ?a)) B
                         𝖺B[e ?b B ?f, d ?a, e ?a] B
                         (𝖺B-1[e ?b, ?f, d ?a] B e ?a)"
        proof -
          have "(e ?b B μ) B B.inv (η ?a) = (e ?b B μ) B (e ?b B ?f) B ?a B B.inv (η ?a)"
            using 1 B.comp_arr_dom B.comp_cod_arr by simp
          also have "... = ((e ?b B μ) B ?a) B ((e ?b B ?f) B B.inv (η ?a))"
            using 1 B.interchange by simp
          finally have "(e ?b B μ) B B.inv (η ?a)
                          = ((e ?b B μ) B ?a) B ((e ?b B ?f) B B.inv (η ?a))"
            by blast
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = ((e ?b B 𝗋B[?g]) B
                         (e ?b B μ B ?a)) B
                         𝖺B[e ?b, ?f, ?a] B
                         ((e ?b B ?f) B B.inv (η ?a)) B
                         𝖺B[e ?b B ?f, d ?a, e ?a] B
                         (𝖺B-1[e ?b, ?f, d ?a] B e ?a)"
          using 1 B.assoc_naturality [of "e ?b" μ "?a"] B.comp_assoc by simp
        also have "... = (e ?b B μ) B
                         (e ?b B 𝗋B[?f]) B
                         𝖺B[e ?b, ?f, ?a] B
                         ((e ?b B ?f) B B.inv (η ?a)) B
                         𝖺B[e ?b B ?f, d ?a, e ?a] B
                         (𝖺B-1[e ?b, ?f, d ?a] B e ?a)"
        proof -
          have "(e ?b B 𝗋B[?g]) B (e ?b B μ B ?a) = e ?b B e ?b B 𝗋B[?g] B (μ B ?a)"
            using 1 B.interchange [of "e ?b" "e ?b" "𝗋B[?g]" "μ B ?a"] by simp
          also have "... = e ?b B e ?b B μ B 𝗋B[?f]"
            using 1 B.runit_naturality by simp
          also have "... = (e ?b B μ) B (e ?b B 𝗋B[?f])"
            using 1 B.interchange [of "e ?b" "e ?b" μ "𝗋B[?f]"] by simp
          finally have "(e ?b B 𝗋B[?g]) B (e ?b B μ B ?a) = (e ?b B μ) B (e ?b B 𝗋B[?f])"
            by blast
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = (e (trg μ)  IS.map μ)  η1 (dom μ)"
        proof -
          have "arr ((e (trg μ)  IS.map μ)  η1 (dom μ))"
            using μ by simp
          hence 2: "arr ((e ?b B μ) 
                         ((e ?b B 𝗋B[?f]) B
                         𝖺B[e ?b, ?f, ?a] B
                         ((e ?b B ?f) B B.inv (η ?a)) B
                         𝖺B[e ?b B ?f, d ?a, e ?a] B
                         (B.inv 𝖺B[e ?b, ?f, d ?a] B e ?a)))"
            unfolding unit1_def
            using μ 1 arr_charSbC P0_props trg_def hcomp_def dom_charSbC by simp
          show ?thesis
            unfolding unit1_def
            using μ 1 arr_charSbC P0_props trg_def hcomp_def dom_charSbC
            apply simp
            using 2 comp_eqI by blast
        qed
        finally show "η1 (cod μ)  (PoE.map μ  e (src μ))
                        = (e (trg μ)  IS.map μ)  η1 (dom μ)"
          by simp
      qed
      show "a. obj a  (e a  IS.unit a)  𝗋-1[e a]  𝗅[e a] = unit1 a  (PoE.unit a  e a)"
      proof -
        fix a
        assume a: "obj a"
        have 0: "B.obj a"
          using a obj_char by blast
        interpret adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB e a d a η a ε a
          using 0 chosen_adjoint_equivalence(1) [of a] by simp
        have src: "srcB ((e a B 𝗅B-1[d a]) B B.inv (ε a)) = P0 a"
          using 0 B.src_vcomp [of "e a B 𝗅B-1[d a]" "B.inv (ε a)"] emb.map_def by simp
        have trg: "trgB ((e a B 𝗅B-1[d a]) B B.inv (ε a)) = P0 a"
          using 0 B.trg_vcomp [of "e a B 𝗅B-1[d a]" "B.inv (ε a)"] emb.map_def by simp
        have 1: "arr a  srcB a = a"
          using a src_def by auto
        have 2: "seq (P a) ((e a B 𝗅B-1[d a]) B B.inv (ε a))"
        proof -
          have "B.arr (P a B ((e a B 𝗅B-1[d a]) B B.inv (ε a)))"
            using a 0 P_def emb.map_def
            by (elim B.objE, intro B.seqI) auto
          moreover have "arr (P a)"
            using 0 by auto
          moreover have "arr ((e a B 𝗅B-1[d a]) B B.inv (ε a))"
            apply (unfold arr_charSbC)
            by (simp add: 0 P0_props(6) src trg)
          ultimately show ?thesis
            using seq_charSbC by simp
        qed
        have 3: "obj (P0 a)"
          using 0 P0_props arr_charSbC obj_char by blast
        have 4: "B.seq (P a) ((e a B 𝗅B-1[d a]) B B.inv (ε a))"
          using 2 seq_charSbC by simp

        have "η1 a  (PoE.unit a  e a) = η1 a  (P a  (e a B 𝗅B-1[d a]) B B.inv (ε a) B e a)"
        proof -
          have "η1 a  (PoE.unit a  e a) = η1 a  (PoE.unit a B e a)"
            using 0 a hcomp_eqI by force
          moreover have "PoE.unit a = P a  (e a B 𝗅B-1[d a]) B B.inv (ε a)"
            using a obj_char PoE.unit_char' emb.unit_char' prj_unit_char by simp
          ultimately show ?thesis by argo
        qed
        also have "... = η1 a  (P a B (e a B 𝗅B-1[d a]) B B.inv (ε a) B e a)"
          using 2 comp_simp by force
        also have "... = η1 a B (P a B (e a B 𝗅B-1[d a]) B B.inv (ε a) B e a)"
        proof -
          have "arr (η1 a)"
            using a by auto
          moreover have "arr (P a B (e a B 𝗅B-1[d a]) B B.inv (ε a) B e a)"
            by (metis (no_types, lifting) a 0 2 B.src_vcomp B.vseq_implies_hpar(1)
                arr_charSbC arr_compI equivalence_data_in_hom(3) equivalence_data_simpsB(6)
                hcomp_closed in_hom_charSbC src)
          moreover have "B.seq (η1 a) (P a B (e a B 𝗅B-1[d a]) B B.inv (ε a) B e a)"
            using 0 3 P_def obj_char
            by (elim B.objE, intro B.seqI) auto (* 10 sec *)
          ultimately show ?thesis
            using 2 comp_char by meson
        qed
        also have "... = (e a B 𝗋B[a]) B
                         𝖺B[e a, a, a] B
                         ((e a B a) B B.inv (η a)) B
                         𝖺B[e a B a, d a, e a] B
                         (𝖺B-1[e a, a, d a] B e a) B
                         (((e a B a B d a) B (e a B 𝗅B-1[d a])) B B.inv (ε a) B e a)"
          using 0 1 unit1_def P_def emb.map0_def emb.map_def B.comp_assoc B.obj_def' by simp
        also have "... = (e a B 𝗋B[a]) B
                         𝖺B[e a, a, a] B
                         ((e a B a) B B.inv (η a)) B
                         𝖺B[e a B a, d a, e a] B
                         ((𝖺B-1[e a, a, d a] B e a) B ((e a B 𝗅B-1[d a]) B B.inv (ε a) B e a))"
        proof -
          have "(e a B a B d a) B (e a B 𝗅B-1[d a]) = e a B 𝗅B-1[d a]"
            using 0 B.comp_cod_arr by simp
          thus ?thesis by simp
        qed
        also have "... = (e a B 𝗋B[a]) B
                         𝖺B[e a, a, a] B
                         ((e a B a) B B.inv (η a)) B
                         (𝖺B[e a B a, d a, e a] B
                         ((𝗋B-1[e a] B d a) B e a)) B
                         (B.inv (ε a) B e a)"
        proof -
          have "(𝖺B-1[e a, a, d a] B e a) B ((e a B 𝗅B-1[d a]) B B.inv (ε a) B e a)
                  = (𝖺B-1[e a, a, d a] B e a) B
                    ((e a B 𝗅B-1[d a]) B e a) B (B.inv (ε a) B e a)"
            using 0 B.whisker_right by simp
          also have "... = ((𝖺B-1[e a, a, d a] B e a) B ((e a B 𝗅B-1[d a]) B e a)) B
                           (B.inv (ε a) B e a)"
            using B.comp_assoc by simp
          also have "... = (𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) B e a) B (B.inv (ε a) B e a)"
          proof -
            have "B.seq 𝖺B-1[e a, a, d a] (e a B 𝗅B-1[d a])"
              using a 0 by (elim B.objE, intro B.seqI) auto
            thus ?thesis
              using 0 B.whisker_right by simp
          qed
          also have "... = ((𝗋B-1[e a] B d a) B e a) B (B.inv (ε a) B e a)"
          proof -
            have "𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a]) = 𝗋B-1[e a] B d a"
            proof -
              have "e a B 𝗅B[d a] = (𝗋B[e a] B d a) B 𝖺B-1[e a, a, d a]"
                using 0 1 B.triangle' [of "e a" "d a"] by simp
              moreover have "B.hseq (e a) 𝗅B[d a]"
                using 0 by auto
              moreover have "srcB (e a) = a"
                using 0 by simp
              ultimately have "(𝗋B-1[e a] B d a) B (e a B 𝗅B[d a]) = 𝖺B-1[e a, a, d a]"
                using 0 B.invert_side_of_triangle(1)
                          [of "e a B 𝗅B[d a]" "𝗋B[e a] B d a" "𝖺B-1[e a, srcB (e a), d a]"]
                by simp
              moreover have "B.arr 𝖺B-1[e a, a, d a]"
                using 0 B.assoc'_in_hom [of "e a" a "d a"] by fastforce
              ultimately have "𝗋B-1[e a] B d a = 𝖺B-1[e a, a, d a] B (e a B 𝗅B-1[d a])"
                using 0 B.invert_side_of_triangle(2)
                          [of "𝖺B-1[e a, a, d a]" "𝗋B-1[e a] B d a" "e a B 𝗅B[d a]"]
                by simp
              thus ?thesis by simp
            qed
            thus ?thesis by simp
          qed
          finally have "(𝖺B-1[e a, a, d a] B e a) B ((e a B 𝗅B-1[d a]) B B.inv (ε a) B e a) =
                          ((𝗋B-1[e a] B d a) B e a) B (B.inv (ε a) B e a)"
            by simp
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = (e a B 𝗋B[a]) B
                         𝖺B[e a, a, a] B
                         (((e a B a) B B.inv (η a)) B
                         (𝗋B-1[e a] B d a B e a)) B
                         𝖺B[e a, d a, e a] B
                         (B.inv (ε a) B e a)"
          using 0 B.assoc_naturality [of "𝗋B-1[e a]" "d a" "e a"] B.comp_assoc by simp
        also have "... = (e a B 𝗋B[a]) B
                         𝖺B[e a, a, a] B
                         (𝗋B-1[e a] B a) B
                         ((e a B B.inv (η a)) B
                         𝖺B[e a, d a, e a] B
                         (B.inv (ε a) B e a))"
        proof -
          have "((e a B a) B B.inv (η a)) B (𝗋B-1[e a] B d a B e a)
                  = (e a B a) B 𝗋B-1[e a] B B.inv (η a) B (d a B e a)"
            using 0 B.interchange [of "e a B a" "𝗋B-1[e a]" "B.inv (η a)" "d a B e a"]
            by force
          also have "... = 𝗋B-1[e a] B e a B a B B.inv (η a)"
            using 0 B.comp_arr_dom B.comp_cod_arr by simp
          also have "... = (𝗋B-1[e a] B a) B (e a B B.inv (η a))"
            using 0 B.interchange [of "𝗋B-1[e a]" "e a" a "B.inv (η a)"] by fastforce
          ultimately have "((e a B a) B B.inv (η a)) B (𝗋B-1[e a] B d a B e a)
                             = (𝗋B-1[e a] B a) B (e a B B.inv (η a))"
            by simp
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = ((e a B 𝗋B[a]) B
                         𝖺B[e a, a, a] B
                         (𝗋B-1[e a] B a)) B
                         𝗋B-1[e a] B 𝗅B[e a]"
          by (metis B.comp_assoc adjoint_equivalence_in_bicategory_def
              adjunction_in_bicategory.triangle_right dual_adjoint_equivalence)
        also have "... = 𝗋B-1[e a] B 𝗅B[e a]"
        proof -
          have "(e a B 𝗋B[a]) B 𝖺B[e a, a, a] B (𝗋B-1[e a] B a) = e a B a"
          proof -
            have "(e a B 𝗋B[a]) B 𝖺B[e a, a, a] B (𝗋B-1[e a] B a)
                    = ((e a B 𝗋B[a]) B 𝖺B[e a, a, a]) B (𝗋B-1[e a] B a)"
              using B.comp_assoc by simp
            also have "... = (𝗋B[e a] B a) B (𝗋B-1[e a] B a)"
              using 0 B.runit_char(2) [of "e a"] B.unitor_coincidence by simp
            also have "... = e a B a"
              using 0 B.whisker_right [of a "𝗋B[e a]" "𝗋B-1[e a]"] B.comp_arr_inv' by auto
            finally show ?thesis
              by simp
          qed
          thus ?thesis
            using 0 B.comp_cod_arr by simp
        qed
        also have "... = (e a  IS.unit a)  𝗋-1[e a]  𝗅[e a]"
        proof -
          have "(e a  IS.unit a)  𝗋-1[e a]  𝗅[e a] = (e a B a) B 𝗋B-1[e a] B 𝗅B[e a]"
          proof -
            have 1: "arr 𝗅B[e a]  arr 𝗋B-1[e a]"
              using a lunit_simp runit'_simp lunit_simps runit'_simps by simp
            moreover have "arr (𝗋B-1[e a] B 𝗅B[e a])"
              using 1 arr_charSbC dom_charSbC cod_charSbC comp_char by auto
            moreover have 2: "hseq (e a) (IS.unit a)"
              using a IS.unit_char by (intro hseqI') auto
            moreover have "B.seq (e a  IS.unit a) (𝗋B-1[e a] B 𝗅B[e a])"
              using 0 2 hcomp_char arr_charSbC IS.unit_char'
              apply (intro B.seqI)
                  apply auto[4]
              using a by force
            ultimately show ?thesis
              using a comp_char IS.unit_char' lunit_simp runit'_simp hcomp_char
              by auto (* 4 sec *)
          qed
          also have "... = 𝗋B-1[e a] B 𝗅B[e a]"
            using 0 B.comp_cod_arr by simp
          finally show ?thesis by simp
        qed
        finally show "(e a  IS.unit a)  𝗋-1[e a]  𝗅[e a] = η1 a  (PoE.unit a  e a)"
          by argo
      qed
      show "f g. ide f; ide g; src g = trg f 
                     (e (trg g)  IS.cmp (g, f)) 
                     𝖺[e (trg g), IS.map g, IS.map f] 
                     (η1 g  IS.map f) 
                     inv 𝖺[PoE.map g, e (src g), IS.map f] 
                     (PoE.map g  η1 f) 
                     𝖺[PoE.map g, PoE.map f, e (src f)]
                       = η1 (g  f)  (PoE.cmp (g, f)  e (src f))"
      proof -
        fix f g
        assume f: "ide f" and g: "ide g" and fg: "src g = trg f"
        let ?a = "srcB f"
        let ?b = "srcB g"
        let ?c = "trgB g"
        have hseq: "ide f  ide g  arr f  arr g  src g = trg f"
          using f g fg ide_charSbC by auto
        have hseqB: "B.ide f  B.ide g  B.arr f  B.arr g  srcB g = trgB f"
          using f g fg ide_charSbC src_def trg_def by auto
        have 1: "?a = src f  ?b = src g  ?c = trg g"
          using f g fg src_def trg_def by simp
        have "(e (trg g)  IS.cmp (g, f)) 
              𝖺[e (trg g), IS.map g, IS.map f] 
              (η1 g  IS.map f) 
              inv 𝖺[PoE.map g, e (src g), IS.map f] 
              (PoE.map g  η1 f) 
              𝖺[PoE.map g, PoE.map f, e (src f)]
                = (e ?c  IS.cmp (g, f)) 
                  𝖺[e ?c, IS.map g, IS.map f] 
                  (η1 g  IS.map f) 
                  inv 𝖺[PoE.map g, e ?b, IS.map f] 
                  (PoE.map g  η1 f) 
                  𝖺[PoE.map g, PoE.map f, e ?a]"
          using 1 by simp
        also have "... = (e ?c B 𝗋B[g] B (g B B.inv (η ?b))
                            B 𝗋B[f] B (f B B.inv (η ?a))) B
                         𝖺B[e ?c, g B d ?b B e ?b, f B d ?a B e ?a] B
                         (𝖺B[e ?c, g, d ?b B e ?b] B f B d ?a B e ?a) B
                         (𝖺B[e ?c B g, d ?b, e ?b] B f B d ?a B e ?a) B
                         ((𝖺B-1[e ?c, g, d ?b] B e ?b) B (f B d ?a B e ?a)) B
                         𝖺B-1[P g, e ?b, f B d ?a B e ?a] B
                         (P g B 𝖺B[e ?b, f, d ?a B e ?a]) B
                         (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                         (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                         𝖺B[P g, P f, e ?a]"
        proof -
          have "(e ?c  IS.cmp (g, f)) 
                𝖺[e ?c, IS.map g, IS.map f] 
                (η1 g  IS.map f) 
                inv 𝖺[PoE.map g, e ?b, IS.map f] 
                (PoE.map g  unit1 f) 
                𝖺[PoE.map g, PoE.map f, e ?a]
                  = (e ?c  g  f) 
                    𝖺[e ?c, g, f] 
                    (η1 g  f) 
                    inv 𝖺[P (E g), e ?b, f] 
                    (P (E g)  η1 f) 
                    𝖺[P (E g), P (E f), e ?a]"
            using f g fg by simp
          also have "... = (e ?c  g  f) B
                           𝖺[e ?c, g, f] B
                           (η1 g  f) B
                           inv 𝖺[P (E g), e ?b, f] B
                           (P (E g)  η1 f) B
                           𝖺[P (E g), P (E f), e ?a]"
          proof -
            have "arr ..."
            proof -
              have "«... : (P (E g)  P (E f))  e ?a  e ?c  g  f»"
              proof (intro comp_in_homI)
                show "«𝖺[P (E g), P (E f), e ?a] :
                         (P (E g)  P (E f))  e ?a  P (E g)  P (E f)  e ?a»"
                  using f g fg VVV.arr_charSbC VV.arr_charSbC src_def obj_src assoc_in_hom
                  by force
                show "«P (E g)  η1 f : P (E g)  P (E f)  e ?a  P (E g)  e ?b  f»"
                  using f g fg 1 unit1_in_homS [of f] vconn_implies_hpar(2)
                  by (intro hcomp_in_vhom) auto
                show "«inv 𝖺[P (E g), e ?b, f] : P (E g)  e ?b  f  (P (E g)  e ?b)  f»"
                proof -
                  have "«𝖺[P (E g), e ?b, f] : (P (E g)  e ?b)  f  P (E g)  e ?b  f»"
                    using f g fg 1 VVV.arr_charSbC VV.arr_charSbC assoc_in_hom by simp
                  moreover have "iso 𝖺[P (E g), e ?b, f]"
                    using f g fg 1 VVV.arr_charSbC VV.arr_charSbC iso_assoc by simp
                  ultimately show ?thesis
                    using inv_in_hom by simp
                qed
                show "«unit1 g  f : (P (E g)  e ?b)  f  (e ?c  g)  f»"
                  using f g fg 1 unit1_in_homS [of g]
                  by (intro hcomp_in_vhom) auto
                have 2: "ide (e (trg g))"
                  using g by simp
                have 3: "src (e (trg g)) = trg g"
                  using g trg_def hseqE in_hom_charSbC obj_trg by auto
                show "«𝖺[e ?c, g, f] : (e ?c  g)  f  e ?c  g  f»"
                  using f g fg 1 2 3 VVV.arr_charSbC VV.arr_charSbC assoc_in_hom [of "e ?c" g f]
                  by simp
                show "«e ?c  g  f : e ?c  g  f  e ?c  g  f»"
                  using f g fg 1 2 3
                  by (intro hcomp_in_vhom) auto
              qed
              thus ?thesis by blast
            qed
            thus ?thesis
              using comp_eqI by blast
          qed
          also have "... = (e ?c  g  f) B
                           𝖺[e ?c, g, f] B
                           (η1 g  f) B
                           B.inv 𝖺[P (E g), e ?b, f] B
                           (P (E g)  η1 f) B
                           𝖺[P (E g), P (E f), e ?a]"
          proof -
            (*
             * TODO: Maybe assoc should be a definition in subcategory, rather than
             * an abbreviation.  The expansion seems to create problems here.
             *)
            have "iso 𝖺[P (E g), e ?b, f]"
              using f g fg hseq hseqB src_def trg_def emb.map0_def emb.map_def arr_charSbC
                    obj_char P0_props(6)
              by (intro iso_assoc) auto
            thus ?thesis
              using f g fg inv_charSbC [of "𝖺[P (E g), e ?b, f]"] by simp
          qed
          also have "... = (e ?c B g B f) B
                           𝖺[e ?c, g, f] B
                           (η1 g B f) B
                           B.inv 𝖺[P (E g), e ?b, f] B
                           (P (E g) B η1 f) B
                           𝖺[P (E g), P (E f), e ?a]"
          proof -
            have "arr (unit0 ?c)"
              using hseq obj_trg trg_def by auto
            moreover have "arr (g B f)"
              using hseq hseqB arr_charSbC by auto
            ultimately show ?thesis
              unfolding hcomp_def
              using f g fg hseqB src_def trg_def P0_props(1) emb.map0_def emb.map_def
              by simp
          qed
          also have "... = (e ?c B g B f) B
                           𝖺B[e ?c, g, f] B
                           (η1 g B f) B
                           B.inv 𝖺B[P (E g), e ?b, f] B
                           (P (E g) B η1 f) B
                           𝖺B[P (E g), P (E f), e ?a]"
          proof -
            have "arr (e ?a)  arr (e ?b)  arr (e ?c)"
              using hseq hseqB src_def trg_def obj_char src.preserves_arr trg.preserves_arr
              by auto
            thus ?thesis
              using f g fg VVV.arr_charSbC VV.arr_charSbC src_def trg_def hseqB
                    emb.map0_def emb.map_def
              by simp
          qed
          also have "... = (e ?c B g B f) B
                           𝖺B[e ?c, g, f] B
                           (η1 g B f) B
                           (B.inv 𝖺B[P g, e ?b, f] B
                           (P g B η1 f) B
                           𝖺B[P g, P f, e ?a])"
            using f g fg 1 emb.map_def by simp
          also have "... = (e ?c B g B f) B
                           𝖺B[e ?c, g, f] B
                           (η1 g B f) B
                           (𝖺B-1[P g, e ?b, f] B
                           (P g B η1 f) B
                           𝖺B[P g, P f, e ?a])"
            using hseqB P.preserves_ide ide_charSbC B.assoc'_eq_inv_assoc [of "P g" "e ?b" f]
            by simp
          also have "... = (e ?c B g B f) B
                           𝖺B[e ?c, g, f] B
                           ((e ?c B 𝗋B[g]) B
                            𝖺B[e ?c, g, ?b] B
                            ((e ?c B g) B B.inv (η ?b)) B
                            𝖺B[e ?c B g, d ?b, e ?b] B
                            (𝖺B-1[e ?c, g, d ?b] B e ?b)
                               B f) B
                           𝖺B-1[P g, e ?b, f] B
                           (P g B (e ?b B 𝗋B[f]) B
                                   𝖺B[e ?b, f, ?a] B
                                   ((e ?b B f) B B.inv (η ?a)) B
                                   𝖺B[e ?b B f, d ?a, e ?a] B
                                   (𝖺B-1[e ?b, f, d ?a] B e ?a)) B
                           𝖺B[P g, P f, e ?a]"
            unfolding unit1_def
            using hseqB B.comp_assoc by auto
          also have "... = ((e ?c B g B f) B
                           𝖺B[e ?c, g, f]) B
                           ((e ?c B 𝗋B[g]) B f) B
                           (𝖺B[e ?c, g, ?b] B f) B
                           (((e ?c B g) B B.inv (η ?b)) B f) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B f) B
                           𝖺B-1[P g, e ?b, f] B
                           (P g B e ?b B 𝗋B[f]) B
                           (P g B 𝖺B[e ?b, f, ?a]) B
                           (P g B (e ?b B f) B B.inv (η ?a)) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
            using hseqB ide_charSbC P.preserves_ide B.whisker_left B.whisker_right B.comp_assoc
            by auto
          also have "... = 𝖺B[e ?c, g, f] B
                           ((e ?c B 𝗋B[g]) B f) B
                           ((𝖺B[e ?c, g, ?b] B f) B
                           (((e ?c B g) B B.inv (η ?b)) B f)) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B f) B
                           𝖺B-1[P g, e ?b, f] B
                           (P g B e ?b B 𝗋B[f]) B
                           (P g B 𝖺B[e ?b, f, ?a]) B
                           (P g B (e ?b B f) B B.inv (η ?a)) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
          proof -
            have "(e ?c B g B f) B 𝖺B[e ?c, g, f] = 𝖺B[e ?c, g, f]"
              using hseqB B.comp_cod_arr by auto
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = 𝖺B[e ?c, g, f] B
                           (((e ?c B 𝗋B[g]) B f) B
                           ((e ?c B g B B.inv (η ?b)) B f)) B
                           (𝖺B[e ?c, g, d ?b B e ?b] B f) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B f) B
                           𝖺B-1[P g, e ?b, f] B
                           (P g B e ?b B 𝗋B[f]) B
                           (P g B 𝖺B[e ?b, f, ?a]) B
                           (P g B (e ?b B f) B B.inv (η ?a)) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
          proof -
            have "(𝖺B[e ?c, g, ?b] B f) B (((e ?c B g) B B.inv (η ?b)) B f)
                    = ((e ?c B g B B.inv (η ?b)) B f) B (𝖺B[e ?c, g, d ?b B e ?b] B f)"
            proof -
              have "(𝖺B[e ?c, g, ?b] B f) B (((e ?c B g) B B.inv (η ?b)) B f)
                      = 𝖺B[e ?c, g, ?b] B ((e ?c B g) B B.inv (η ?b)) B f"
                using hseqB B.whisker_right by auto
              also have "... = (e ?c B g B B.inv (η ?b)) B 𝖺B[e ?c, g, d ?b B e ?b] B f"
                using hseqB B.assoc_naturality [of "e ?c" g "B.inv (η ?b)"] by simp
              also have "... = ((e ?c B g B B.inv (η ?b)) B f) B
                                (𝖺B[e ?c, g, d ?b B e ?b] B f)"
                using hseqB B.whisker_right by auto
              finally show ?thesis by simp
            qed
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (𝖺B[e ?c, g, f] B
                           ((e ?c B 𝗋B[g] B (g B B.inv (η ?b))) B f)) B
                           (𝖺B[e ?c, g, d ?b B e ?b] B f) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B f) B
                           𝖺B-1[P g, e ?b, f] B
                           (P g B e ?b B 𝗋B[f]) B
                           (P g B 𝖺B[e ?b, f, ?a]) B
                           (P g B (e ?b B f) B B.inv (η ?a)) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
            using hseqB B.whisker_left B.whisker_right B.comp_assoc by simp
          also have "... = (e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B f) B
                           𝖺B[e ?c, g B d ?b B e ?b, f] B
                           (𝖺B[e ?c, g, d ?b B e ?b] B f) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B f) B
                           𝖺B-1[P g, e ?b, f] B
                           (P g B e ?b B 𝗋B[f]) B
                           ((P g B 𝖺B[e ?b, f, ?a]) B
                           (P g B (e ?b B f) B B.inv (η ?a))) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
            using hseqB B.comp_assoc
                  B.assoc_naturality [of "e ?c" "𝗋B[g] B (g B B.inv (η ?b))" f]
            by fastforce
          also have "... = (e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B f) B
                           𝖺B[e ?c, g B d ?b B e ?b, f] B
                           (𝖺B[e ?c, g, d ?b B e ?b] B f) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B f) B
                           𝖺B-1[P g, e ?b, f] B
                           ((P g B e ?b B 𝗋B[f]) B
                           (P g B e ?b B f B B.inv (η ?a))) B
                           (P g B 𝖺B[e ?b, f, d ?a B e ?a]) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
          proof -
            have "(P g B 𝖺B[e ?b, f, ?a]) B (P g B (e ?b B f) B B.inv (η ?a))
                    = (P g B e ?b B f B B.inv (η ?a)) B (P g B 𝖺B[e ?b, f, d ?a B e ?a])"
            proof -
              have "(P g B 𝖺B[e ?b, f, ?a]) B (P g B (e ?b B f) B B.inv (η ?a))
                      = P g B 𝖺B[e ?b, f, ?a] B ((e ?b B f) B B.inv (η ?a))"
                using hseqB ide_charSbC P.preserves_ide B.whisker_left by auto
              also have "... = P g B (e ?b B f B B.inv (η ?a)) B 𝖺B[e ?b, f, d ?a B e ?a]"
                using hseqB B.assoc_naturality [of "e ?b" f "B.inv (η ?a)"] by auto
              also have "... = (P g B e ?b B f B B.inv (η ?a)) B
                               (P g B 𝖺B[e ?b, f, d ?a B e ?a])"
                using hseqB ide_charSbC P.preserves_ide B.whisker_left by auto
              finally show ?thesis by simp
            qed
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B f) B
                           𝖺B[e ?c, g B d ?b B e ?b, f] B
                           (𝖺B[e ?c, g, d ?b B e ?b] B f) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B f) B
                           (𝖺B-1[P g, e ?b, f] B
                           (P g B e ?b B 𝗋B[f] B (f B B.inv (η ?a)))) B
                           (P g B 𝖺B[e ?b, f, d ?a B e ?a]) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
            using hseqB ide_charSbC P.preserves_ide B.whisker_left B.comp_assoc by auto
          also have "... = (e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B f) B
                           𝖺B[e ?c, g B d ?b B e ?b, f] B
                           (𝖺B[e ?c, g, d ?b B e ?b] B f) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f) B
                           (((𝖺B-1[e ?c, g, d ?b] B e ?b) B f) B
                           ((P g B e ?b) B 𝗋B[f] B (f B B.inv (η ?a)))) B
                           𝖺B-1[P g, e ?b, f B d ?a B e ?a] B
                           (P g B 𝖺B[e ?b, f, d ?a B e ?a]) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
            using hseqB B.comp_assoc
                  B.assoc'_naturality [of "P g" "e ?b" "𝗋B[f] B (f B B.inv (η ?a))"]
            by simp
          also have "... = (e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B f) B
                           𝖺B[e ?c, g B d ?b B e ?b, f] B
                           (𝖺B[e ?c, g, d ?b B e ?b] B f) B
                           ((𝖺B[e ?c B g, d ?b, e ?b] B f) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B 𝗋B[f] B (f B B.inv (η ?a)))) B
                           𝖺B-1[P g, e ?b, f B d ?a B e ?a] B
                           (P g B 𝖺B[e ?b, f, d ?a B e ?a]) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
          proof -
            have "((𝖺B-1[e ?c, g, d ?b] B e ?b) B f) B
                     ((P g B e ?b) B 𝗋B[f] B (f B B.inv (η ?a)))
                    = (𝖺B-1[e ?c, g, d ?b] B e ?b) B 𝗋B[f] B (f B B.inv (η ?a))"
            proof -
              have "((𝖺B-1[e ?c, g, d ?b] B e ?b) B f) B
                    ((P g B e ?b) B 𝗋B[f] B (f B B.inv (η ?a)))
                      = (𝖺B-1[e ?c, g, d ?b] B e ?b) B (P g B e ?b) B
                        (f B 𝗋B[f] B (f B B.inv (η ?a)))"
              proof -
                have "B.seq (𝖺B-1[e ?c, g, d ?b] B e ?b) (P g B e ?b)"
                  using hseqB P.preserves_ide P_def src_def trg_def by auto
                moreover have "B.seq f (𝗋B[f] B (f B B.inv (η ?a)))"
                  using hseqB by simp
                ultimately show ?thesis
                  using B.interchange
                          [of "𝖺B-1[e ?c, g, d ?b] B e ?b" "P g B e ?b"
                              f "𝗋B[f] B (f B B.inv (η ?a))"]
                  by presburger
              qed
              also have "... = (𝖺B-1[e ?c, g, d ?b] B (e ?c B g B d ?b) B e ?b)
                                  B f B 𝗋B[f] B (f B B.inv (η ?a))"
                using hseqB B.whisker_right P_def by auto
              also have "... = (𝖺B-1[e ?c, g, d ?b] B e ?b) B 𝗋B[f] B (f B B.inv (η ?a))"
                using hseqB B.comp_arr_dom B.comp_cod_arr by simp
              finally show ?thesis by simp
            qed
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B f) B
                           𝖺B[e ?c, g B d ?b B e ?b, f] B
                           ((𝖺B[e ?c, g, d ?b B e ?b] B f) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B 𝗋B[f] B (f B B.inv (η ?a)))) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B (f B d ?a B e ?a)) B
                           𝖺B-1[P g, e ?b, f B d ?a B e ?a] B
                           (P g B 𝖺B[e ?b, f, d ?a B e ?a]) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
          proof -
            have "(𝖺B[e ?c B g, d ?b, e ?b] B f) B ((𝖺B-1[e ?c, g, d ?b] B e ?b)
                     B 𝗋B[f] B (f B B.inv (η ?a)))
                    = (𝖺B[e ?c B g, d ?b, e ?b] B 𝗋B[f] B (f B B.inv (η ?a))) B
                        ((𝖺B-1[e ?c, g, d ?b] B e ?b) B (f B d ?a B e ?a))"
            proof -
              have "(𝖺B[e ?c B g, d ?b, e ?b] B f) B
                    ((𝖺B-1[e ?c, g, d ?b] B e ?b) B 𝗋B[f] B (f B B.inv (η ?a)))
                      = 𝖺B[e ?c B g, d ?b, e ?b] B (𝖺B-1[e ?c, g, d ?b] B e ?b)
                          B f B 𝗋B[f] B (f B B.inv (η ?a))"
                using hseqB B.interchange by simp
              also have "... = 𝖺B[e ?c B g, d ?b, e ?b] B (𝖺B-1[e ?c, g, d ?b] B e ?b)
                                  B (𝗋B[f] B (f B B.inv (η ?a))) B (f B d ?a B e ?a)"
                using hseqB B.comp_cod_arr B.comp_arr_dom by simp
              also have "... = (𝖺B[e ?c B g, d ?b, e ?b] B 𝗋B[f] B (f B B.inv (η ?a))) B
                               ((𝖺B-1[e ?c, g, d ?b] B e ?b) B (f B d ?a B e ?a))"
                using hseqB B.interchange by auto
              finally show ?thesis by simp
            qed
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B f) B
                           (𝖺B[e ?c, g B d ?b B e ?b, f] B
                           (𝖺B[e ?c, g, d ?b B e ?b] B 𝗋B[f] B (f B B.inv (η ?a)))) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f B d ?a B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B (f B d ?a B e ?a)) B
                           𝖺B-1[P g, e ?b, f B d ?a B e ?a] B
                           (P g B 𝖺B[e ?b, f, d ?a B e ?a]) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
          proof -
            have "(𝖺B[e ?c, g, d ?b B e ?b] B f) B
                  (𝖺B[e ?c B g, d ?b, e ?b] B 𝗋B[f] B (f B B.inv (η ?a)))
                    = 𝖺B[e ?c, g, d ?b B e ?b] B 𝖺B[e ?c B g, d ?b, e ?b]
                        B f B 𝗋B[f] B (f B B.inv (η ?a))"
              using hseqB B.interchange by simp
            also have "... = 𝖺B[e ?c, g, d ?b B e ?b] B 𝖺B[e ?c B g, d ?b, e ?b]
                                B (𝗋B[f] B (f B B.inv (η ?a))) B (f B d ?a B e ?a)"
              using hseqB B.comp_cod_arr B.comp_arr_dom by simp
            also have "... = (𝖺B[e ?c, g, d ?b B e ?b] B 𝗋B[f] B (f B B.inv (η ?a))) B
                             (𝖺B[e ?c B g, d ?b, e ?b] B f B d ?a B e ?a)"
              using hseqB B.interchange by auto
            finally have "(𝖺B[e ?c, g, d ?b B e ?b] B f) B (𝖺B[e ?c B g, d ?b, e ?b]
                              B 𝗋B[f] B (f B B.inv (η ?a)))
                            = (𝖺B[e ?c, g, d ?b B e ?b] B 𝗋B[f] B (f B B.inv (η ?a))) B
                              (𝖺B[e ?c B g, d ?b, e ?b] B f B d ?a B e ?a)"
              by simp
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = ((e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B f) B
                           ((e ?c B (g B d ?b B e ?b) B 𝗋B[f] B (f B B.inv (η ?a))))) B
                           𝖺B[e ?c, g B d ?b B e ?b, f B d ?a B e ?a] B
                           (𝖺B[e ?c, g, d ?b B e ?b] B f B d ?a B e ?a) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f B d ?a B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B (f B d ?a B e ?a)) B
                           𝖺B-1[P g, e ?b, f B d ?a B e ?a] B
                           (P g B 𝖺B[e ?b, f, d ?a B e ?a]) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
          proof -
            have "𝖺B[e ?c, g B d ?b B e ?b, f] B
                  (𝖺B[e ?c, g, d ?b B e ?b] B 𝗋B[f] B (f B B.inv (η ?a)))
                    = 𝖺B[e ?c, g B d ?b B e ?b, f] B ((e ?c B g B d ?b B e ?b) B
                      𝖺B[e ?c, g, d ?b B e ?b]
                         B (𝗋B[f] B (f B B.inv (η ?a))) B (f B d ?a B e ?a))"
              using hseqB B.comp_arr_dom B.comp_cod_arr by simp
            also have "... = 𝖺B[e ?c, g B d ?b B e ?b, f] B
                             ((e ?c B g B d ?b B e ?b) B 𝗋B[f] B (f B B.inv (η ?a))) B
                             (𝖺B[e ?c, g, d ?b B e ?b] B f B d ?a B e ?a)"
              using hseqB B.interchange by simp
            also have "... = (𝖺B[e ?c, g B d ?b B e ?b, f] B
                             ((e ?c B g B d ?b B e ?b) B 𝗋B[f] B (f B B.inv (η ?a)))) B
                             (𝖺B[e ?c, g, d ?b B e ?b] B f B d ?a B e ?a)"
              using B.comp_assoc by simp
            also have "... = (((e ?c B (g B d ?b B e ?b) B 𝗋B[f] B (f B B.inv (η ?a)))) B
                             𝖺B[e ?c, g B d ?b B e ?b, f B d ?a B e ?a]) B
                             (𝖺B[e ?c, g, d ?b B e ?b] B f B d ?a B e ?a)"
              using hseqB B.assoc_naturality [of "e ?c" "g B d ?b B e ?b"
                                                 "𝗋B[f] B (f B B.inv (η ?a))"]
              by auto
            also have "... = ((e ?c B (g B d ?b B e ?b) B 𝗋B[f] B (f B B.inv (η ?a)))) B
                             𝖺B[e ?c, g B d ?b B e ?b, f B d ?a B e ?a] B
                             (𝖺B[e ?c, g, d ?b B e ?b] B f B d ?a B e ?a)"
              using B.comp_assoc by simp
            finally have "𝖺B[e ?c, g B d ?b B e ?b, f] B (𝖺B[e ?c, g, d ?b B e ?b]
                              B 𝗋B[f] B (f B B.inv (η ?a)))
                            = ((e ?c B (g B d ?b B e ?b) B 𝗋B[f] B (f B B.inv (η ?a)))) B
                              𝖺B[e ?c, g B d ?b B e ?b, f B d ?a B e ?a] B
                              (𝖺B[e ?c, g, d ?b B e ?b] B f B d ?a B e ?a)"
              by simp
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (e ?c B 𝗋B[g] B (g B B.inv (η ?b))
                              B 𝗋B[f] B (f B B.inv (η ?a))) B
                           𝖺B[e ?c, g B d ?b B e ?b, f B d ?a B e ?a] B
                           (𝖺B[e ?c, g, d ?b B e ?b] B f B d ?a B e ?a) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f B d ?a B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B (f B d ?a B e ?a)) B
                           𝖺B-1[P g, e ?b, f B d ?a B e ?a] B
                           (P g B 𝖺B[e ?b, f, d ?a B e ?a]) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
          proof -
            have "(e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B f) B
                  ((e ?c B (g B d ?b B e ?b) B 𝗋B[f] B (f B B.inv (η ?a))))
                    = e ?c B (𝗋B[g] B (g B B.inv (η ?b))) B (g B d ?b B e ?b)
                         B f B 𝗋B[f] B (f B B.inv (η ?a))"
              using hseqB B.whisker_left B.interchange by fastforce
            also have "... = e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B 𝗋B[f] B (f B B.inv (η ?a))"
              using hseqB B.comp_arr_dom B.comp_cod_arr by auto
            finally have "(e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B f) B
                          ((e ?c B (g B d ?b B e ?b) B 𝗋B[f] B (f B B.inv (η ?a))))
                            = e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B 𝗋B[f] B (f B B.inv (η ?a))"
              by simp
            thus ?thesis by simp
          qed
          finally show ?thesis by blast
        qed
        also have "... = (e ?c B 𝗋B[g B f]) B
                         𝖺B[e ?c, g B f, ?a] B
                         ((e ?c B g B f) B B.inv (η ?a)) B
                         𝖺B[e ?c B g B f, d ?a, e ?a] B
                         (𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                         ((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                         ((e ?c B g B 𝗅B[f B d ?a]) B e ?a) B
                         ((e ?c B g B B.inv (η ?b) B f B d ?a) B e ?a) B
                         ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                         (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                         (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                         ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
        proof -
          (* Working from the above expression to the common form, as in the previous part. *)
          have "(e ?c B 𝗋B[g B f]) B
                𝖺B[e ?c, g B f, ?a] B
                ((e ?c B g B f) B B.inv (η ?a)) B
                𝖺B[e ?c B g B f, d ?a, e ?a] B
                (𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                ((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                ((e ?c B g B 𝗅B[f B d ?a]) B e ?a) B
                ((e ?c B g B B.inv (η ?b) B f B d ?a) B e ?a) B
                ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)
                  = (e ?c B 𝗋B[g B f]) B
                    (𝖺B[e ?c, g B f, ?a] B
                    ((e ?c B g B f) B B.inv (η ?a))) B
                    𝖺B[e ?c B g B f, d ?a, e ?a] B
                    (𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                    ((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                    ((e ?c B g B 𝗅B[f B d ?a]) B e ?a) B
                    ((e ?c B g B B.inv (η ?b) B f B d ?a) B e ?a) B
                    ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                    (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                    (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                    ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
            using B.comp_assoc by simp
          also have "... = ((e ?c B 𝗋B[g B f]) B
                           (e ?c B (g B f) B B.inv (η ?a))) B
                           𝖺B[e ?c, g B f, d ?a B e ?a] B
                           𝖺B[e ?c B g B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                           ((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                           ((e ?c B g B 𝗅B[f B d ?a]) B e ?a) B
                           ((e ?c B g B B.inv (η ?b) B f B d ?a) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
            using hseqB B.assoc_naturality [of "e ?c" "g B f" "B.inv (η ?a)"] B.comp_assoc
            by simp
          also have "... = (e ?c B 𝗋B[g B f] B ((g B f) B B.inv (η ?a))) B
                           𝖺B[e ?c, g B f, d ?a B e ?a] B
                           𝖺B[e ?c B g B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                           (((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                           ((e ?c B g B 𝗅B[f B d ?a]) B e ?a)) B
                           ((e ?c B g B B.inv (η ?b) B f B d ?a) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
            using hseqB B.whisker_left B.comp_assoc by auto
          also have "... = (e ?c B 𝗋B[g B f] B ((g B f) B B.inv (η ?a))) B
                           𝖺B[e ?c, g B f, d ?a B e ?a] B
                           𝖺B[e ?c B g B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                           ((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                           ((e ?c B 𝗋B[g] B f B d ?a) B e ?a) B
                           (((e ?c B 𝖺B-1[g, ?b, f B d ?a]) B e ?a) B
                           ((e ?c B g B B.inv (η ?b) B f B d ?a) B e ?a)) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
          proof -
            have "((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                    ((e ?c B g B 𝗅B[f B d ?a]) B e ?a)
                    = ((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                        ((e ?c B 𝗋B[g] B f B d ?a) B e ?a) B
                          ((e ?c B 𝖺B-1[g, ?b, f B d ?a]) B e ?a)"
              using hseqB B.whisker_right B.whisker_left B.triangle' comp_assoc by auto
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (e ?c B 𝗋B[g B f] B ((g B f) B B.inv (η ?a))) B
                           𝖺B[e ?c, g B f, d ?a B e ?a] B
                           𝖺B[e ?c B g B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                           ((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                           (((e ?c B 𝗋B[g] B f B d ?a) B e ?a) B
                           ((e ?c B (((g B B.inv (η ?b)) B f B d ?a))) B e ?a)) B
                           (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
          proof -
            have "((e ?c B 𝖺B-1[g, ?b, f B d ?a]) B e ?a) B
                  ((e ?c B g B B.inv (η ?b) B f B d ?a) B e ?a)
                    = (e ?c B 𝖺B-1[g, ?b, f B d ?a] B (g B B.inv (η ?b) B f B d ?a)) B e ?a"
              using hseqB B.whisker_right B.whisker_left by auto
            also have "... = (e ?c B
                               (((g B B.inv (η ?b)) B f B d ?a)) B
                               𝖺B-1[g, d ?b B e ?b, f B d ?a])
                                 B e ?a"
              using hseqB B.assoc'_naturality [of g "B.inv (η ?b)" "f B d ?a"] by auto
            also have "... = ((e ?c B (((g B B.inv (η ?b)) B f B d ?a))) B e ?a) B
                             (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a)"
              using hseqB B.whisker_right B.whisker_left by auto
            finally have "((e ?c B 𝖺B-1[g, ?b, f B d ?a]) B e ?a) B
                          ((e ?c B g B B.inv (η ?b) B f B d ?a) B e ?a)
                            = ((e ?c B (((g B B.inv (η ?b)) B f B d ?a))) B e ?a) B
                              (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a)"
              by simp
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (e ?c B 𝗋B[g B f] B ((g B f) B B.inv (η ?a))) B
                           𝖺B[e ?c, g B f, d ?a B e ?a] B
                           𝖺B[e ?c B g B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                           (((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                           ((e ?c B (𝗋B[g] B (g B B.inv (η ?b))) B f B d ?a) B e ?a)) B
                           (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
            using hseqB B.whisker_right B.whisker_left B.comp_assoc by auto
          also have "... = (e ?c B 𝗋B[g B f] B ((g B f) B B.inv (η ?a))) B
                           𝖺B[e ?c, g B f, d ?a B e ?a] B
                           𝖺B[e ?c B g B f, d ?a, e ?a] B
                           ((𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                           ((e ?c B (((𝗋B[g] B (g B B.inv (η ?b))) B f) B d ?a)) B e ?a)) B
                           ((e ?c B 𝖺B-1[g B d ?b B e ?b, f, d ?a]) B e ?a) B
                           (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
          proof -
            have "((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                  ((e ?c B (𝗋B[g] B (g B B.inv (η ?b))) B f B d ?a) B e ?a)
                    = (e ?c B 𝖺B-1[g, f, d ?a] B
                      ((𝗋B[g] B (g B B.inv (η ?b))) B f B d ?a)) B e ?a"
              using hseqB B.whisker_left B.whisker_right by auto
            also have "... = (e ?c B
                               (((𝗋B[g] B (g B B.inv (η ?b))) B f) B d ?a) B
                                  𝖺B-1[g B d ?b B e ?b, f, d ?a])
                                 B e ?a"
              using hseqB B.assoc'_naturality [of "𝗋B[g] B (g B B.inv (η ?b))" f "d ?a"]
              by auto
            also have "... = ((e ?c B (((𝗋B[g] B (g B B.inv (η ?b))) B f) B d ?a))
                                 B e ?a) B
                             ((e ?c B 𝖺B-1[g B d ?b B e ?b, f, d ?a]) B e ?a)"
              using hseqB B.whisker_left B.whisker_right by auto
            finally have "((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                          ((e ?c B (𝗋B[g] B (g B B.inv (η ?b))) B f B d ?a) B e ?a)
                            = ((e ?c B (((𝗋B[g] B (g B B.inv (η ?b))) B f) B d ?a)) B e ?a) B
                              ((e ?c B 𝖺B-1[g B d ?b B e ?b, f, d ?a]) B e ?a)"
              by simp
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (e ?c B 𝗋B[g B f] B ((g B f) B B.inv (η ?a))) B
                           𝖺B[e ?c, g B f, d ?a B e ?a] B
                           (𝖺B[e ?c B g B f, d ?a, e ?a] B
                           (((e ?c B ((𝗋B[g] B (g B B.inv (η ?b))) B f)) B d ?a) B e ?a)) B
                           (𝖺B-1[e ?c, (g B (d ?b B e ?b)) B f, d ?a] B e ?a) B
                           ((e ?c B 𝖺B-1[g B d ?b B e ?b, f, d ?a]) B e ?a) B
                           (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
          proof -
            have "(𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                  ((e ?c B (((𝗋B[g] B (g B B.inv (η ?b))) B f) B d ?a)) B e ?a)
                    = 𝖺B-1[e ?c, g B f, d ?a] B
                      (e ?c B ((𝗋B[g] B (g B B.inv (η ?b)) B f) B d ?a)) B e ?a"
              using hseqB B.whisker_right by auto
            also have "... = ((e ?c B ((𝗋B[g] B (g B B.inv (η ?b))) B f)) B d ?a) B
                             𝖺B-1[e ?c, (g B (d ?b B e ?b)) B f, d ?a] B e ?a"
              using hseqB B.assoc'_naturality [of "e ?c" "(𝗋B[g] B (g B B.inv (η ?b))) B f" "d ?a"]
              by fastforce
            also have "... = (((e ?c B ((𝗋B[g] B (g B B.inv (η ?b))) B f)) B d ?a)
                                B e ?a) B
                             (𝖺B-1[e ?c, (g B (d ?b B e ?b)) B f, d ?a] B e ?a)"
              using hseqB B.whisker_right by auto
            finally have "(𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                          ((e ?c B (((𝗋B[g] B (g B B.inv (η ?b))) B f) B d ?a)) B e ?a)
                            = (((e ?c B ((𝗋B[g] B (g B B.inv (η ?b))) B f)) B d ?a) B e ?a) B
                              (𝖺B-1[e ?c, (g B (d ?b B e ?b)) B f, d ?a] B e ?a)"
              by simp
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (e ?c B 𝗋B[g B f] B ((g B f) B B.inv (η ?a))) B
                           (𝖺B[e ?c, g B f, d ?a B e ?a] B
                           ((e ?c B ((𝗋B[g] B (g B B.inv (η ?b))) B f)) B d ?a B e ?a)) B
                           𝖺B[e ?c B (g B d ?b B e ?b) B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, (g B (d ?b B e ?b)) B f, d ?a] B e ?a) B
                           ((e ?c B 𝖺B-1[g B d ?b B e ?b, f, d ?a]) B e ?a) B
                           (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
          proof -
            have "𝖺B[e ?c B g B f, d ?a, e ?a] B
                  (((e ?c B ((𝗋B[g] B (g B B.inv (η ?b))) B f)) B d ?a) B e ?a)
                    = ((e ?c B ((𝗋B[g] B (g B B.inv (η ?b))) B f)) B d ?a B e ?a) B
                      𝖺B[e ?c B (g B d ?b B e ?b) B f, d ?a, e ?a]"
              using hseqB
                    B.assoc_naturality [of "(e ?c B ((𝗋B[g] B (g B B.inv (η ?b))) B f))"
                                           "d ?a" "e ?a"]
              by fastforce
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (e ?c B 𝗋B[g B f] B ((g B f) B B.inv (η ?a))) B
                           (e ?c B ((𝗋B[g] B (g B B.inv (η ?b))) B f) B d ?a B e ?a) B
                           𝖺B[e ?c, (g B d ?b B e ?b) B f, d ?a B e ?a] B
                           𝖺B[e ?c B (g B d ?b B e ?b) B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, (g B (d ?b B e ?b)) B f, d ?a] B e ?a) B
                           ((e ?c B 𝖺B-1[g B d ?b B e ?b, f, d ?a]) B e ?a) B
                           (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
            using hseqB B.comp_assoc
                  B.assoc_naturality [of "e ?c" "(𝗋B[g] B (g B B.inv (η ?b))) B f"
                                         "d ?a B e ?a"]
            by force
          also have "... = (e ?c B (g B 𝗋B[f] B (f B B.inv (η ?a)))) B
                           ((e ?c B 𝖺B[g, f, d ?a B e ?a]) B
                           (e ?c B ((𝗋B[g] B (g B B.inv (η ?b))) B f) B d ?a B e ?a)) B
                           𝖺B[e ?c, (g B d ?b B e ?b) B f, d ?a B e ?a] B
                           𝖺B[e ?c B (g B d ?b B e ?b) B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, (g B (d ?b B e ?b)) B f, d ?a] B e ?a) B
                           ((e ?c B 𝖺B-1[g B d ?b B e ?b, f, d ?a]) B e ?a) B
                           (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
          proof -
            have "e ?c B 𝗋B[g B f] B ((g B f) B B.inv (η ?a))
                    = (e ?c B (g B 𝗋B[f] B (f B B.inv (η ?a)))) B
                      (e ?c B 𝖺B[g, f, d ?a B e ?a])"
              using hseqB B.runit_hcomp B.whisker_left B.comp_assoc
                    B.assoc_naturality [of g f "B.inv (η ?a)"]
              by auto
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = ((e ?c B (g B 𝗋B[f] B (f B B.inv (η ?a)))) B
                           (e ?c B (𝗋B[g] B (g B B.inv (η ?b)) B f B d ?a B e ?a))) B
                           (e ?c B 𝖺B[g B d ?b B e ?b, f, d ?a B e ?a]) B
                           𝖺B[e ?c, (g B d ?b B e ?b) B f, d ?a B e ?a] B
                           𝖺B[e ?c B (g B d ?b B e ?b) B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, (g B (d ?b B e ?b)) B f, d ?a] B e ?a) B
                           ((e ?c B 𝖺B-1[g B d ?b B e ?b, f, d ?a]) B e ?a) B
                           (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
          proof -
            have "(e ?c B 𝖺B[g, f, d ?a B e ?a]) B
                  (e ?c B (𝗋B[g] B (g B B.inv (η ?b)) B f) B d ?a B e ?a)
                    = e ?c B 𝖺B[g, f, d ?a B e ?a] B
                              ((𝗋B[g] B (g B B.inv (η ?b)) B f) B d ?a B e ?a)"
              using hseqB B.whisker_left by auto
            also have "... = e ?c B (𝗋B[g] B (g B B.inv (η ?b)) B f B d ?a B e ?a) B
                                     𝖺B[g B d ?b B e ?b, f, d ?a B e ?a]"
              using hseqB B.assoc_naturality [of "𝗋B[g] B (g B B.inv (η ?b))" f "d ?a B e ?a"]
              by auto
            also have "... = (e ?c B (𝗋B[g] B (g B B.inv (η ?b)) B f B d ?a B e ?a)) B
                             (e ?c B 𝖺B[g B d ?b B e ?b, f, d ?a B e ?a])"
              using hseqB B.whisker_left by auto
            finally have "(e ?c B 𝖺B[g, f, d ?a B e ?a]) B
                          (e ?c B (𝗋B[g] B (g B B.inv (η ?b)) B f) B d ?a B e ?a)
                            = (e ?c B (𝗋B[g] B (g B B.inv (η ?b)) B f B d ?a B e ?a)) B
                              (e ?c B 𝖺B[g B d ?b B e ?b, f, d ?a B e ?a])"
              by simp
            thus ?thesis
              using B.comp_assoc by simp
          qed
          also have "... = (e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B
                                    𝗋B[f] B (f B B.inv (η ?a))) B
                           (e ?c B 𝖺B[g B d ?b B e ?b, f, d ?a B e ?a]) B
                           𝖺B[e ?c, (g B d ?b B e ?b) B f, d ?a B e ?a] B
                           𝖺B[e ?c B (g B d ?b B e ?b) B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, (g B (d ?b B e ?b)) B f, d ?a] B e ?a) B
                           ((e ?c B 𝖺B-1[g B d ?b B e ?b, f, d ?a]) B e ?a) B
                           (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
          proof -
            have "(e ?c B (g B 𝗋B[f] B (f B B.inv (η ?a)))) B
                  (e ?c B (𝗋B[g] B (g B B.inv (η ?b)) B f B d ?a B e ?a))
                    = e ?c B (g B 𝗋B[f] B (f B B.inv (η ?a))) B
                              (𝗋B[g] B (g B B.inv (η ?b)) B f B d ?a B e ?a)"
              using hseqB B.whisker_left by auto
            also have "... = e ?c B g B 𝗋B[g] B (g B B.inv (η ?b))
                                  B (𝗋B[f] B (f B B.inv (η ?a))) B (f B d ?a B e ?a)"
              using hseqB B.interchange by auto
            also have "... = e ?c B (g B 𝗋B[g]) B (g B B.inv (η ?b))
                                  B 𝗋B[f] B (f B B.inv (η ?a)) B (f B d ?a B e ?a)"
              using B.comp_assoc by simp
            also have "... = e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B 𝗋B[f] B (f B B.inv (η ?a))"
              using hseqB B.comp_arr_dom B.comp_cod_arr by auto
            finally have "(e ?c B (g B 𝗋B[f] B (f B B.inv (η ?a)))) B
                          (e ?c B (𝗋B[g] B (g B B.inv (η ?b)) B f B d ?a B e ?a))
                            = e ?c B 𝗋B[g] B (g B B.inv (η ?b)) B 𝗋B[f] B (f B B.inv (η ?a))"
              by simp
            thus ?thesis
              by simp
          qed
          also have "... = (e ?c B 𝗋B[g] B (g B B.inv (η ?b))
                                 B 𝗋B[f] B (f B B.inv (η ?a))) B
                           𝖺B[e ?c, g B d ?b B e ?b, f B d ?a B e ?a] B
                           (𝖺B[e ?c, g, d ?b B e ?b] B f B d ?a B e ?a) B
                           (𝖺B[e ?c B g, d ?b, e ?b] B f B d ?a B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B e ?b) B (f B d ?a B e ?a)) B
                           𝖺B-1[P g, e ?b, f B d ?a B e ?a] B
                           (P g B 𝖺B[e ?b, f, d ?a B e ?a]) B
                           (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                           (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                           𝖺B[P g, P f, e ?a]"
          proof -
            let ?LHS = "(e ?c B 𝖺B[g B d ?b B e ?b, f, d ?a B e ?a]) B
                        𝖺B[e ?c, (g B d ?b B e ?b) B f, d ?a B e ?a] B
                        𝖺B[e ?c B (g B d ?b B e ?b) B f, d ?a, e ?a] B
                        (𝖺B-1[e ?c, (g B (d ?b B e ?b)) B f, d ?a] B e ?a) B
                        ((e ?c B 𝖺B-1[g B d ?b B e ?b, f, d ?a]) B e ?a) B
                        (((e ?c B 𝖺B-1[g, d ?b B e ?b, f B d ?a])) B e ?a) B
                        ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                        (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                        (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                        ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"

            let ?RHS = "𝖺B[e ?c, g B d ?b B e ?b, f B d ?a B e ?a] B
                        (𝖺B[e ?c, g, d ?b B e ?b] B f B d ?a B e ?a) B
                        (𝖺B[e ?c B g, d ?b, e ?b] B f B d ?a B e ?a) B
                        ((𝖺B-1[e ?c, g, d ?b] B e ?b) B (f B d ?a B e ?a)) B
                        𝖺B-1[P g, e ?b, f B d ?a B e ?a] B
                        (P g B 𝖺B[e ?b, f, d ?a B e ?a]) B
                        (P g B 𝖺B[e ?b B f, d ?a, e ?a]) B
                        (P g B 𝖺B-1[e ?b, f, d ?a] B e ?a) B
                        𝖺B[P g, P f, e ?a]"
            have "?LHS = ?RHS"
            proof -
              let ?LHSt = "(e ?c  𝖺[g  d ?b  e ?b, f, d ?a  e ?a]) 
                           𝖺[e ?c, (g  d ?b  e ?b)  f, d ?a  e ?a] 
                           𝖺[e ?c  (g  d ?b  e ?b)  f, d ?a, e ?a] 
                           (𝖺-1[e ?c, (g  (d ?b  e ?b))  f, d ?a]  e ?a) 
                           ((e ?c  𝖺-1[g  d ?b  e ?b, f, d ?a])  e ?a) 
                           (((e ?c  𝖺-1[g, d ?b  e ?b, f  d ?a]))  e ?a) 
                           ((e ?c  g  𝖺-1[d ?b, e ?b, f  d ?a])  e ?a) 
                           (𝖺[e ?c, g, d ?b  PRJ f]  e ?a) 
                           (𝖺[e ?c  g, d ?b, PRJ f]  e ?a) 
                           ((𝖺-1[e ?c, g, d ?b ]  PRJ f)  e ?a)"

              let ?RHSt = "𝖺[e ?c, g  d ?b  e ?b, f  d ?a  e ?a] 
                           (𝖺[e ?c, g, d ?b  e ?b]  f  d ?a  e ?a) 
                           (𝖺[e ?c  g, d ?b, e ?b]  f  d ?a  e ?a) 
                           ((𝖺-1[e ?c, g, d ?b]  e ?b)  ( f  d ?a  e ?a)) 
                           𝖺-1[PRJ g, e ?b, f  d ?a  e ?a] 
                           (PRJ g  𝖺[e ?b, f, d ?a  e ?a]) 
                           (PRJ g  𝖺[e ?b  f, d ?a, e ?a]) 
                           (PRJ g  𝖺-1[e ?b, f, d ?a]  e ?a) 
                           𝖺[PRJ g, PRJ f, e ?a]"

              have "?LHS = ?LHSt"
                using f g fg hseqB B.α_def B.α'.map_ide_simp B.VVV.ide_charSbC B.VVV.arr_charSbC
                      B.VV.ide_charSbC B.VV.arr_charSbC P_def
                by auto
              also have "... = ?RHSt"
                using hseqB by (intro EV.eval_eqI, auto)
              also have "... = ?RHS"
                using f g fg hseqB B.α_def B.α'.map_ide_simp B.VVV.ide_charSbC B.VVV.arr_charSbC
                      B.VV.ide_charSbC B.VV.arr_charSbC P_def
                by auto
              finally show ?thesis by blast
            qed
            thus ?thesis using hseqB by auto
          qed
          finally show ?thesis by simp
        qed
        also have "... = η1 (g  f)  (PoE.cmp (g, f)  e (src f))"
        proof -
          have "η1 (g  f)  (PoE.cmp (g, f)  e (src f))
                  = η1 (g  f)  (P (E (g  f))  P (ΦE (g, f))  ΦP (E g, E f)  e (src f))"
            using f g fg PoE.cmp_def VV.arr_charSbC VV.dom_charSbC by simp
          also have "... = η1 (g  f)  (P (E (g  f))  P (g B f)  ΦP (E g, E f)  e (src f))"
            using f g fg emb.cmp_def VV.arr_charSbC by simp
          also have "... = η1 (g  f)  (P (E (g  f))  ΦP (E g, E f)  e (src f))"
            using f g fg hseqB comp_cod_arr emb.map_def hseq_char' prj.cmp_simps'(1,5)
            by auto
          also have "... = η1 (g  f)  (P (g B f)  ΦP (g, f)  e (src f))"
            using hseqB hseq emb.map_def hcomp_char hseqI' by auto
          also have "... = η1 (g  f)  (ΦP (g, f)  (P g  P f)  e (src f))"
            using hseqB B.VV.arr_charSbC B.VV.dom_charSbC B.VV.cod_charSbC ΦP.naturality [of "(g, f)"]
                  P.FF_def
            by auto
          also have "... = η1 (g  f)  (ΦP (g, f)  e (src f))"
          proof -
            have "ΦP (g, f)  (P g  P f) = ΦP (g, f)"
              using f g fg comp_arr_dom [of "ΦP (g, f)" "P g  P f"] VV.arr_charSbC
              by (metis (no_types, lifting) prj.cmp_simps'(1) ΦP_simps(4) hseqB)
            thus ?thesis by simp
          qed
          also have "... = η1 (g  f) B (ΦP (g, f)  e (src f))"
          proof -
            (* TODO: For some reason, this requires an epic struggle. *)
            have "seq (η1 (g  f)) (ΦP (g, f)  e (src f))"
            proof -
              have "cod (ΦP (g, f)) = P (g  f)"
                using f g fg ΦP_simps(5) [of g f] VV.arr_charSbC
                by (metis (no_types, lifting) hcomp_eqI hseq hseqB hseqI')
              moreover have "hseq (ΦP (g, f)) (e (src f))"
                using f g fg hseqB 1 B.VV.arr_charSbC by auto
              ultimately show ?thesis
                using f g fg hseqB 1 B.VV.arr_charSbC P0_props prj.cmp_simps emb.map_def
                      ΦP_simps(5) [of g f] hcomp_eqI B.VV.dom_charSbC B.VV.cod_charSbC P.FF_def
                by auto
            qed
            thus ?thesis
              using comp_eqI by simp
          qed
          also have "... = η1 (g B f) B (ΦP (g, f) B e (src f))"
            using f g fg hseqB 1 B.VV.arr_charSbC hcomp_eqI by simp
          also have "... = ((e ?c B 𝗋B[g B f]) B
                           𝖺B[e ?c, g B f, ?a] B
                           ((e ?c B g B f) B B.inv (η ?a)) B
                           𝖺B[e ?c B g B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, g B f, d ?a] B e ?a)) B
                           (ΦP (g, f) B e ?a)"
            unfolding unit1_def
            using hseqB 1 comp_char by simp
          also have "... = ((e ?c B 𝗋B[g B f]) B
                           𝖺B[e ?c, g B f, ?a] B
                           ((e ?c B g B f) B B.inv (η ?a)) B
                           𝖺B[e ?c B g B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, g B f, d ?a] B e ?a)) B
                           (((e ?c B 𝖺B-1[g, f, d ?a]) B
                           (e ?c B g B
                              𝗅B[f B d ?a] B
                              (B.inv (η ?b) B f B d ?a)) B
                              (e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B
                              𝖺B[e ?c, g, d ?b B P f] B
                              𝖺B[e ?c B g, d ?b, P f] B
                              (𝖺B-1[e ?c, g, d ?b] B P f))
                                 B e ?a)"
          proof -
            have "B.VV.ide (g, f)"
              using f g fg ide_charSbC src_def trg_def B.VV.ide_charSbC by auto
            hence "ΦP (g, f) = CMP g f"
              unfolding ΦP_def
              using f g fg ide_charSbC CMP.map_simp_ide by simp
            also have "... = (e ?c B 𝖺B-1[g, f, d ?a]) B
                             (e ?c B g B 𝗅B[f B d ?a] B
                             (B.inv (η ?b) B f B d ?a)) B
                             (e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B
                             𝖺B[e ?c, g, d ?b B P f] B
                             𝖺B[e ?c B g, d ?b, P f] B
                             (𝖺B-1[e ?c, g, d ?b] B P f)"
              using hseqB CMP_def by auto
            finally have "ΦP (g, f) = (e ?c B 𝖺B-1[g, f, d ?a]) B
                                      (e ?c B g B 𝗅B[f B d ?a] B
                                      (B.inv (η ?b) B f B d ?a)) B
                                      (e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B
                                      𝖺B[e ?c, g, d ?b B P f] B
                                      𝖺B[e ?c B g, d ?b, P f] B
                                      (𝖺B-1[e ?c, g, d ?b] B P f)"
              by blast
            thus ?thesis by simp
          qed
          also have "... = ((e ?c B 𝗋B[g B f]) B
                           𝖺B[e ?c, g B f, src f] B
                           ((e ?c B g B f) B B.inv (η ?a)) B
                           𝖺B[e ?c B g B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, g B f, d ?a] B e ?a)) B
                           ((e ?c B 𝖺B-1[g, f, d ?a]) B
                            (e ?c B g B 𝗅B[f B d ?a]) B
                            (e ?c B g B B.inv (η ?b) B f B d ?a) B
                            (e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B
                            𝖺B[e ?c, g, d ?b B P f] B
                            𝖺B[e ?c B g, d ?b, P f] B
                            (𝖺B-1[e ?c, g, d ?b] B P f)
                               B e ?a)"
            using f g fg hseqB src_def trg_def B.whisker_left B.comp_assoc by simp
          also have "... = (e ?c B 𝗋B[g B f]) B
                           𝖺B[e ?c, g B f, src f] B
                           ((e ?c B g B f) B B.inv (η ?a)) B
                           𝖺B[e ?c B g B f, d ?a, e ?a] B
                           (𝖺B-1[e ?c, g B f, d ?a] B e ?a) B
                           ((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                           ((e ?c B g B 𝗅B[f B d ?a]) B e ?a) B
                           ((e ?c B g B B.inv (η ?b) B f B d ?a) B e ?a) B
                           ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                           (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                           (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                           ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
          proof -
            have "B.arr ((e ?c B 𝖺B-1[g, f, d ?a]) B
                        (e ?c B g B 𝗅B[f B d ?a]) B
                        (e ?c B g B B.inv (η ?b) B f B d ?a) B
                        (e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B
                        𝖺B[e ?c, g, d ?b B P f] B
                        𝖺B[e ?c B g, d ?b, P f] B
                        (𝖺B-1[e ?c, g, d ?b] B P f))"
              using hseqB ide_charSbC P.preserves_ide P_def by auto
            hence "(e ?c B 𝖺B-1[g, f, d ?a]) B
                   (e ?c B g B 𝗅B[f B d ?a]) B
                   (e ?c B g B B.inv (η ?b) B f B d ?a) B
                   (e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B
                   𝖺B[e ?c, g, d ?b B P f] B
                   𝖺B[e ?c B g, d ?b, P f] B
                   (𝖺B-1[e ?c, g, d ?b] B P f)
                      B e ?a
                     = ((e ?c B 𝖺B-1[g, f, d ?a]) B e ?a) B
                       ((e ?c B g B 𝗅B[f B d ?a]) B e ?a) B
                       ((e ?c B g B B.inv (η ?b) B f B d ?a) B e ?a) B
                       ((e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B e ?a) B
                       (𝖺B[e ?c, g, d ?b B P f] B e ?a) B
                       (𝖺B[e ?c B g, d ?b, P f] B e ?a) B
                       ((𝖺B-1[e ?c, g, d ?b] B P f) B e ?a)"
              using hseqB B.whisker_right by fastforce
            thus ?thesis
              using B.comp_assoc by simp
          qed
          finally show ?thesis
            using f ide_charSbC src_def by simp
        qed
        finally show "(e (trg g)  IS.cmp (g, f)) 
                      𝖺[e (trg g), IS.map g, IS.map f] 
                      (η1 g  IS.map f) 
                      inv 𝖺[PoE.map g, e (src g), IS.map f] 
                      (PoE.map g  η1 f) 
                      𝖺[PoE.map g, PoE.map f, e (src f)]
                        = η1 (g  f)  (PoE.cmp (g, f)  e (src f))"
          by blast
      qed
    qed

    abbreviation (input) counit0
    where "counit0  d"

    definition counit1
    where "counit1 f = 𝖺B[d (trgB f), e (trgB f), f B d (srcB f)] B
                       (η (trgB f) B f B d (srcB f)) B
                       𝗅B-1[f B d (srcB f)]"

    abbreviation (input) ε0
    where "ε0  counit0"

    abbreviation (input) ε1
    where "ε1  counit1"

    lemma counit1_in_hom [intro]:
    assumes "B.ide f"
    shows "«ε1 f : f B d (srcB f) B d (trgB f) B e (trgB f) B f B d (srcB f)»"
      using assms B.iso_is_arr
      by (unfold counit1_def, intro B.comp_in_homI' B.seqI B.hseqI' B.hcomp_in_vhom) auto

    lemma counit1_simps [simp]:
    assumes "B.ide f"
    shows "B.arr (ε1 f)"
    and "srcB (ε1 f) = P0 (srcB f)" and "trgB (ε1 f) = trgB f"
    and "B.dom (ε1 f) = f B d (srcB f)"
    and "B.cod (ε1 f) = d (trgB f) B e (trgB f) B f B d (srcB f)"
    and "B.iso (ε1 f)"
      using assms counit1_in_hom
           apply auto
      using B.vconn_implies_hpar(1)
        apply fastforce
      using B.vconn_implies_hpar(2)
        apply fastforce
      unfolding counit1_def
      apply (intro B.isos_compose)
      by auto

    lemma technical:
    assumes "B.ide f" and "B.ide g" and "srcB g = trgB f"
    shows "(ε1 g B P f) B 𝖺B-1[g, d (srcB g), P f] B (g B ε1 f)
             = (𝖺B[d (trgB g), e (trgB g), g B d (srcB g)] B P f) B
               (𝖺B[d (trgB g) B e (trgB g), g, d (srcB g)] B P f) B
               𝖺B-1[(d (trgB g) B e (trgB g)) B g, d (srcB g), P f] B
               (((d (trgB g) B e (trgB g)) B g) B d (srcB g) B P f) B
               (((d (trgB g) B e (trgB g)) B g)
                  B 𝖺B[d (srcB g), e (srcB g), f B d (srcB f)]) B
               ((η (trgB g) B g) B 𝗅B-1[g] B (η (srcB g) B f B d (srcB f))) B
               (g B 𝖺B[srcB g, f, d (srcB f)] B (𝗅B-1[f] B d (srcB f)))"
    proof -
      let ?a = "srcB f"
      let ?b = "srcB g"
      let ?c = "trgB g"
      have "(ε1 g B P f) B 𝖺B-1[g, d ?b, P f] B (g B ε1 f)
              = (𝖺B[d ?c, e ?c, g B d ?b] B
                (η ?c B g B d ?b) B
                𝗅B-1[g B d ?b]
                   B P f) B
                𝖺B-1[g, d ?b, P f] B
                (g B 𝖺B[d ?b, e ?b, f B d ?a] B
                      (η ?b B f B d ?a) B
                      𝗅B-1[f B d ?a])"
        using assms counit1_def by simp
      also have "... = (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                       ((η ?c B g B d ?b) B P f) B
                       (𝗅B-1[g B d ?b] B P f) B
                       𝖺B-1[g, d ?b, P f] B
                       (g B 𝖺B[d ?b, e ?b, f B d ?a]) B
                       (g B η ?b B f B d ?a) B
                       (g B 𝗅B-1[f B d ?a])"
      proof -
        have "𝖺B[d ?c, e ?c, g B d ?b] B
              (η ?c B g B d ?b) B
              𝗅B-1[g B d ?b]
                B P f
                = (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                  ((η ?c B g B d ?b) B P f) B
                  (𝗅B-1[g B d ?b] B P f)"
          using assms B.iso_is_arr B.whisker_right P_def by auto
        moreover have "g B 𝖺B[d ?b, e ?b, f B d ?a] B
                           (η ?b B f B d ?a) B
                           𝗅B-1[f B d ?a]
                         = (g B 𝖺B[d ?b, e ?b, f B d ?a]) B
                           (g B η ?b B f B d ?a) B
                           (g B 𝗅B-1[f B d ?a])"
          using assms B.iso_is_arr B.whisker_left by simp
        ultimately show ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                       ((η ?c B g B d ?b) B P f) B
                       (𝖺B[?c, g, d ?b] B (𝗅B-1[g] B d ?b) B P f) B
                       𝖺B-1[g, d ?b, P f] B
                       (g B 𝖺B[d ?b, e ?b, f B d ?a]) B
                       (g B η ?b B f B d ?a) B
                       (g B 𝗅B-1[f B d ?a])"
        using assms B.lunit_hcomp [of g "d ?b"] B.comp_assoc by simp
      also have "... = (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                       (((η ?c B g B d ?b) B P f) B
                       (𝖺B[?c, g, d ?b] B P f)) B
                       ((𝗅B-1[g] B d ?b) B P f) B
                       𝖺B-1[g, d ?b, P f] B
                       (g B 𝖺B[d ?b, e ?b, f B d ?a]) B
                       (g B η ?b B f B d ?a) B
                       (g B 𝗅B-1[f B d ?a])"
        using assms B.whisker_right P_def B.comp_assoc by simp
      also have "... = (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                       (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                       (((η ?c B g) B d ?b) B P f) B
                       (((𝗅B-1[g] B d ?b) B P f) B
                       𝖺B-1[g, d ?b, P f]) B
                       (g B 𝖺B[d ?b, e ?b, f B d ?a]) B
                       (g B η ?b B f B d ?a) B
                       (g B 𝗅B-1[f B d ?a])"
      proof -
        have "((η ?c B g B d ?b) B P f) B (𝖺B[?c, g, d ?b] B P f) =
              (η ?c B g B d ?b) B 𝖺B[?c, g, d ?b] B P f"
          using assms B.whisker_right P_def B.iso_is_arr by simp
        also have "... = 𝖺B[d ?c B e ?c, g, d ?b] B ((η ?c B g) B d ?b) B P f"
          using assms B.assoc_naturality [of "η ?c" g "d ?b"] B.iso_is_arr by simp
        also have "... = (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                         (((η ?c B g) B d ?b) B P f)"
          using assms B.whisker_right P_def B.iso_is_arr by simp
        finally have "((η ?c B g B d ?b) B P f) B (𝖺B[?c, g, d ?b] B P f) =
                        (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                        (((η ?c B g) B d ?b) B P f)"
          by blast
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                       (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                       ((((η ?c B g) B d ?b) B P f) B
                       𝖺B-1[?c B g, d ?b, P f]) B
                       (𝗅B-1[g] B d ?b B P f) B
                       (g B 𝖺B[d ?b, e ?b, f B d ?a]) B
                       (g B η ?b B f B d ?a) B
                       (g B 𝗅B-1[f B d ?a])"
      proof -
        have "((𝗅B-1[g] B d ?b) B P f) B 𝖺B-1[g, d ?b, P f] =
              𝖺B-1[?c B g, d ?b, P f] B (𝗅B-1[g] B d ?b B P f)"
          using assms B.assoc'_naturality [of "𝗅B-1[g]" "d ?b" "P f"] by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                       (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                       𝖺B-1[(d ?c B e ?c) B g, d ?b, P f] B
                       (((η ?c B g) B d ?b B P f) B
                       (𝗅B-1[g] B d ?b B P f)) B
                       (g B 𝖺B[d ?b, e ?b, f B d ?a]) B
                       (g B η ?b B f B d ?a) B
                       (g B 𝗅B-1[f B d ?a])"
      proof -
        have "(((η ?c B g) B d ?b) B P f) B 𝖺B-1[?c B g, d ?b, P f] =
              𝖺B-1[(d ?c B e ?c) B g, d ?b, P f] B ((η ?c B g) B d ?b B P f)"
          using assms B.assoc'_naturality [of "η ?c B g" "d ?b" "P f"] B.iso_is_arr
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                       (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                       𝖺B-1[(d ?c B e ?c) B g, d ?b, P f] B
                       (((η ?c B g) B 𝗅B-1[g] B d ?b B P f) B
                       (g B 𝖺B[d ?b, e ?b, f B d ?a])) B
                       (g B η ?b B f B d ?a) B
                       (g B 𝗅B-1[f B d ?a])"
        using assms B.whisker_right B.iso_is_arr P_def B.comp_assoc by simp
      also have "... = (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                       (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                       𝖺B-1[(d ?c B e ?c) B g, d ?b, P f] B
                       (((d ?c B e ?c) B g) B d ?b B P f) B
                       (((η ?c B g) B 𝗅B-1[g] B 𝖺B[d ?b, e ?b, f B d ?a]) B
                       (g B η ?b B f B d ?a)) B
                       (g B 𝗅B-1[f B d ?a])"
      proof -
        have "((η ?c B g) B 𝗅B-1[g] B d ?b B P f) B
              (g B 𝖺B[d ?b, e ?b, f B d ?a])
                = (((d ?c B e ?c) B g) B d ?b B P f) B
                  ((η ?c B g) B 𝗅B-1[g] B 𝖺B[d ?b, e ?b, f B d ?a])"
          using assms B.comp_arr_dom B.comp_cod_arr B.iso_is_arr P_def
                B.interchange [of "(η ?c B g) B 𝗅B-1[g]" g]
                B.interchange [of "(d ?c B e ?c) B g" "(η ?c B g) B 𝗅B-1[g]"]
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                       (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                       𝖺B-1[(d ?c B e ?c) B g, d ?b, P f] B
                       (((d ?c B e ?c) B g) B d ?b B P f) B
                       (((d ?c B e ?c) B g) B 𝖺B[d ?b, e ?b, f B d ?a]) B
                       (((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a) B
                       (g B 𝗅B-1[f B d ?a]))"
      proof -
        have "((η ?c B g) B 𝗅B-1[g] B 𝖺B[d ?b, e ?b, f B d ?a]) B (g B η ?b B f B d ?a)
                = (((d ?c B e ?c) B g) B 𝖺B[d ?b, e ?b, f B d ?a]) B
                  ((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a)"
          using assms B.comp_arr_dom B.comp_cod_arr B.iso_is_arr P_def
                B.interchange [of "(η ?c B g) B 𝗅B-1[g]" g]
                B.interchange [of "(d ?c B e ?c) B g" "(η ?c B g) B 𝗅B-1[g]"]
          by simp
        thus ?thesis
          using B.comp_assoc by simp
      qed
      also have "... = (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                       (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                       𝖺B-1[(d ?c B e ?c) B g, d ?b, P f] B
                       (((d ?c B e ?c) B g) B d ?b B P f) B
                       (((d ?c B e ?c) B g) B 𝖺B[d ?b, e ?b, f B d ?a]) B
                       ((η ?c B g) B 𝗅B-1[g] B (η ?b B f B d ?a)) B
                       (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a))"
      proof -
        have "((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a) B (g B 𝗅B-1[f B d ?a])
                = (η ?c B g) B 𝗅B-1[g] B (η ?b B f B d ?a) B 𝗅B-1[f B d ?a]"
          using assms B.comp_arr_dom B.iso_is_arr P_def
                B.interchange [of "(η ?c B g) B 𝗅B-1[g]" g "η ?b B f B d ?a" "𝗅B-1[f B d ?a]"]
          by simp
        also have "... = (η ?c B g) B 𝗅B-1[g] B
                           (η ?b B f B d ?a) B
                           𝖺B[?b, f, d ?a] B
                           (𝗅B-1[f] B d ?a)"
          using assms B.lunit_hcomp by simp
        also have "... = ((η ?c B g) B 𝗅B-1[g] B (η ?b B f B d ?a)) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a))"
          using assms B.comp_arr_dom B.iso_is_arr P_def
                B.interchange [of "(η ?c B g) B 𝗅B-1[g]" g]
          by simp
        finally have "((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a) B (g B 𝗅B-1[f B d ?a])
                        = ((η ?c B g) B 𝗅B-1[g] B (η ?b B f B d ?a)) B
                          (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a))"
          by blast
        thus ?thesis by simp
      qed
      finally show ?thesis by blast
    qed

    sublocale counit: pseudonatural_equivalence
                        VB HB 𝖺B 𝗂B srcB trgB VB HB 𝖺B 𝗂B srcB trgB
                        E  P EoP.cmp IC.map IC.cmp counit0 counit1
    proof
      show "a. B.obj a  B.ide (d a)"
        by simp
      show "a. B.obj a  B.equivalence_map (d a)"
      proof -
        fix a
        assume a: "B.obj a"
        interpret Adj: adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB
                         e a d a η a ε a
          using a chosen_adjoint_equivalence by simp
        show "B.equivalence_map (d a)"
          using Adj.equivalence_in_bicategory_axioms B.equivalence_map_def Adj.dual_equivalence
          by blast
      qed
      show "a. B.obj a  «d a : srcB ((E  P) a) B srcB (IC.map a)»"
        unfolding emb.map_def by fastforce
      show "f. B.ide f  B.iso (ε1 f)"
        by simp
      show "f. B.ide f  «ε1 f : IC.map f B d (srcB f) B d (trgB f) B (E  P) f»"
        unfolding counit1_def P_def emb.map_def
        using P_def P_simps(1)
        by (intro B.comp_in_homI' B.seqI B.hseqI') (auto simp add: B.iso_is_arr)
      show "μ. B.arr μ  ε1 (B.cod μ) B (IC.map μ B d (srcB μ))
                               = (d (trgB μ) B (E  P) μ) B ε1 (B.dom μ)"
      proof -
        fix μ
        assume μ: "B.arr μ"
        let ?a = "srcB μ"
        let ?b = "trgB μ"
        let ?f = "B.dom μ"
        let ?g = "B.cod μ"
        have "ε1 ?g B (IC.map μ B d ?a)
                = (𝖺B[d ?b, e ?b, ?g B d ?a] B
                  (η ?b B ?g B d ?a) B
                  𝗅B-1[?g B d ?a]) B
                  (μ B d ?a)"
          using μ counit1_def P_def emb.map_def arr_charSbC P0_props(1) by simp
        also have "... = (d ?b B e ?b B μ B d ?a) B
                         𝖺B[d ?b, e ?b, ?f B d ?a] B
                         (η ?b B ?f B d ?a) B
                         𝗅B-1[?f B d ?a]"
        proof -
          have "(𝖺B[d ?b, e ?b, ?g B d ?a] B
                (η ?b B ?g B d ?a) B
                𝗅B-1[?g B d ?a]) B
                (μ B d ?a)
                  = 𝖺B[d ?b, e ?b, ?g B d ?a] B
                    (η ?b B ?g B d ?a) B
                    𝗅B-1[?g B d ?a] B
                    (μ B d ?a)"
            using B.comp_assoc by simp
          also have "... = 𝖺B[d ?b, e ?b, ?g B d ?a] B
                           ((η ?b B ?g B d ?a) B
                           (?b B μ B d ?a)) B
                           𝗅B-1[?f B d ?a]"
            using μ B.lunit'_naturality [of "μ B d ?a"] B.comp_assoc by simp
          also have "... = (𝖺B[d ?b, e ?b, ?g B d ?a] B
                           ((d ?b B e ?b) B μ B d ?a)) B
                           (η ?b B ?f B d ?a) B
                           𝗅B-1[?f B d ?a]"
          proof -
            have "(η ?b B ?g B d ?a) B (?b B μ B d ?a) =
                    η ?b B ?b B ?g B μ B d ?a B d ?a"
              using μ B.iso_is_arr
                    B.interchange [of "η ?b" ?b "(?g B d ?a)" "μ B d ?a"]
                    B.interchange [of "?g" μ "d ?a" "d ?a"]
              by simp
            also have "... = (d ?b B e ?b) B η ?b B μ B ?f B d ?a B d ?a"
              using μ B.comp_arr_dom B.comp_cod_arr apply simp
              by (metis B.arrI equivalence_data_in_homB(7) equivalence_data_simpsB(17-18)
                  B.obj_trg)
            also have "... = ((d ?b B e ?b) B μ B d ?a) B (η ?b B ?f B d ?a)"
              using μ B.iso_is_arr B.interchange [of "d ?b B e ?b" "η ?b" "μ B d ?a" "?f B d ?a"]
                    B.interchange [of μ "?f" "d ?a" "d ?a"]
              by simp
            finally have "(η ?b B ?g B d ?a) B (?b B μ B d ?a)
                            = ((d ?b B e ?b) B μ B d ?a) B (η ?b B ?f B d ?a)"
              by blast
            thus ?thesis using B.comp_assoc by simp
          qed
          also have "... = (d ?b B e ?b B μ B d ?a) B
                           𝖺B[d ?b, e ?b, ?f B d ?a] B
                           (η ?b B ?f B d ?a) B
                           𝗅B-1[?f B d ?a]"
            using μ B.assoc_naturality [of "d ?b" "e ?b" "μ B d ?a"] B.comp_assoc by simp
          finally show ?thesis by simp
        qed
        also have "... = (d ?b B (E  P) μ) B ε1 ?f"
          using μ counit1_def P_def emb.map_def arr_charSbC P0_props(1) P_simpsB(1)
          by (simp add: P0_props(6))
        finally show "ε1 (?g) B (IC.map μ B d ?a) = (d ?b B (E  P) μ) B ε1 ?f"
          by blast
      qed
      show "f g. B.ide f; B.ide g; srcB g = trgB f 
                    (d (trgB g) B EoP.cmp (g, f)) B
                    𝖺B[d (trgB g), (E  P) g, (E  P) f] B
                    (ε1 g B (E  P) f) B
                    B.inv 𝖺B[IC.map g, d (srcB g), (E  P) f] B (IC.map g B ε1 f) B
                    𝖺B[IC.map g, IC.map f, d (srcB f)]
                      = ε1 (g B f) B (IC.cmp (g, f) B d (srcB f))"
      proof -
        fix f g
        assume f: "B.ide f" and g: "B.ide g" and fg: "srcB g = trgB f"
        let ?a = "srcB f"
        let ?b = "srcB g"
        let ?c = "trgB g"
        have "(d ?c B EoP.cmp (g, f)) B
              𝖺B[d ?c, (E  P) g, (E  P) f] B
              (ε1 g B (E  P) f) B
              B.inv 𝖺B[IC.map g, d ?b, (E  P) f] B
              (IC.map g B ε1 f) B
              𝖺B[IC.map g, IC.map f, d ?a]
                = (d ?c B P (g B f) B CMP.map (g, f) B (P g B P f)) B
                  𝖺B[d ?c, P g, P f] B
                  (ε1 g B P f) B
                  𝖺B-1[g, d ?b, P f] B
                  (g B ε1 f) B
                  𝖺B[g, f, d ?a]"
        proof -
          have "srcB (e ?c B g B d ?b) = trgB (e ?b B f B d ?a)"
            using f g fg src_def trg_def arr_charSbC P_simpsB(1)
            by (simp add: P0_props(1,6) arr_charSbC)
          moreover have "B.inv 𝖺B[g, d ?b, P f] = 𝖺B-1[g, d ?b, P f]"
            using f g fg by (simp add: P_def)
          ultimately show ?thesis
            using f g fg emb.map_def P.preserves_reflects_arr P.preserves_reflects_arr
                  P.preserves_reflects_arr ΦP_def emb.cmp_def
                  EoP.cmp_def B.VV.arr_charSbC B.VV.dom_charSbC VV.arr_charSbC
            by simp
        qed
        also have "... = (d ?c B CMP.map (g, f)) B
                         𝖺B[d ?c, P g, P f] B
                         ((ε1 g B P f) B
                         𝖺B-1[g, d ?b, P f] B
                         (g B ε1 f)) B
                         𝖺B[g, f, d ?a]"
        proof -
          have "«CMP.map (g, f) : P g  P f  P (g B f)»"
            using f g fg VV.arr_charSbC VV.cod_charSbC P.FF_def ΦP_def ΦP_in_hom(2) by auto
          hence "«CMP.map (g, f) : P g B P f B P (g B f)»"
            using in_hom_charSbC hcomp_char by auto
          hence "P (g B f) B CMP.map (g, f) B (P g B P f) = CMP.map (g, f)"
            using B.comp_arr_dom B.comp_cod_arr by auto
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = (d ?c B CMP.map (g, f)) B
                         𝖺B[d ?c, P g, P f] B
                         (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                         (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                         𝖺B-1[(d ?c B e ?c) B g, d ?b, P f] B
                         (((d ?c B e ?c) B g) B d ?b B P f) B
                         (((d ?c B e ?c) B g) B 𝖺B[d ?b, e ?b, f B d ?a]) B
                         ((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
          using f g fg technical B.comp_assoc by simp
        also have "... = (d ?c B (e ?c B 𝖺B-1[g, f, d ?a]) B
                         (e ?c B g B 𝗅B[f B d ?a]) B
                         (e ?c B g B B.inv (η ?b) B f B d ?a) B
                         (e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B
                         𝖺B[e ?c, g, d ?b B P f] B
                         𝖺B[e ?c B g, d ?b, P f] B
                         (𝖺B-1[e ?c, g, d ?b] B P f)) B
                         𝖺B[d ?c, P g, P f] B
                         (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                         (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                         𝖺B-1[(d ?c B e ?c) B g, d ?b, P f] B
                         (((d ?c B e ?c) B g) B d ?b B P f) B
                         (((d ?c B e ?c) B g) B 𝖺B[d ?b, e ?b, f B d ?a]) B
                         ((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
          using f g fg CMP.map_simp_ide CMP_def B.VV.ide_charSbC B.VV.arr_charSbC
                B.whisker_left B.whisker_left B.comp_assoc
          by simp
        also have "... = (d ?c B e ?c B 𝖺B-1[g, f, d ?a]) B
                         (d ?c B e ?c B g B 𝗅B[f B d ?a]) B
                         (d ?c B e ?c B g B B.inv (η ?b) B f B d ?a) B
                         ((d ?c B e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B
                         (d ?c B 𝖺B[e ?c, g, d ?b B P f]) B
                         (d ?c B 𝖺B[e ?c B g, d ?b, P f]) B
                         (d ?c B 𝖺B-1[e ?c, g, d ?b] B P f) B
                         𝖺B[d ?c, P g, P f] B
                         (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                         (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                         𝖺B-1[(d ?c B e ?c) B g, d ?b, P f] B
                         (((d ?c B e ?c) B g) B d ?b B P f) B
                         (((d ?c B e ?c) B g) B 𝖺B[d ?b, e ?b, f B d ?a])) B
                         ((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
          using f g fg B.whisker_left P_def B.comp_assoc by simp (* 11 sec *)
        also have "... = (d ?c B e ?c B 𝖺B-1[g, f, d ?a]) B
                         (d ?c B e ?c B g B 𝗅B[f B d ?a]) B
                         ((d ?c B e ?c B g B B.inv (η ?b) B f B d ?a) B
                         𝖺B[d ?c, e ?c, g B (d ?b B e ?b) B f B d ?a]) B
                         𝖺B[d ?c B e ?c, g, (d ?b B e ?b) B f B d ?a] B
                         ((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
        proof -
          have "(d ?c B e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B
                (d ?c B 𝖺B[e ?c, g, d ?b B P f]) B
                (d ?c B 𝖺B[e ?c B g, d ?b, P f]) B
                (d ?c B 𝖺B-1[e ?c, g, d ?b] B P f) B
                𝖺B[d ?c, P g, P f] B
                (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                𝖺B-1[(d ?c B e ?c) B g, d ?b, P f] B
                (((d ?c B e ?c) B g) B d ?b B P f) B
                (((d ?c B e ?c) B g) B 𝖺B[d ?b, e ?b, f B d ?a])
                  = 𝖺B[d ?c, e ?c, g B (d ?b B e ?b) B f B d ?a] B
                    𝖺B[d ?c B e ?c, g, (d ?b B e ?b) B f B d ?a]"
          proof -
            have "(d ?c B e ?c B g B 𝖺B-1[d ?b, e ?b, f B d ?a]) B
                  (d ?c B 𝖺B[e ?c, g, d ?b B P f]) B
                  (d ?c B 𝖺B[e ?c B g, d ?b, P f]) B
                  (d ?c B 𝖺B-1[e ?c, g, d ?b] B P f) B
                  𝖺B[d ?c, P g, P f] B
                  (𝖺B[d ?c, e ?c, g B d ?b] B P f) B
                  (𝖺B[d ?c B e ?c, g, d ?b] B P f) B
                  𝖺B-1[(d ?c B e ?c) B g, d ?b, P f] B
                  (((d ?c B e ?c) B g) B d ?b B P f) B
                  (((d ?c B e ?c) B g) B 𝖺B[d ?b, e ?b, f B d ?a])
                    = (d ?c  e ?c  g  𝖺-1[d ?b, e ?b, f  d ?a]) 
                       (d ?c  𝖺[e ?c, g, d ?b  PRJ f]) 
                       (d ?c  𝖺[e ?c  g, d ?b, PRJ f]) 
                       (d ?c  𝖺-1[e ?c, g, d ?b]  PRJ f) 
                       𝖺[d ?c, PRJ g, PRJ f] 
                       (𝖺[d ?c, e ?c, g  d ?b]  PRJ f) 
                       (𝖺[d ?c  e ?c, g, d ?b]  PRJ f) 
                       𝖺-1[(d ?c  e ?c)  g, d ?b, PRJ f] 
                       (((d ?c  e ?c)  g)  d ?b  PRJ f) 
                       (((d ?c  e ?c)  g) 
                             𝖺[d ?b, e ?b, f  d ?a])"
              using f g fg B.α_def B.α'.map_ide_simp B.VVV.ide_charSbC B.VVV.arr_charSbC
                    B.VV.ide_charSbC B.VV.arr_charSbC P_def B.𝔩_ide_simp
              by auto
            also have "... = 𝖺[d ?c, e ?c,
                                g  (d ?b  e ?b)  f  d ?a] 
                              𝖺[d ?c  e ?c, g,
                                (d ?b  e ?b)  f  d ?a]"
              using f g fg by (intro EV.eval_eqI, auto)
            also have "... = 𝖺B[d ?c, e ?c, g B (d ?b B e ?b) B f B d ?a] B
                             𝖺B[d ?c B e ?c, g, (d ?b B e ?b) B f B d ?a]"
              using f g fg B.α_def B.α'.map_ide_simp B.VVV.ide_charSbC B.VVV.arr_charSbC
                    B.VV.ide_charSbC B.VV.arr_charSbC P_def B.𝔩_ide_simp
              by auto
            finally show ?thesis by blast
          qed
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = (d ?c B e ?c B 𝖺B-1[g, f, d ?a]) B
                         (d ?c B e ?c B g B 𝗅B[f B d ?a]) B
                         𝖺B[d ?c, e ?c, g B ?b B f B d ?a] B
                         (((d ?c B e ?c) B g B B.inv (η ?b) B f B d ?a) B
                         𝖺B[d ?c B e ?c, g, (d ?b B e ?b) B f B d ?a]) B
                         ((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
        proof -
          have "(d ?c B e ?c B g B B.inv (η ?b) B f B d ?a) B
                𝖺B[d ?c, e ?c, g B (d ?b B e ?b) B f B d ?a] =
                  𝖺B[d ?c, e ?c, g B ?b B f B d ?a] B
                  ((d ?c B e ?c) B g B B.inv (η ?b) B f B d ?a)"
            using f g fg
                  B.assoc_naturality
                    [of "d ?c" "e ?c" "g B B.inv (η ?b) B f B d ?a"]
            by simp
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = (d ?c B e ?c B 𝖺B-1[g, f, d ?a]) B
                         (d ?c B e ?c B g B 𝗅B[f B d ?a]) B
                         𝖺B[d ?c, e ?c, g B ?b B f B d ?a] B
                         𝖺B[d ?c B e ?c, g, ?b B f B d ?a] B
                         ((((d ?c B e ?c) B g) B B.inv (η ?b) B f B d ?a) B
                         ((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a)) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
        proof -
          have "((d ?c B e ?c) B g B B.inv (η ?b) B f B d ?a) B
                𝖺B[d ?c B e ?c, g, (d ?b B e ?b) B f B d ?a] =
                  𝖺B[d ?c B e ?c, g, ?b B f B d ?a] B
                  (((d ?c B e ?c) B g) B B.inv (η ?b) B f B d ?a)"
            using f g fg
                  B.assoc_naturality
                    [of "d ?c B e ?c" g "B.inv (η ?b) B f B d ?a"]
            by simp
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = (d ?c B e ?c B 𝖺B-1[g, f, d ?a]) B
                         (d ?c B e ?c B g B 𝗅B[f B d ?a]) B
                         𝖺B[d ?c, e ?c, g B ?b B f B d ?a] B
                         𝖺B[d ?c B e ?c, g, ?b B f B d ?a] B
                         ((η ?c B g) B 𝗅B-1[g] B ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
        proof -
          have "(((d ?c B e ?c) B g) B B.inv (η ?b) B f B d ?a) B
                ((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a)
                  = (η ?c B g) B 𝗅B-1[g] B ?b B f B d ?a"
          proof -
            have "(((d ?c B e ?c) B g) B B.inv (η ?b) B f B d ?a) B
                  ((η ?c B g) B 𝗅B-1[g] B η ?b B f B d ?a)
                    = ((d ?c B e ?c) B g) B (η ?c B g) B 𝗅B-1[g] B
                      (B.inv (η ?b) B f B d ?a) B (η ?b B f B d ?a)"
              using f g fg
                    B.interchange [of "(d ?c B e ?c) B g" "(η ?c B g) B 𝗅B-1[g]"
                                      "B.inv (η ?b) B f B d ?a"
                                      "η ?b B f B d ?a"]
              by fastforce
            also have "... = (η ?c B g) B 𝗅B-1[g] B
                               (B.inv (η ?b) B f B d ?a) B (η ?b B f B d ?a)"
              using f g fg B.comp_cod_arr [of "(η ?c B g) B 𝗅B-1[g]" "(d ?c B e ?c) B g"]
              by (auto simp add: B.iso_is_arr)
            also have "... = (η ?c B g) B 𝗅B-1[g] B ?b B f B d ?a"
            proof -
              have "(B.inv (η ?b) B f B d ?a) B (η ?b B f B d ?a)
                      = B.inv (η ?b) B η ?b B f B d ?a"
              proof -
                have "B.seq (B.inv (η ?b)) (η ?b)"
                  using f g fg chosen_adjoint_equivalence(5) by fastforce
                thus ?thesis
                  using f g fg B.whisker_right by simp
              qed
              also have "... = ?b B f B d ?a"
                using f g fg chosen_adjoint_equivalence(5) B.comp_inv_arr' by simp
              finally have "(B.inv (η ?b) B f B d ?a) B (η ?b B f B d ?a)
                              = ?b B f B d ?a"
                by blast
              thus ?thesis
                using B.comp_assoc by simp
            qed
            finally show ?thesis by simp
          qed
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = (d ?c B e ?c B 𝖺B-1[g, f, d ?a]) B
                         (d ?c B e ?c B g B 𝗅B[f B d ?a]) B
                         𝖺B[d ?c, e ?c, g B ?b B f B d ?a] B
                         (𝖺B[d ?c B e ?c, g, ?b B f B d ?a] B
                         ((η ?c B g) B ?b B f B d ?a)) B
                         (𝗅B-1[g] B ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
          using f g fg B.comp_assoc
                B.whisker_right [of "?b B f B d ?a" "η ?c B g" "𝗅B-1[g]"]
            by fastforce
        also have "... = (d ?c B e ?c B 𝖺B-1[g, f, d ?a]) B
                         ((d ?c B e ?c B g B 𝗅B[f B d ?a]) B
                         𝖺B[d ?c, e ?c, g B ?b B f B d ?a]) B
                         (η ?c B g B ?b B f B d ?a) B
                         𝖺B[?c, g, ?b B f B d ?a] B
                         (𝗅B-1[g] B ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
          using f g fg B.iso_is_arr B.comp_assoc
                B.assoc_naturality [of "η ?c" g "?b B f B d ?a"]
          by simp
        also have "... = (d ?c B e ?c B 𝖺B-1[g, f, d ?a]) B
                         𝖺B[d ?c, e ?c, g B f B d ?a] B
                         (((d ?c B e ?c) B g B 𝗅B[f B d ?a]) B
                         (η ?c B g B ?b B f B d ?a)) B
                         𝖺B[?c, g, ?b B f B d ?a] B
                         (𝗅B-1[g] B ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
         proof -
          have "(d ?c B e ?c B g B 𝗅B[f B d ?a]) B
                𝖺B[d ?c, e ?c, g B ?b B f B d ?a] =
                  𝖺B[d ?c, e ?c, g B f B d ?a] B
                  ((d ?c B e ?c) B g B 𝗅B[f B d ?a])"
            using f g fg B.iso_is_arr
                  B.assoc_naturality [of "d ?c" "e ?c" "g B 𝗅B[f B d ?a]"]
            by simp
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = ((d ?c B e ?c B 𝖺B-1[g, f, d ?a]) B
                         𝖺B[d ?c, e ?c, g B f B d ?a]) B
                         (η ?c B g B f B d ?a) B
                         (?c B g B 𝗅B[f B d ?a]) B
                         𝖺B[?c, g, ?b B f B d ?a] B
                         (𝗅B-1[g] B ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
        proof -
          have "((d ?c B e ?c) B g B 𝗅B[f B d ?a]) B (η ?c B g B ?b B f B d ?a)
                   = (η ?c B g B f B d ?a) B (?c B g B 𝗅B[f B d ?a])"
            using f g fg B.comp_arr_dom B.comp_cod_arr B.iso_is_arr
                  B.interchange [of "d ?c B e ?c" "η ?c"
                                    "g B 𝗅B[f B d ?a]" "g B ?b B f B d ?a"]
                  B.interchange [of "η ?c" ?c
                                    "g B f B d ?a" "g B 𝗅B[f B d ?a]"]
            by auto
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = 𝖺B[d ?c, e ?c, (g B f) B d ?a] B
                         (((d ?c B e ?c) B 𝖺B-1[g, f, d ?a]) B
                         (η ?c B g B f B d ?a)) B
                         (?c B g B 𝗅B[f B d ?a]) B
                         𝖺B[?c, g, ?b B f B d ?a] B
                         (𝗅B-1[g] B ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
        proof -
          have "(d ?c B e ?c B 𝖺B-1[g, f, d ?a]) B
                𝖺B[d ?c, e ?c, g B f B d ?a] =
                  𝖺B[d ?c, e ?c, (g B f) B d ?a] B
                  ((d ?c B e ?c) B 𝖺B-1[g, f, d ?a])"
            using f g fg B.assoc_naturality [of "d ?c" "e ?c" "𝖺B-1[g, f, d ?a]"] by simp
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = 𝖺B[d ?c, e ?c, (g B f) B d ?a] B
                         (η ?c B (g B f) B d ?a) B
                         (?c B 𝖺B-1[g, f, d ?a]) B
                         (?c B g B 𝗅B[f B d ?a]) B
                         𝖺B[?c, g, ?b B f B d ?a] B
                         (𝗅B-1[g] B ?b B f B d ?a) B
                         (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                         𝖺B[g, f, d ?a]"
        proof -
          have "((d ?c B e ?c) B 𝖺B-1[g, f, d ?a]) B (η ?c B g B f B d ?a)
                  = (η ?c B (g B f) B d ?a) B (?c B 𝖺B-1[g, f, d ?a])"
            using f g fg B.comp_arr_dom B.comp_cod_arr B.iso_is_arr
                  B.interchange [of "d ?c B e ?c" "η ?c" "𝖺B-1[g, f, d ?a]" "g B f B d ?a"]
                  B.interchange [of "η ?c" ?c "(g B f) B d ?a" "𝖺B-1[g, f, d ?a]"]
            by auto
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = ε1 (g B f) B (IC.cmp (g, f) B d ?a)"
        proof -
          have "ε1 (g B f) B (IC.cmp (g, f) B d ?a)
                  = 𝖺B[d ?c, e ?c, (g B f) B d ?a] B
                    (η ?c B (g B f) B d ?a) B
                    𝗅B-1[(g B f) B d ?a] B ((g B f) B d ?a)"
            using f g fg counit1_def B.comp_assoc by simp
          also have "... = 𝖺B[d ?c, e ?c, (g B f) B d ?a] B
                           (η ?c B (g B f) B d ?a) B
                           (?c B 𝖺B-1[g, f, d ?a]) B
                           (?c B g B 𝗅B[f B d ?a]) B
                           𝖺B[?c, g, ?b B f B d ?a] B
                           (𝗅B-1[g] B ?b B f B d ?a) B
                           (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                           𝖺B[g, f, d ?a]"
          proof -
            have "𝗅B-1[(g B f) B d ?a] B ((g B f) B d ?a)
                    = (?c B 𝖺B-1[g, f, d ?a]) B
                      (?c B g B 𝗅B[f B d ?a]) B
                      𝖺B[?c, g, ?b B f B d ?a] B
                      (𝗅B-1[g] B ?b B f B d ?a) B
                      (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                      𝖺B[g, f, d ?a]"
            proof -
              have "𝗅B-1[(g B f) B d ?a] B ((g B f) B d ?a) =
                      𝗅-1[(g  f)  d ?a]  ((g  f)  d ?a)"
                using f g fg B.𝔩_ide_simp by auto
              also have "... = (?c0  𝖺-1[g, f, d ?a]) 
                                (?c0  g  𝗅[f  d ?a]) 
                                𝖺[?c0, g, ?b0  f  d ?a] 
                                (𝗅-1[g]  ?b0  f  d ?a) 
                                (g  𝖺[?b0, f, d ?a]  (𝗅-1[f]  d ?a)) 
                                𝖺[g, f, d ?a]"
                using f g fg by (intro EV.eval_eqI, auto)
              also have "... = (?c B 𝖺B-1[g, f, d ?a]) B
                               (?c B g B 𝗅B[f B d ?a]) B
                               𝖺B[?c, g, ?b B f B d ?a] B
                               (𝗅B-1[g] B ?b B f B d ?a) B
                               (g B 𝖺B[?b, f, d ?a] B (𝗅B-1[f] B d ?a)) B
                               𝖺B[g, f, d ?a]"
                using f g fg B.α_def B.α'.map_ide_simp B.VVV.ide_charSbC B.VVV.arr_charSbC
                      B.VV.ide_charSbC B.VV.arr_charSbC P_def B.𝔩_ide_simp
                by auto
              finally show ?thesis by blast
            qed
            thus ?thesis by simp
          qed
          finally show ?thesis by simp
        qed
        finally show "(d ?c B EoP.cmp (g, f)) B
                      𝖺B[d ?c, (E  P) g, (E  P) f] B
                      (ε1 g B (E  P) f) B
                      B.inv 𝖺B[IC.map g, d ?b, (E  P) f] B (IC.map g B ε1 f) B
                      𝖺B[IC.map g, IC.map f, d ?a]
                        = ε1 (g B f) B (IC.cmp (g, f) B d ?a)"
          by blast
      qed
      show "a. B.obj a 
                  (d a B EoP.unit a) B 𝗋B-1[d a] B 𝗅B[d a] = ε1 a B (IC.unit a B d a)"
      proof -
        fix a
        assume a: "B.obj a"
        interpret adjoint_equivalence_in_bicategory VB HB 𝖺B 𝗂B srcB trgB e a d a η a ε a
          using a chosen_adjoint_equivalence by simp
        have 0: "srcB a = a  trgB a = a"
          using a by auto
        have 1: "obj (P0 a)"
          using a P0_props(1) by simp
        have "(d a B EoP.unit a) B 𝗋B-1[d a] B 𝗅B[d a]
                = (d a B ((e a B 𝗅B-1[d a]) B B.inv (ε a)) B srcB (a B d (srcB a))) B
                  𝗋B-1[d a] B 𝗅B[d a]"
        proof -
          have "B.hseq (e (trgB a)) (a B d (srcB a))"
            using a by (elim B.objE, intro B.hseqI') auto
          moreover have "arr (srcB (d (srcB a)))"
            using a arr_charSbC P0_props(1) obj_char B.obj_simps(1) by auto
          moreover have "arr (srcB (a B d (srcB a)))"
            using a arr_charSbC P0_props(1) obj_char calculation(1) by fastforce
          moreover have "srcB (a B d (srcB a))  Obj  trgB (e (trgB a))  Obj"
            using a B.obj_simps P0_props obj_char arr_charSbC by simp
          moreover have "P0 a B P0 a  Obj"
            using 1 arr_charSbC P0_props(1) obj_char
            by (metis (no_types, lifting) B.cod_trg B.obj_def' B.trg.as_nat_trans.is_natural_2)
          moreover have "emb.unit (P0 (srcB a)) = P0 (srcB a)"
            using a 0 1 emb.unit_char' P.map0_def src_def by simp
          ultimately show ?thesis
            using a emb.map_def EoP.unit_char' prj_unit_char emb.unit_char' P.map0_def P_def
                  src_def arr_charSbC obj_char
            by simp
        qed
        also have "... = (d a B (e a B 𝗅B-1[d a]) B B.inv (ε a) B P0 a) B
                         𝗋B-1[d a] B 𝗅B[d a]"
          using a 1 B.comp_assoc B.obj_simps by auto
        also have "... = (d a B (e a B 𝗅B-1[d a]) B B.inv (ε a)) B 𝗋B-1[d a] B 𝗅B[d a]"
          using a B.comp_arr_dom by simp
        also have "... = (d a B e a B 𝗅B-1[d a]) B (d a B B.inv (ε a)) B 𝗋B-1[d a] B 𝗅B[d a]"
          using a B.whisker_left B.comp_assoc by simp
        also have "... = ((d a B e a B 𝗅B-1[d a]) B 𝖺B[d a, e a, d a]) B (η a B d a)"
          using a triangle_right B.comp_assoc
                B.invert_side_of_triangle(1) [of "𝗋B-1[d a] B 𝗅B[d a]" "d a B ε a"]
          by simp
        also have "... = 𝖺B[d a, e a, a B d a] B ((d a B e a) B 𝗅B-1[d a]) B (η a B d a)"
        proof -
          have "(d a B e a B 𝗅B-1[d a]) B 𝖺B[d a, e a, d a] =
                𝖺B[d a, e a, a B d a] B ((d a B e a) B 𝗅B-1[d a])"
            using a B.assoc_naturality [of "d a" "e a" "𝗅B-1[d a]"] by simp
          thus ?thesis
            using B.comp_assoc by simp
        qed
        also have "... = 𝖺B[d a, e a, a B d a] B (η a B 𝗅B-1[d a])"
          using a B.interchange [of "d a B e a" "η a" "𝗅B-1[d a]" "d a"]
                B.comp_arr_dom B.comp_cod_arr
          by simp
        also have "... = 𝖺B[d a, e a, a B d a] B (η a B a B d a) B 𝖺B[a, a, d a] B
                         (𝗅B-1[a] B d a)"
        proof -
          have "𝖺B[a, a, d a] B (𝗅B-1[a] B d a) = a B 𝗅B-1[d a]"
          proof -
            (* TODO: I wanted to prove this directly, but missed some necessary trick. *)
            have "𝖺B[a, a, d a] B (𝗅B-1[a] B d a) = 𝖺[a0, a0, d a]  (𝗅-1[a0]  d a)"
              using a ide_char B.α_def B.α'.map_ide_simp B.VVV.ide_charSbC B.VVV.arr_charSbC
                    B.VV.ide_charSbC B.VV.arr_charSbC P_def B.𝔩_ide_simp
              by auto
            also have "... = a0  𝗅-1[d a]"
              using a ide_charSbC by (intro EV.eval_eqI, auto)
            also have "... = a B 𝗅B-1[d a]"
              using a ide_char B.α_def B.α'.map_ide_simp B.VVV.ide_charSbC B.VVV.arr_charSbC
                    B.VV.ide_charSbC B.VV.arr_charSbC P_def B.𝔩_ide_simp
              by auto
            finally show ?thesis by blast
          qed
          hence "𝖺B[d a, e a, a B d a] B (η a B a B d a) B 𝖺B[a, a, d a] B
                 (𝗅B-1[a] B d a)
                   = 𝖺B[d a, e a, a B d a] B (η a B a B d a) B (a B 𝗅B-1[d a])"
            by simp
          also have "... = 𝖺B[d a, e a, a B d a] B (η a B 𝗅B-1[d a])"
            using a B.interchange [of "η a" a "a B d a" "𝗅B-1[d a]"] B.comp_arr_dom B.comp_cod_arr
            by simp
          finally show ?thesis by simp
        qed
        also have "... = ε1 a B (IC.unit a B d a)"
        proof -
          have "ε1 a B (IC.unit a B d a) =
                𝖺B[d a, e a, a B d a] B (η a B a B d a) B 𝗅B-1[a B d a] B (a B d a)"
            using a 0 counit1_def IC.unit_char' B.comp_assoc by simp
          also have "... = 𝖺B[d a, e a, a B d a] B (η a B a B d a) B 𝗅B-1[a B d a]"
          proof -
            have "B.arr 𝗅B-1[a B d a]"
              using a
              by (metis B.ide_hcomp B.lunit'_simps(1) B.objE equivalence_data_simpsB(8) ide_right)
            moreover have "B.dom 𝗅B-1[a B d a] = a B d a"
              using a
              by (metis B.ide_hcomp B.lunit'_simps(4) B.objE antipar(1) equivalence_data_simpsB(5)
                  ide_right)
            ultimately show ?thesis
              using a B.comp_arr_dom [of "𝗅B-1[a B d a]" "a B d a"] by simp
          qed
          also have "... = 𝖺B[d a, e a, a B d a] B (η a B a B d a) B 𝖺B[a, a, d a] B
                           (𝗅B-1[a] B d a)"
            using a B.lunit_hcomp(4) [of a "d a"] by auto
          finally have "ε1 a B (IC.unit a B d a) =
                        𝖺B[d a, e a, a B d a] B (η a B a B d a) B 𝖺B[a, a, d a] B
                        (𝗅B-1[a] B d a)"
            by simp
          thus ?thesis by simp
        qed
        finally show "(d a B EoP.unit a) B 𝗋B-1[d a] B 𝗅B[d a] = ε1 a B (IC.unit a B d a)"
          by blast
      qed
    qed

    interpretation equivalence_of_bicategories VB HB 𝖺B 𝗂B srcB trgB comp hcomp 𝖺 𝗂B src trg
                     E ΦE P ΦP unit0 unit1 counit0 counit1
      ..

    lemma induces_equivalence:
    shows "equivalence_of_bicategories VB HB 𝖺B 𝗂B srcB trgB comp hcomp 𝖺 𝗂B src trg
             E ΦE P ΦP unit0 unit1 counit0 counit1"
      ..

  end

  subsection "Equivalence Pseudofunctors, Bijective on Objects"

  text ‹
    Here we carry out the extension of an equivalence pseudofunctor F› to an equivalence
    of bicategories in the special case that the object map of F› is bijective.
    The bijectivity assumption simplifies the construction of the unit and counit of the
    equivalence (the components at objects are in fact identities), as well as the proofs
    of the associated coherence conditions.
  ›

  locale equivalence_pseudofunctor_bij_on_obj =
    equivalence_pseudofunctor +
  assumes bij_on_obj: "bij_betw map0 (Collect C.obj) (Collect D.obj)"
  begin

    abbreviation F0
    where "F0  map0"

    definition G0
    where "G0 = inv_into (Collect C.obj) F0"

    lemma G0_props:
    shows "D.obj b  C.obj (G0 b)  F0 (G0 b) = b"
    and "C.obj a  D.obj (F0 a)  G0 (F0 a) = a"
    proof -
      have surj: "F0 ` (Collect C.obj) = Collect D.obj"
        using bij_on_obj bij_betw_imp_surj_on by blast
      assume b: "D.obj b"
      show "C.obj (G0 b)  F0 (G0 b) = b"
      proof
        show "C.obj (G0 b)"
          unfolding G0_def
          using b by (metis inv_into_into mem_Collect_eq surj)
        show "F0 (G0 b) = b"
          unfolding G0_def
          using b by (simp add: f_inv_into_f surj)
      qed
      next
      have bij: "bij_betw G0 (Collect D.obj) (Collect C.obj)"
        using bij_on_obj G0_def bij_betw_inv_into by auto
      have surj: "G0 ` (Collect D.obj) = Collect C.obj"
        using bij_on_obj G0_def
        by (metis bij_betw_def bij_betw_inv_into)
      assume a: "C.obj a"
      show "D.obj (F0 a)  G0 (F0 a) = a"
        using a bij surj G0_def bij_betw_imp_inj_on bij_on_obj by force
    qed

    text ‹
      We extend G0 to all arrows of D› using chosen adjoint equivalences, which extend F›,
      between homC (a, a')› and homD (F a, F a')›.  The use of adjoint equivalences restricts
      choices that prevent us from validating the necessary coherence conditions.
    ›

    definition G
    where "G ν = (if D.arr ν then
                     equivalence_pseudofunctor_at_hom.G1 VC srcC trgC VD srcD trgD F
                       (G0 (srcD ν)) (G0 (trgD ν)) ν
                  else C.null)"

    lemma G_in_hom [intro]:
    assumes "D.arr ν"
    shows "«G ν : G0 (srcD ν) C G0 (trgD ν)»"
    and "«G ν : G (D.dom ν) C G (D.cod ν)»"
    proof -
      have 1: "srcD ν = F0 (G0 (srcD ν))  trgD ν = F0 (G0 (trgD ν))"
        using assms G0_props by simp
      interpret homC: subcategory VC λμ. «μ : G0 (srcD ν) C G0 (trgD ν)»
        using assms C.hhom_is_subcategory by simp
      interpret homD: subcategory VD λν'. «ν' : F0 (G0 (srcD ν)) D F0 (G0 (trgD ν))»
        using assms 1 D.in_hhom_def D.hhom_is_subcategory by simp
      interpret Faa': equivalence_pseudofunctor_at_hom
                        VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                        F Φ G0 (srcD ν) G0 (trgD ν)
        using assms G0_props(1) by unfold_locales auto
      have 2: "homD.arr ν"
        using assms 1 homD.arr_charSbC by simp
      show "«G ν : G (D.dom ν) C G (D.cod ν)»"
        unfolding G_def
        using assms 2 homC.arr_charSbC homC.dom_charSbC homC.cod_charSbC homD.dom_charSbC homD.cod_charSbC
              Faa'.ηε.F.preserves_arr Faa'.ηε.F.preserves_dom Faa'.ηε.F.preserves_cod
        by (intro C.in_homI) auto
      thus "C.in_hhom (G ν) (G0 (srcD ν)) (G0 (trgD ν))"
      proof -
        have "homC.arr (Faa'.G1 ν)"
          using 2 by simp
        thus ?thesis
          using G_def assms homC.arrE by presburger
      qed
    qed

    lemma G_simps [simp]:
    assumes "D.arr ν"
    shows "C.arr (G ν)"
    and "srcC (G ν) = G0 (srcD ν)" and "trgC (G ν) = G0 (trgD ν)"
    and "C.dom (G ν) = G (D.dom ν)" and "C.cod (G ν) = G (D.cod ν)"
      using assms G_in_hom by auto

    interpretation G: "functor" VD VC G
    proof
      show "f. ¬ D.arr f  G f = C.null"
        unfolding G_def by simp
      fix ν
      assume ν: "D.arr ν"
      show "C.arr (G ν)"
        using ν by simp
      show "C.dom (G ν) = G (D.dom ν)"
        using ν by simp
      show "C.cod (G ν) = G (D.cod ν)"
        using ν by simp
      next
      fix μ ν
      assume μν: "D.seq μ ν"
      have 1: "D.arr μ  D.arr ν  srcD μ = srcD ν  trgD μ = trgD ν 
               srcD μ = F0 (G0 (srcD μ))  trgD μ = F0 (G0 (trgD μ)) 
               srcD ν = F0 (G0 (srcD ν))  trgD ν = F0 (G0 (trgD ν))"
        using μν G0_props D.vseq_implies_hpar by auto
      interpret homC: subcategory VC λμ. «μ : G0 (srcD ν) C G0 (trgD ν)»
        using 1 C.hhom_is_subcategory by simp
      interpret homD: subcategory VD λν'. «ν' : F0 (G0 (srcD ν)) D F0 (G0 (trgD ν))»
        using 1 D.hhom_is_subcategory
        by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
      interpret homD': subcategory VD
                         λν'. D.arr ν'  srcD ν' = srcD ν  trgD ν' = trgD ν
        using 1 D.hhom_is_subcategory
        by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
      interpret Faa': equivalence_pseudofunctor_at_hom
                        VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                        F Φ G0 (srcD ν) G0 (trgD ν)
        using μν 1 G0_props(1) by unfold_locales auto
      have 2: "homD.seq μ ν"
        using μν 1 homD.arr_charSbC homD.comp_def by fastforce
      have "G (homD.comp μ ν) = homC.comp (G μ) (G ν)"
        unfolding G_def
        using 1 2 G0_props Faa'.G1_props homD.arr_charSbC
              Faa'.ηε.F.as_nat_trans.preserves_comp_2
        by auto
      thus "G (μ D ν) = G μ C G ν"
        using μν 1 2 G_def homC.comp_simp homD.comp_simp D.src_vcomp D.trg_vcomp
              Faa'.ηε.F.preserves_arr
        by metis
    qed

    lemma functor_G:
    shows "functor VD VC G"
      ..

    interpretation G: faithful_functor VD VC G
    proof
      fix f f'
      assume par: "D.par f f'"
      assume eq: "G f = G f'"
      have 1: "srcD f = srcD f'  trgD f = trgD f'"
        using par by (metis D.src_dom D.trg_dom)
      interpret homC: subcategory VC λμ. «μ : G0 (srcD f) C G0 (trgD f)»
        using par C.hhom_is_subcategory by simp
      interpret homD: subcategory VD λν'. «ν' : F0 (G0 (srcD f)) D F0 (G0 (trgD f))»
        using par D.hhom_is_subcategory
        by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
      interpret F: equivalence_pseudofunctor_at_hom
                        VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                        F Φ G0 (srcD f) G0 (trgD f)
        using par G0_props(1) by unfold_locales auto
      interpret F': equivalence_pseudofunctor_at_hom
                        VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                        F Φ G0 (srcD f') G0 (trgD f')
        using par G0_props(1) by unfold_locales auto
      have 2: "homD.par f f'"
        using par 1 homD.arr_charSbC homD.dom_charSbC homD.cod_charSbC G0_props by simp
      have "F.G1 f = F.G1 f'"
        using par eq G_def G_simps(2-3) by metis
      thus "f = f'"
        using F.ηε.F_is_faithful
        by (simp add: 2 faithful_functor_axioms_def faithful_functor_def)
    qed

    interpretation FG: composite_functor VD VC VD G F ..
    interpretation FG: faithful_functor VD VD "F o G"
      using faithful_functor_axioms G.faithful_functor_axioms faithful_functors_compose
      by blast

    interpretation GF: composite_functor VC VD VC F G ..
    interpretation GF: faithful_functor VC VC "G o F"
      using faithful_functor_axioms G.faithful_functor_axioms faithful_functors_compose
      by blast

    interpretation G: weak_arrow_of_homs VD srcD trgD VC srcC trgC G
    proof
      have *: "b. D.obj b  C.isomorphic (G b) (srcC (G b))"
      proof -
        fix b
        assume b: "D.obj b"
        interpret homC: subcategory VC λμ. «μ : G0 b C G0 b»
          using b C.hhom_is_subcategory by simp
        interpret homD: subcategory VD λν'. «ν' : F0 (G0 b) D F0 (G0 b)»
          using b D.hhom_is_subcategory
          by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
        interpret Faa': equivalence_pseudofunctor_at_hom
                          VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                          F Φ G0 b G0 b
          using b G0_props(1) by unfold_locales auto
        have 1: "C.in_hhom (G0 b) (G0 b) (G0 b)"
          using b G0_props by force

        text ‹
          Using the unit constraints of F› and the fact that F0 (G0 b) = b›,
          we obtain an isomorphism «?ψ : b ⇒D F (G0 b)»›, which is also
          an isomorphism in homD.
        ›

        let  = "unit (G0 b)"
        have ψ: "« : b D F (G0 b)»  D.iso "
          using b G0_props unit_char by auto
        have ψ_in_hhom: "D.in_hhom  b b"
          using b ψ 1 D.src_dom D.trg_dom by fastforce
        have 2: "homD.arr   homD.arr (D.inv )"
          by (metis D.in_hhomE D.inv_in_hhom G0_props(1) ψ ψ_in_hhom homD.arr_charSbC)
        have ψ': "homD.in_hom  b (Faa'.F1 (G0 b))"
        proof
          show "homD.arr "
            using 2 by simp
          show "homD.dom  = b"
            using 2 ψ homD.dom_charSbC by auto
          show "homD.cod  = Faa'.F1 (G0 b)"
          proof -
            have "« : b D Faa'.F1 (G0 b)»  D.iso "
              unfolding Faa'.F1_def using b ψ 1 G0_props by auto
            thus ?thesis
              using 2 ψ homD.cod_charSbC by auto
          qed
        qed

        text ‹
          Transposing ?ψ› via the adjunction ηε› yields an isomorphism from
          G b› to G0 b› in homC, hence in C›.
        ›

        have **: "homC.isomorphic (G b) (G0 b)"
        proof -
          have "homC.in_hom (Faa'.ηε.ψ (G0 b) ) (Faa'.G1 b) (G0 b)"
            by (metis "1" C.ide_src C.in_hhom_def Faa'.ηε.ψ_in_hom ψ' homC.ideISbC)
          moreover have "homC.iso (Faa'.ηε.ψ (G0 b) )"
          proof (unfold Faa'.ηε.ψ_def)
            have "homC.iso_in_hom (homC.comp (Faa'.ε (G0 b))
                    (Faa'.G1 )) (Faa'.G1 b) (G0 b)"
            proof (intro homC.comp_iso_in_hom)
              show "homC.iso_in_hom (Faa'.G1 (unit (G0 b)))
                      (Faa'.G1 b) (Faa'.ηε.FG.map (G0 b))"
                using ψ ψ' 1 2 homD.iso_charSbC Faa'.ηε.F.preserves_iso homD.iso_charSbC
                      Faa'.ηε.F.preserves_iso homC.in_hom_charSbC homC.arr_charSbC
                by auto
              show "homC.iso_in_hom (Faa'.ε (G0 b)) (Faa'.ηε.FG.map (G0 b)) (G0 b)"
                using 1 C.ide_src C.ide_trg C.in_hhom_def Faa'.ηε.ε.components_are_iso
                      Faa'.ηε.ε.preserves_hom homC.arrI homC.ideISbC homC.ide_in_hom
                      homC.map_simp homC.iso_in_hom_def
                by metis
            qed
            thus "homC.iso (homC.comp (Faa'.ε (G0 b)) (Faa'.G1 (unit (G0 b))))" by auto
          qed
          ultimately show ?thesis
            unfolding G_def
            using b homC.isomorphic_def D.obj_def D.obj_simps(3) by auto
        qed
        hence "homC.isomorphic (G b) (srcC (G b))"
          using b G_in_hom(1) by auto
        moreover have "f g. homC.isomorphic f g  C.isomorphic f g"
          using homC.iso_charSbC homC.isomorphic_def C.isomorphic_def homC.in_hom_charSbC homC.arr_charSbC
          by auto
        ultimately show "C.isomorphic (G b) (srcC (G b))"
          by simp
      qed
      fix ν
      assume ν: "D.arr ν"
      show "C.isomorphic (G (srcD ν)) (srcC (G ν))"
        using ν * by force
      show "C.isomorphic (G (trgD ν)) (trgC (G ν))"
        using ν * by force
    qed

    lemma weak_arrow_of_homs_G:
    shows "weak_arrow_of_homs VD srcD trgD VC srcC trgC G"
      ..

    sublocale HDoGG: composite_functor D.VV.comp C.VV.comp VC
                       G.FF λμν. fst μν C snd μν
      ..
    sublocale GoHD: composite_functor D.VV.comp VD VC λμν. fst μν D snd μν G
      ..

    text ‹
      To get the unit η› of the equivalence of bicategories, we piece together the
      units from the local equivalences.  The components at objects will in fact be identities.
    ›

    definition η
    where "η ν = (if D.arr ν then
                    equivalence_pseudofunctor_at_hom.η VC srcC trgC VD srcD trgD F
                      (G0 (srcD ν)) (G0 (trgD ν)) ν
                  else D.null)"

    lemma η_in_hom:
    assumes "D.arr ν"
    shows [intro]: "«η ν : srcD ν D trgD ν»"
    and [intro]: "«η ν : D.dom ν D F (G (D.cod ν))»"
    and "D.ide ν  D.iso (η ν)"
    proof -
      interpret homC: subcategory VC λμ. «μ : G0 (srcD ν) C G0 (trgD ν)»
        using assms C.hhom_is_subcategory by simp
      interpret homD: subcategory VD λν'. «ν' : F0 (G0 (srcD ν)) D F0 (G0 (trgD ν))»
        using assms D.hhom_is_subcategory
        by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
      interpret Faa': equivalence_pseudofunctor_at_hom VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                        F Φ G0 (srcD ν) G0 (trgD ν)
        using assms G0_props(1) by unfold_locales auto
      have 1: "«η ν : D.dom ν D F (G (D.cod ν))»  (D.ide ν  D.iso (η ν))"
      proof -
        have "srcD ν = F0 (G0 (srcD ν))  trgD ν = F0 (G0 (trgD ν))"
          using assms G0_props by simp
        hence "homD.arr ν"
          using assms homD.arr_charSbC homD.ide_char by simp
        hence "homD.in_hom (Faa'.η ν) (D.dom ν) (F (G (D.cod ν))) 
               (D.ide ν  homD.iso (Faa'.η ν))"
          using assms Faa'.ηε.η.preserves_hom [of ν "D.dom ν" "D.cod ν"]
                homD.map_def G_def Faa'.F1_def
          apply simp
          by (simp add: D.arr_iff_in_hom homD.arr_charSbC homD.cod_closed homD.dom_closed
              homD.ide_charSbC homD.in_hom_charSbC)
        thus ?thesis
          unfolding η_def
          using homD.in_hom_charSbC homD.iso_charSbC by auto
      qed
      show "«η ν : D.dom ν D F (G (D.cod ν))»"
        using 1 by simp
      thus "D.in_hhom (η ν) (srcD ν) (trgD ν)"
        using assms D.src_dom D.trg_dom
        by (metis D.arrI D.in_hhom_def D.vconn_implies_hpar(1-2))
      show "D.ide ν  D.iso (η ν)"
        using 1 by simp
    qed

    lemma η_simps [simp]:
    assumes "D.arr ν"
    shows "D.arr (η ν)"
    and "srcD (η ν) = srcD ν" and "trgD (η ν) = trgD ν"
    and "D.dom (η ν) = D.dom ν" and "D.cod (η ν) = F (G (D.cod ν))"
    and "D.ide ν  D.iso (η ν)"
      using assms η_in_hom by auto

    lemma η_naturality:
    assumes "D.arr ν"
    shows "η (D.cod ν) D ν = η ν" and "F (G ν) D η (D.dom ν) = η ν"
    and "ν D D.inv (η (D.dom ν)) = D.inv (η (D.cod ν)) D F (G ν)"
    proof -
      interpret homC: subcategory VC λμ'. «μ' : G0 (srcD ν) C G0 (trgD ν)»
        using assms C.hhom_is_subcategory by simp
      interpret homD: subcategory VD λν'. «ν' : F0 (G0 (srcD ν)) D F0 (G0 (trgD ν))»
        using D.hhom_is_subcategory by simp
      interpret Faa': equivalence_pseudofunctor_at_hom
                        VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                        F Φ G0 (srcD ν) G0 (trgD ν)
        using assms G0_props by unfold_locales auto
      show 1: "η (D.cod ν) D ν = η ν"
        unfolding η_def
        using assms Faa'.ηε.η.is_natural_2 homD.comp_char G_def homD.cod_simp
              G_in_hom(1) homC.arr_charSbC homD.arr_charSbC homD.cod_closed
        apply simp
        by (metis (no_types, lifting) D.ext Faa'.ηε.F.preserves_reflects_arr
            Faa'.ηε.η.preserves_reflects_arr)
      show 2: "F (G ν) D η (D.dom ν) = η ν"
        unfolding η_def
        using assms D.src_dom D.src_cod D.trg_dom D.trg_cod Faa'.ηε.η.is_natural_1
              homD.comp_char G_def Faa'.F1_def homD.dom_simp homD.cod_simp
        apply simp
        by (metis (no_types, lifting) D.not_arr_null Faa'.ηε.η.is_extensional
            η_def η_simps(1) homD.null_char)
      show "ν D D.inv (η (D.dom ν)) = D.inv (η (D.cod ν)) D F (G ν)"
        using assms 1 2
        by (simp add: D.invert_opposite_sides_of_square)
    qed

    text ‹
      The fact that G0 is chosen to be right-inverse to F› implies that η› is an
      ordinary natural isomorphism (with respect to vertical composition) from ID to FG›.
    ›

    interpretation η: natural_isomorphism VD VD D.map FG.map η
      using η_simps η_naturality η_def
      by unfold_locales auto

    text ‹
      In view of the bijectivity assumption, we can obtain the counit ε› the same way.
    ›

    definition ε
    where "ε μ = (if C.arr μ then
                    equivalence_pseudofunctor_at_hom.ε VC srcC trgC VD srcD trgD F
                      (srcC μ) (trgC μ) μ
                  else C.null)"

    lemma ε_in_hom:
    assumes "C.arr μ"
    shows [intro]: "«ε μ : srcC μ C trgC μ»"
    and [intro]: "«ε μ : G (F (C.dom μ)) C C.cod μ»"
    and "C.ide μ  C.iso (ε μ)"
    proof -
      interpret homC: subcategory VC λμ'. «μ' : srcC μ C trgC μ»
        using C.hhom_is_subcategory by simp
      interpret homD: subcategory VD λν'. «ν' : F0 (srcC μ) D F0 (trgC μ)»
        using assms D.hhom_is_subcategory
        by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
      interpret Faa': equivalence_pseudofunctor_at_hom
                        VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                        F Φ G0 (F0 (srcC μ)) G0 (F0 (trgC μ))
        using assms G0_props
        by unfold_locales simp_all
      have 1: "«ε μ : G (F (C.dom μ)) C C.cod μ»  (C.ide μ  C.iso (ε μ))"
      proof -
        have "homC.arr μ"
          using assms homC.arr_charSbC by simp
        hence "homC.in_hom μ (C.dom μ) (C.cod μ)"
          using assms homC.in_hom_charSbC homC.dom_charSbC homC.cod_charSbC by auto
        hence "homC.in_hom (Faa'.ε μ) (G (F (C.dom μ))) (C.cod μ) 
               (C.ide μ  homC.iso (Faa'.ε μ))"
          using assms Faa'.ηε.ε.preserves_hom [of μ "C.dom μ" "C.cod μ"]
                Faa'.ηε.ε.components_are_iso homC.map_def G_def Faa'.F1_def G0_props
          by (simp add: homC.ideISbC)
        thus ?thesis
          unfolding ε_def
          using assms homC.in_hom_charSbC homC.iso_charSbC G0_props(2) by simp
      qed
      show "«ε μ : G (F (C.dom μ)) C C.cod μ»"
        using 1 by simp
      thus "C.in_hhom (ε μ) (srcC μ) (trgC μ)"
        using assms C.src_cod C.trg_cod
        by (metis C.arrI C.in_hhom_def C.vconn_implies_hpar(1-4))
      show "C.ide μ  C.iso (ε μ)"
        using 1 by simp
    qed

    lemma ε_simps [simp]:
    assumes "C.arr μ"
    shows "C.arr (ε μ)"
    and "srcC (ε μ) = srcC μ" and "trgC (ε μ) = trgC μ"
    and "C.dom (ε μ) = G (F (C.dom μ))" and "C.cod (ε μ) = C.cod μ"
    and "C.ide μ  C.iso (ε μ)"
      using assms ε_in_hom by auto

    lemma ε_naturality:
    assumes "C.arr μ"
    shows "ε (C.cod μ) C G (F μ) = ε μ" and "μ C ε (C.dom μ) = ε μ"
    and "G (F μ) C C.inv (ε (C.dom μ)) = C.inv (ε (C.cod μ)) C μ"
    proof -
      interpret homC: subcategory VC λμ'. «μ' : srcC μ C trgC μ»
        using assms C.hhom_is_subcategory by simp
      interpret homD: subcategory VD λν. «ν : F0 (srcC μ) D F0 (trgC μ)»
        using assms D.hhom_is_subcategory
        by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def)
      interpret Faa': equivalence_pseudofunctor_at_hom
                        VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                        F Φ srcC μ trgC μ
        using assms by unfold_locales simp_all
      show 1: "ε (C.cod μ) C G (F μ) = ε μ"
        unfolding ε_def
        using assms C.src_dom C.src_cod C.trg_dom C.trg_cod Faa'.ηε.ε.is_natural_2
              homC.comp_char G_def Faa'.F1_def G0_props homC.dom_charSbC homC.cod_charSbC
        apply simp
        by (metis (no_types, lifting) C.in_hhomI Faa'.ηε.ε.preserves_reflects_arr homC.arr_charSbC
            homC.not_arr_null homC.null_char)
      show 2: "μ C ε (C.dom μ) = ε μ"
        unfolding ε_def
        using assms C.src_dom C.src_cod C.trg_dom C.trg_cod Faa'.ηε.ε.is_natural_1
              homC.comp_char G_def Faa'.F1_def G0_props homC.dom_charSbC homC.cod_charSbC
        apply simp
        by (metis (no_types, lifting) C.in_hhomI Faa'.ηε.ε.preserves_reflects_arr
            homC.arr_charSbC homC.not_arr_null homC.null_char)
      show "G (F μ) C C.inv (ε (C.dom μ)) = C.inv (ε (C.cod μ)) C μ"
        using assms 1 2 C.invert_opposite_sides_of_square by simp
    qed

    interpretation ε: natural_isomorphism VC VC GF.map C.map ε
      using ε_simps ε_naturality ε_def
      by unfold_locales auto

    interpretation GFG: composite_functor VD VC VC G GF.map ..
    interpretation FGF: composite_functor VC VD VD F FG.map ..

    interpretation Goη: natural_transformation VD VC G GFG.map G  η
    proof -
      have "G  D.map = G"
        using G.as_nat_trans.natural_transformation_axioms by auto
      moreover have "G  FG.map = GFG.map"
        by auto
      ultimately show "natural_transformation VD VC G GFG.map (G  η)"
        using η.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms
              horizontal_composite [of VD VD D.map FG.map η VC G G G] by simp
    qed

    interpretation ηoF: natural_transformation VC VD F FGF.map η  F
      using η.natural_transformation_axioms as_nat_trans.natural_transformation_axioms
            horizontal_composite [of VC VD F F F VD D.map FG.map η]
      by simp

    interpretation εoG: natural_transformation VD VC GFG.map G ε  G
      using ε.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms
            horizontal_composite [of VD VC G G G VC GF.map C.map ε]
      by simp

    interpretation Foε: natural_transformation VC VD FGF.map F F  ε
    proof -
      have "F  C.map = F"
        using as_nat_trans.natural_transformation_axioms by auto
      moreover have "F  GF.map = FGF.map"
        by auto
      ultimately show "natural_transformation VC VD FGF.map F (F  ε)"
        using ε.natural_transformation_axioms as_nat_trans.natural_transformation_axioms
              horizontal_composite [of VC VC GF.map C.map ε VD F F F]
        by simp
    qed

    interpretation εoG_Goη: vertical_composite VD VC G GFG.map G G  η ε  G ..
    interpretation Foε_ηoF: vertical_composite VC VD F FGF.map F η  F F  ε ..

    text ‹
      Bijectivity results in an ordinary adjunction between the vertical categories.
    ›

    lemma adjunction_ηε:
    shows "unit_counit_adjunction VC VD G F η ε"
    proof
      show "εoG_Goη.map = G"
      proof
        fix ν
        have "¬ D.arr ν  εoG_Goη.map ν = G ν"
          unfolding εoG_Goη.map_def
          by (simp add: G.is_extensional)
        moreover have "D.arr ν  εoG_Goη.map ν = G ν"
        proof -
          assume ν: "D.arr ν"
          interpret homC: subcategory VC λμ'. «μ' : G0 (srcD ν) C G0 (trgD ν)»
            using ν C.hhom_is_subcategory by simp
          interpret homD: subcategory VD λν'. «ν' : F0 (G0 (srcD ν)) D F0 (G0 (trgD ν))»
            using D.hhom_is_subcategory by simp
          interpret Faa': equivalence_pseudofunctor_at_hom
                            VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                            F Φ G0 (srcD ν) G0 (trgD ν)
            using ν G0_props by unfold_locales simp_all
          show "εoG_Goη.map ν = G ν"
          proof -
            have 1: "C.arr (Faa'.G1 (D.cod ν))"
              using ν homD.arr_charSbC G_def
              by (metis Faa'.ηε.F.preserves_reflects_arr G_in_hom(1) homC.arr_charSbC
                  homC.inclusion homD.cod_closed)
            have 2: "D.arr (Faa'.η ν)"
              using ν homD.arr_charSbC
              by (metis D.in_hhomI D.obj_src D.obj_trg Faa'.ηε.η.preserves_reflects_arr
                  G0_props(1) homD.inclusion)
            have 3: "srcC (Faa'.G1 (D.cod ν)) = G0 (srcD ν)"
              using ν
              by (metis D.arr_cod D.src_cod D.trg_cod G_def G_simps(2))
            have 4: "trgC (Faa'.G1 (D.cod ν)) = G0 (trgD ν)"
              using ν
              by (metis D.arr_cod D.src_cod D.trg_cod G_def G_simps(3))
            have 5: "homD.arr ν"
              using ν homD.arr_charSbC G0_props by simp
            have "εoG_Goη.map ν = ε (G (D.cod ν)) C G (η ν)"
              using ν εoG_Goη.map_def by simp
            also have "... = Faa'.ε (Faa'.G1 (D.cod ν)) C Faa'.G1 (Faa'.η ν)"
              using ν 1 2 3 4 η_def ε_def G_def G_simps(2-3) η_simps(2-3) by auto
            also have "... = homC.comp ((Faa'.ε  Faa'.G1) (homD.cod ν)) ((Faa'.G1  Faa'.η) ν)"
            proof -
              have "Faa'.ε (Faa'.G1 (D.cod ν)) C Faa'.G1 (Faa'.η ν) = Faa'.ηε.εFoFη.map ν"
              proof -
                have "C.seq (Faa'.ε (Faa'.G1 (homD.cod ν))) (Faa'.G1 (Faa'.η ν))"
                proof -
                  have "homC.seq (Faa'.ε (Faa'.G1 (homD.cod ν))) (Faa'.G1 (Faa'.η ν))"
                    using ν 5 homC.arr_charSbC homC.seq_charSbC Faa'.ηε.F.preserves_arr
                    apply (intro homC.seqI)
                      apply auto
                    using Faa'.ηε.ε.preserves_reflects_arr homC.inclusion homD.arr_cod
                    by presburger
                  thus ?thesis
                    using homC.seq_charSbC by simp
                qed
                moreover have "D.in_hhom (Faa'.η ν) (F0 (G0 (srcD ν))) (F0 (G0 (trgD ν)))"
                  using ν 5 Faa'.ηε.η.preserves_reflects_arr homD.arrE by blast
                moreover have "D.in_hhom (D.cod ν) (F0 (G0 (srcD ν))) (F0 (G0 (trgD ν)))"
                  using ν 5 homD.arrE homD.cod_closed by blast
                ultimately show ?thesis
                  using ν 5 Faa'.ηε.εFoFη.map_def homD.arr_charSbC homC.comp_char homD.cod_charSbC
                  by simp
              qed
              thus ?thesis
                using ν 5 Faa'.ηε.εFoFη.map_def by simp
            qed
            also have "... = G ν"
              using ν 5 G_def Faa'.ηε.triangle_F Faa'.ηε.εFoFη.map_simp_1 [of ν] by simp
            finally show ?thesis by simp
          qed
        qed
        ultimately show "εoG_Goη.map ν = G ν" by auto
      qed
      show "Foε_ηoF.map = F"
      proof
        fix μ
        have "¬ C.arr μ  Foε_ηoF.map μ = F μ"
          unfolding Foε_ηoF.map_def
          by (simp add: is_extensional)
        moreover have "C.arr μ  Foε_ηoF.map μ = F μ"
        proof -
          assume μ: "C.arr μ"
          interpret homC: subcategory VC λμ'. «μ' : srcC μ C trgC μ»
            using μ C.hhom_is_subcategory by simp
          interpret homD: subcategory VD λν. «ν : F0 (srcC μ) D F0 (trgC μ)»
            using μ D.in_hhom_def D.hhom_is_subcategory by simp
          interpret Faa': equivalence_pseudofunctor_at_hom
                            VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                            F Φ srcC μ trgC μ
            using μ by unfold_locales auto
          show "Foε_ηoF.map μ = F μ"
          proof -
            have 1: "homC.arr μ"
              using μ homC.arr_charSbC by simp
            have "Foε_ηoF.map μ = F (ε (C.cod μ)) D η (F μ)"
              using μ Foε_ηoF.map_def by simp
            also have "... = Faa'.F1 (Faa'.ε (C.cod μ)) D Faa'.η (Faa'.F1 μ)"
              unfolding η_def ε_def G_def
              using μ G0_props Faa'.F1_def
              by auto
            also have "... = homD.comp ((Faa'.F1  Faa'.ε) (homC.cod μ)) ((Faa'.η  Faa'.F1) μ)"
            proof -
              have 2: "homC.arr (C.cod μ)"
                using 1 homC.cod_charSbC [of μ] homC.arr_cod homC.dom_charSbC homC.cod_charSbC
                by simp
              moreover have "D.seq (Faa'.F1 (Faa'.ε (C.cod μ))) (Faa'.η (Faa'.F1 μ))"
              proof -
                have "homD.seq (Faa'.F1 (Faa'.ε (C.cod μ))) (Faa'.η (Faa'.F1 μ))"
                  using μ 1 2 homD.arr_charSbC homD.seq_charSbC Faa'.ηε.G.preserves_arr
                        homC.cod_simp homC.dom_cod
                  apply (intro homD.seqI)
                  by auto metis
                thus ?thesis
                  using homD.seq_charSbC by simp
              qed
              ultimately show ?thesis
                using 1 homD.comp_char homC.dom_charSbC homC.cod_charSbC by simp
            qed
            also have "... = Faa'.ηε.GεoηG.map μ"
              unfolding Faa'.ηε.GεoηG.map_def
              using 1 by simp
            also have "... = F μ"
              using μ Faa'.ηε.triangle_G homC.arr_charSbC Faa'.ηε.εFoFη.map_def
                    homD.arr_charSbC homC.comp_char homD.comp_char Faa'.F1_def
              by simp
            finally show ?thesis by simp
          qed
        qed
        ultimately show "Foε_ηoF.map μ = F μ" by auto
      qed
    qed

    interpretation ηε: unit_counit_adjunction VC VD G F η ε
      using adjunction_ηε by simp

    text ‹
      We now use the adjunction between the vertical categories to define the
      compositors for G›.  Without the bijectivity assumption, we would only obtain
      η› and ε› as pseudonatural equivalences, rather than natural isomorphisms,
      which would make everything more complicated.
    ›

    definition ΦG0
    where "ΦG0 f f' = C.inv (ε (G f C G f') C G (Φ (G f, G f') D (η f D η f')))"

    lemma ΦG0_in_hom:
    assumes "D.ide f" and "D.ide f'" and "srcD f = trgD f'"
    shows "«ΦG0 f f' : G f C G f' C G (f D f')»"
    proof (unfold ΦG0_def, intro C.inv_in_hom)
      show 1: "«ε (G f C G f') C G (Φ (G f, G f') D (η f D η f')) :
                  G (f D f') C G f C G f'»"
      proof
        show "«ε (G f C G f') : G (F (G f C G f')) C G f C G f'»"
          using assms ε_in_hom by auto
        show "«G (Φ (G f, G f') D (η f D η f')) : G (f D f') C G (F (G f C G f'))»"
        proof -
          have "«Φ (G f, G f') D (η f D η f') : f D f' D F (G f C G f')»"
            using assms η_in_hom
            by (intro D.comp_in_homI') auto
          thus ?thesis by auto
        qed
      qed
      show "C.iso (ε (G f C G f') C G (Φ (G f, G f') D (η f D η f')))"
      proof (intro C.isos_compose)
        show "C.iso (G (Φ (G f, G f') D (η f D η f')))"
        proof -
          have "D.iso (Φ (G f, G f') D (η f D η f'))"
            using assms η_in_hom
            by (intro D.isos_compose D.seqI) auto
          thus ?thesis by simp
        qed
        show "C.iso (ε (G f C G f'))"
          using assms ε_in_hom by simp
        show "C.seq (ε (G f C G f')) (G (Φ (G f, G f') D (η f D η f')))"
          using 1 by auto
      qed
    qed

    lemma iso_ΦG0:
    assumes "D.ide f" and "D.ide f'" and "srcD f = trgD f'"
    shows "C.iso (ΦG0 f f')"
    proof (unfold ΦG0_def, intro C.iso_inv_iso C.isos_compose)
      show 1: "C.iso (G (Φ (G f, G f') D (η f D η f')))"
        using assms η_in_hom D.in_hhom_def cmp_simps'(1) cmp_simps(4) by auto
      show "C.iso (ε (G f C G f'))"
        using assms ε_in_hom by simp
      show "C.seq (ε (G f C G f')) (G (Φ (G f, G f') D (η f D η f')))"
        using assms 1 by force
    qed

    lemma ΦG0_naturality:
    assumes "D.arr ν" and "D.arr ν'" and "srcD ν = trgD ν'"
    shows "ΦG0 (D.cod ν) (D.cod ν') C (G ν C G ν') =
           G (ν D ν') C ΦG0 (D.dom ν) (D.dom ν')"
    proof -
      let ?X = "ε (G (D.cod ν) C G (D.cod ν')) C
                G (Φ (G (D.cod ν), G (D.cod ν')) D (η (D.cod ν) D η (D.cod ν')))"
      have iso_X: "C.iso ?X"
        using assms C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC FF_def
        by (intro C.isos_compose) auto
      have "ΦG0 (D.cod ν) (D.cod ν') C (G ν C G ν') = C.inv ?X C (G ν C G ν')"
        unfolding ΦG0_def by simp
      also have "... = (C.inv (G (Φ (G (D.cod ν), G (D.cod ν')) D
                                 (η (D.cod ν) D η (D.cod ν')))) C
                        C.inv (ε (G (D.cod ν) C G (D.cod ν')))) C
                       (G ν C G ν')"
        using assms iso_X η_in_hom ε_in_hom
        by (simp add: C.inv_comp_left(1))
      also have "... = (C.inv (G (Φ (G (D.cod ν), G (D.cod ν'))) C
                               G (η (D.cod ν) D η (D.cod ν'))) C
                        C.inv (ε (G (D.cod ν) C G (D.cod ν')))) C
                       (G ν C G ν')"
      proof -
        have "G (Φ (G (D.cod ν), G (D.cod ν')) D (η (D.cod ν) D η (D.cod ν'))) =
              G (Φ (G (D.cod ν), G (D.cod ν'))) C G (η (D.cod ν) D η (D.cod ν'))"
          using assms iso_X η_in_hom by fast
        thus ?thesis by simp
      qed
      also have "... = C.inv (G (η (D.cod ν) D η (D.cod ν'))) C
                       C.inv (G (Φ (G (D.cod ν), G (D.cod ν')))) C
                       C.inv (ε (G (D.cod ν) C G (D.cod ν'))) C
                       (G ν C G ν')"
      proof -
        have "C.iso (G (Φ (G (D.cod ν), G (D.cod ν'))) C G (η (D.cod ν) D η (D.cod ν')))"
          using assms iso_X η_in_hom C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC FF_def
          by (intro C.isos_compose C.seqI) auto
        hence "C.inv (G (Φ (G (D.cod ν), G (D.cod ν'))) C
               G (η (D.cod ν) D η (D.cod ν')))
                 = C.inv (G (η (D.cod ν) D η (D.cod ν'))) C
                   C.inv (G (Φ (G (D.cod ν), G (D.cod ν'))))"
          using assms η_in_hom
          by (simp add: C.inv_comp_right(1))
        thus ?thesis
          using C.comp_assoc by simp
      qed
      also have "... = G (D.inv (η (D.cod ν)) D D.inv (η (D.cod ν'))) C
                       C.inv (G (Φ (G (D.cod ν), G (D.cod ν')))) C
                       C.inv (ε (G (D.cod ν) C G (D.cod ν'))) C
                       (G ν C G ν')"
      proof -
        have "C.inv (G (η (D.cod ν) D η (D.cod ν'))) =
                G (D.inv (η (D.cod ν)) D D.inv (η (D.cod ν')))"
          using assms η_in_hom D.ide_cod D.iso_hcomp D.src_cod D.trg_cod G.preserves_inv
                η_simps(2-3) D.inv_hcomp
          by (metis D.arr_cod)
        thus ?thesis by simp
      qed
      also have "... = G (D.inv (η (D.cod ν)) D D.inv (η (D.cod ν'))) C
                       (C.inv (G (Φ (G (D.cod ν), G (D.cod ν')))) C
                       G (F (G ν C G ν'))) C
                       C.inv (ε (G (D.dom ν) C G (D.dom ν')))"
      proof -
        have "C.inv (ε (G (D.cod ν) C G (D.cod ν'))) C (G ν C G ν')
                = G (F (G ν C G ν')) C C.inv (ε (G (D.dom ν) C G (D.dom ν')))"
        proof -
          have "G (D.dom ν) C G (D.dom ν') = C.dom (G ν C G ν')"
            by (simp add: assms)
          thus ?thesis
            using assms C.hcomp_simps(4) C.hseqI' G.preserves_cod G.preserves_hseq
                  ε_naturality(3)
            by presburger
        qed
        thus ?thesis
          using C.comp_assoc by simp
      qed
      also have "... = (G (D.inv (η (D.cod ν)) D D.inv (η (D.cod ν'))) C
                        G (F (G ν) D F (G ν'))) C
                       G (D.inv (Φ (C.dom (G ν), C.dom (G ν')))) C
                       C.inv (ε (G (D.dom ν) C G (D.dom ν')))"
      proof -
        have "C.inv (G (Φ (G (D.cod ν), G (D.cod ν')))) C G (F (G ν C G ν')) =
              G (D.inv (Φ (C.cod (G ν), C.cod (G ν'))) D F (G ν C G ν'))"
          using assms by (simp add: G.preserves_inv)
        also have "... = G ((F (G ν) D F (G ν')) D D.inv (Φ (C.dom (G ν), C.dom (G ν'))))"
          using assms C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC FF_def
                Φ.inv_naturality [of "(G ν, G ν')"]
          by simp
        also have "... = G (F (G ν) D F (G ν')) C
                         G (D.inv (Φ (C.dom (G ν), C.dom (G ν'))))"
          using assms by simp
        finally show ?thesis
          using C.comp_assoc by simp
      qed
      also have "... = G (ν D ν') C
                       G (D.inv (η (D.dom ν)) D D.inv (η (D.dom ν'))) C
                       G (D.inv (Φ (C.dom (G ν), C.dom (G ν')))) C
                       C.inv (ε (G (D.dom ν) C G (D.dom ν')))"
      proof -
        have "G (D.inv (η (D.cod ν)) D D.inv (η (D.cod ν'))) C G (F (G ν) D F (G ν')) =
                G ((D.inv (η (D.cod ν)) D D.inv (η (D.cod ν'))) D (F (G ν) D F (G ν')))"
          using assms by simp
        also have "... = G (D.inv (η (D.cod ν)) D F (G ν)
                           D D.inv (η (D.cod ν')) D F (G ν'))"
          using assms D.interchange by simp
        also have "... = G (ν D D.inv (η (D.dom ν)) D ν' D D.inv (η (D.dom ν')))"
          using assms η_naturality by simp
        also have "... = G ((ν D ν') D (D.inv (η (D.dom ν)) D D.inv (η (D.dom ν'))))"
          using assms D.interchange by simp
        also have "... = G (ν D ν') C G (D.inv (η (D.dom ν)) D D.inv (η (D.dom ν')))"
          using assms by simp
        finally show ?thesis
          using C.comp_assoc by simp
      qed
      also have "... = G (ν D ν') C ΦG0 (D.dom ν) (D.dom ν')"
      proof -
        have "G (D.inv (η (D.dom ν)) D D.inv (η (D.dom ν'))) C
              G (D.inv (Φ (C.dom (G ν), C.dom (G ν')))) C
              C.inv (ε (G (D.dom ν) C G (D.dom ν')))
                = C.inv (G (η (D.dom ν) D η (D.dom ν'))) C
                  C.inv (G (Φ (C.dom (G ν), C.dom (G ν')))) C
                  C.inv (ε (G (D.dom ν) C G (D.dom ν')))"
        proof -
          have "G (D.inv (η (D.dom ν)) D D.inv (η (D.dom ν'))) =
                  C.inv (G (η (D.dom ν) D η (D.dom ν')))"
            using assms cmp_components_are_iso
            by (metis D.arr_dom D.ide_dom D.inv_hcomp D.iso_hcomp D.src_dom D.trg_dom
                      G.preserves_inv η_simps(2) η_simps(3) η_simps(6))
          moreover have "G (D.inv (Φ (C.dom (G ν), C.dom (G ν')))) =
                           C.inv (G (Φ (C.dom (G ν), C.dom (G ν'))))"
            using assms cmp_components_are_iso
            by (simp add: G.preserves_inv)
          ultimately show ?thesis by simp
        qed
        also have "... = C.inv (ε (G (D.dom ν) C G (D.dom ν')) C
                         G (Φ (C.dom (G ν), C.dom (G ν'))) C
                         G (η (D.dom ν) D η (D.dom ν')))"
        proof -
          have "C.iso (ε (G (D.dom ν) C G (D.dom ν')) C
                       G (Φ (C.dom (G ν), C.dom (G ν'))) C
                       G (η (D.dom ν) D η (D.dom ν')))"
            using assms η_in_hom ε_in_hom cmp_simps'(1,4) C.VV.cod_charSbC
            by (intro C.isos_compose C.seqI) auto
          thus ?thesis
            using assms cmp_components_are_iso C.comp_assoc C.inv_comp_left by simp
        qed
        also have "... = C.inv (ε (G (D.dom ν) C G (D.dom ν')) C
                                G (Φ (C.dom (G ν), C.dom (G ν')) D
                                (η (D.dom ν) D η (D.dom ν'))))"
        proof -
          have "G (Φ (C.dom (G ν), C.dom (G ν'))) C G (η (D.dom ν) D η (D.dom ν')) =
                  G (Φ (C.dom (G ν), C.dom (G ν')) D (η (D.dom ν) D η (D.dom ν')))"
            using assms cmp_simps'(1,4) by auto
          thus ?thesis by simp
        qed
        also have "... = ΦG0 (D.dom ν) (D.dom ν')"
          unfolding ΦG0_def
          using assms by simp
        finally show ?thesis by simp
      qed
      finally show ?thesis by simp
    qed

    interpretation ΦG: transformation_by_components D.VV.comp VC HDoGG.map GoHD.map
                         λfg. ΦG0 (fst fg) (snd fg)
      using D.VV.ide_charSbC D.VV.arr_charSbC ΦG0_in_hom ΦG0_naturality G.FF_def
            D.VV.dom_charSbC D.VV.cod_charSbC
      by unfold_locales auto

    interpretation ΦG: natural_isomorphism D.VV.comp VC HDoGG.map GoHD.map ΦG.map
      using ΦG.map_simp_ide D.VV.ide_charSbC D.VV.arr_charSbC iso_ΦG0 by unfold_locales simp

    abbreviation ΦG
    where "ΦG  ΦG.map"

    lemma ΦG_in_hom [intro]:
    assumes "D.arr ν" and "D.arr ν'" and "srcD ν = trgD ν'"
    shows "C.in_hhom (ΦG.map (ν, ν')) (srcC (G (D.dom ν'))) (trgC (G (D.cod ν)))"
    and "«ΦG.map (ν, ν') : G (D.dom ν) C G (D.dom ν') C G (D.cod ν D D.cod ν')»"
    proof -
      show "«ΦG.map (ν, ν') : G (D.dom ν) C G (D.dom ν') C G (D.cod ν D D.cod ν')»"
      proof -
        have "ΦG.map (ν, ν') = ΦG0 (D.cod ν) (D.cod ν') C (G ν C G ν')"
          using assms D.VV.arr_charSbC ΦG.map_def ΦG0_in_hom G.FF_def D.VV.cod_charSbC
          by simp
        moreover have "«ΦG0 (D.cod ν) (D.cod ν') C (G ν C G ν') :
                         G (D.dom ν) C G (D.dom ν') C G (D.cod ν D D.cod ν')»"
          using assms ΦG0_in_hom by (intro C.comp_in_homI) auto
        ultimately show ?thesis by simp
      qed
      thus "C.in_hhom (ΦG.map (ν, ν')) (srcC (G (D.dom ν'))) (trgC (G (D.cod ν)))"
        using assms C.vconn_implies_hpar(1-2) by auto
    qed

    lemma ΦG_simps [simp]:
    assumes "D.arr ν" and "D.arr ν'" and "srcD ν = trgD ν'"
    shows "C.arr (ΦG.map (ν, ν'))"
    and "srcC (ΦG.map (ν, ν')) = srcC (G (D.dom ν'))"
    and "trgC (ΦG.map (ν, ν')) = trgC (G (D.cod ν))"
    and "C.dom (ΦG.map (ν, ν')) = G (D.dom ν) C G (D.dom ν')"
    and "C.cod  (ΦG.map (ν, ν')) = G (D.cod ν D D.cod ν')"
      using assms ΦG_in_hom
          apply auto
      by fast+

    lemma ΦG_map_simp_ide:
    assumes "D.ide f" and "D.ide f'" and "srcD f = trgD f'"
    shows "ΦG.map (f, f') = G (D.inv (η f) D D.inv (η f')) C G (D.inv (Φ (G f, G f'))) C
                            C.inv (ε (G f C G f'))"
    proof -
      have "ΦG.map (f, f') = C.inv (ε (G f C G f') C G (Φ (G f, G f') D (η f D η f')))"
        using assms ΦG.map_simp_ide D.VV.ide_charSbC D.VV.arr_charSbC ΦG0_def G.FF_def by auto
      also have "... = C.inv (ε (G f C G f') C G (Φ (G f, G f')) C G (η f D η f'))"
        using assms D.VV.ide_charSbC D.VV.arr_charSbC cmp_simps(1,4) by simp
      also have "... = C.inv (G (η f D η f')) C C.inv (G (Φ (G f, G f'))) C
                       C.inv (ε (G f C G f'))"
      proof -
        have "C.iso (ε (G f C G f') C G (Φ (G f, G f')) C G (η f D η f'))"
          using assms C.VV.ide_charSbC C.VV.arr_charSbC FF_def
          by (intro C.isos_compose C.seqI) auto
        thus ?thesis
          using assms C.inv_comp_left(1-2) cmp_components_are_iso C.comp_assoc by simp
      qed
      also have "... = G (D.inv (η f) D D.inv (η f')) C G (D.inv (Φ (G f, G f'))) C
                         C.inv (ε (G f C G f'))"
        using assms cmp_components_are_iso D.ideD(1) D.inv_hcomp D.iso_hcomp
              G.preserves_ide G.preserves_inv G_simps(2-3) η_simps(2-3,6)
        by metis
      finally show ?thesis by blast
    qed

    lemma η_hcomp:
    assumes "D.ide f" and "D.ide f'" and "srcD f = trgD f'"
    shows "η (f D f') = F (ΦG.map (f, f')) D Φ (G f, G f') D (η f D η f')"
    proof -
      have 1: "«F (ΦG.map (f, f')) D Φ (G f, G f') D (η f D η f') :
                  f D f' D F (G (f D f'))»"
        using assms by fastforce
      have 2: "«ε (G f C G f') C G (Φ (G f, G f')) C G (η f D η f') :
                  G (f D f') C G f C G f'»"
        using assms G.preserves_hom cmp_in_hom(2)
        by (intro C.comp_in_homI) auto
      have "F (ΦG.map (f, f')) D Φ (G f, G f') D (η f D η f')
              = F (C.inv (ε (G f C G f') C G (Φ (G f, G f') D (η f D η f')))) D
                  Φ (G f, G f') D (η f D η f')"
        using assms ΦG.map_simp_ide D.VV.ide_charSbC D.VV.arr_charSbC ΦG0_def by simp
      also have "... = F (C.inv (G (η f D η f')) C C.inv (G (Φ (G f, G f'))) C
                          C.inv (ε (G f C G f'))) D
                       Φ (G f, G f') D (η f D η f')"
      proof -
        have "C.inv (ε (G f C G f') C G (Φ (G f, G f') D (η f D η f'))) =
              C.inv (ε (G f C G f') C G (Φ (G f, G f')) C G (η f D η f'))"
          using assms 1 by force
        also have "... = C.inv (G (η f D η f')) C C.inv (G (Φ (G f, G f'))) C
                           C.inv (ε (G f C G f'))"
        proof -
          have "C.iso (ε (G f C G f') C G (Φ (G f, G f')) C G (η f D η f'))"
            using assms 2 by (intro C.isos_compose) auto
          thus ?thesis
            using assms 1 C.inv_comp_left C.comp_assoc by simp
        qed
        finally show ?thesis by simp
      qed
      also have "... = F (C.inv (G (η f D η f'))) D
                         F (C.inv (G (Φ (G f, G f')))) D F (C.inv (ε (G f C G f'))) D
                         Φ (G f, G f') D (η f D η f')"
      proof -
        have "C.arr ((C.inv (G (η f D η f')) C C.inv (G (Φ (G f, G f'))) C
                               C.inv (ε (G f C G f'))))"
          using assms 1 2 cmp_components_are_iso C.VV.ide_charSbC C.VV.arr_charSbC FF_def
          by (intro C.seqI) auto
        thus ?thesis
          using C.inv_comp_left D.comp_assoc by auto
      qed
      also have "... = D.inv (F (G (η f D η f'))) D
                       D.inv (F (G (Φ (G f, G f')))) D
                       D.inv (F (ε (G f C G f'))) D
                       Φ (G f, G f') D
                       (η f D η f')"
        using assms by (simp add: preserves_inv)
      also have "... = D.inv ((η (F (G f) D F (G f'))) D (η f D η f') D
                         D.inv (η (f D f'))) D
                       D.inv (η (F (G f C G f')) D Φ (G f, G f') D
                              D.inv (η (F (G f) D F (G f')))) D
                       D.inv (F (ε (G f C G f'))) D
                       Φ (G f, G f') D
                       (η f D η f')"
      proof -
        have *: "ν. D.arr ν  η (D.cod ν) D ν D D.inv (η (D.dom ν)) = F (G ν)"
          by (metis (full_types) D.arr_dom D.comp_assoc D.ide_dom D.invert_side_of_triangle(2)
              η_naturality(1-2) η_simps(1,6))
        have "F (G (η f D η f')) =
              η (F (G f) D F (G f')) D (η f D η f') D D.inv (η (f D f'))"
          using assms * [of "η f D η f'"] by simp
        moreover have "F (G (Φ (G f, G f'))) =
                         η (F (G f C G f')) D Φ (G f, G f') D D.inv (η (F (G f) D F (G f')))"
          using assms 1 * [of "Φ (G f, G f')"] cmp_simps(5) by fastforce
        ultimately show ?thesis by simp
      qed
      also have "... = η (f D f') D
                       D.inv (η f D η f') D
                       ((D.inv (η (F (G f) D F (G f')))) D (η (F (G f) D F (G f')))) D
                       (D.inv (Φ (G f, G f')) D
                       ((D.inv (η (F (G f C G f')))) D D.inv (F (ε (G f C G f')))) D
                       Φ (G f, G f')) D
                       (η f D η f')"
      proof -
        have "D.inv ((η (F (G f) D F (G f'))) D (η f D η f') D D.inv (η (f D f')))
                = η (f D f') D D.inv (η f D η f') D D.inv (η (F (G f) D F (G f')))"
        proof -
          have "D.iso (η (F (G f) D F (G f')) D (η f D η f') D D.inv (η (f D f')))"
            using assms by (intro D.isos_compose) auto
          thus ?thesis
            using assms D.inv_comp_left D.comp_assoc by simp
        qed
        moreover have
          "D.inv (η (F (G f C G f')) D Φ (G f, G f') D D.inv (η (F (G f) D F (G f')))) =
           η (F (G f) D F (G f')) D D.inv (Φ (G f, G f')) D D.inv (η (F (G f C G f')))"
        proof -
          have "D.iso (η (F (G f C G f')) D Φ (G f, G f') D D.inv (η (F (G f) D F (G f'))))"
          proof -
            have 1: "D.seq (Φ (G f, G f')) (D.inv (η (F (G f) D F (G f'))))"
              using assms by (intro D.seqI) auto
            moreover have "D.seq (η (F (G f C G f')))
                                 (Φ (G f, G f') D D.inv (η (F (G f) D F (G f'))))"
              using assms 1 by auto
            ultimately show ?thesis
              using assms cmp_components_are_iso η_in_hom
                by (intro D.isos_compose) auto
          qed
          thus ?thesis
            using assms D.comp_inv_arr' D.comp_assoc D.inv_comp_left cmp_components_are_iso
            by auto
        qed
        ultimately show ?thesis
          using D.comp_assoc by simp
      qed
      also have "... =  η (f D f') D
                        D.inv (η f D η f') D
                        (D.inv (η (F (G f) D F (G f'))) D (η (F (G f) D F (G f')))) D
                        (D.inv (Φ (G f, G f')) D Φ (G f, G f')) D
                        (η f D η f')"
      proof -
        have "(D.inv (η (F (G f C G f')))) D D.inv (F (ε (G f C G f'))) D Φ (G f, G f')
                = F (G f C G f') D Φ (G f, G f')"
        proof -
          have "(D.inv (η (F (G f C G f')))) D D.inv (F (ε (G f C G f')))
                  = D.inv (F (ε (G f C G f')) D η (F (G f C G f')))"
            using assms by (simp add: D.inv_comp)
          also have "... = F (G f C G f')"
            using assms ηε.triangle_G Foε_ηoF.map_simp_ide by fastforce
          finally show ?thesis using D.comp_assoc by metis
        qed
        also have "... = Φ (G f, G f')"
          using assms D.comp_cod_arr cmp_in_hom(2) [of "G f" "G f'"] by auto
        finally have "(D.inv (η (F (G f C G f')))) D D.inv (F (ε (G f C G f'))) D
                      Φ (G f, G f')
                        = Φ (G f, G f')"
          by blast
        thus ?thesis
          using D.comp_assoc by simp
      qed
      also have "... = η (f D f') D
                       D.inv (η f D η f') D
                       (D.inv (η (F (G f) D F (G f'))) D (η (F (G f) D F (G f')))) D
                       (η f D η f')"
        using assms D.comp_cod_arr D.comp_inv_arr' cmp_components_are_iso by simp
      also have "... = η (f D f') D D.inv (η f D η f') D (η f D η f')"
        using assms η_in_hom D.comp_inv_arr' D.comp_cod_arr by simp
      also have "... = η (f D f')"
          using assms D.comp_inv_arr' [of "η f D η f'"] D.comp_arr_dom by simp
      finally show ?thesis by simp
    qed

    lemma ε_hcomp:
    assumes "C.ide g" and "C.ide g'" and "srcC g = trgC g'"
    shows "ε (g C g') = (ε g C ε g') C C.inv (ΦG.map (F g, F g')) C C.inv (G (Φ (g, g')))"
    proof -
      have 1: "«ε (G (F g) C G (F g')) C G (Φ (G (F g), G (F g')) D (η (F g) D η (F g'))) :
                 G (F g D F g') C G (F g) C G (F g')»"
        using assms by fastforce
      have "(ε g C ε g') C C.inv (ΦG.map (F g, F g')) C C.inv (G (Φ (g, g')))
              = (ε g C ε g') C
                C.inv (C.inv (ε (G (F g) C G (F g')) C G (Φ (G (F g), G (F g')) D
                                (η (F g) D η (F g'))))) C
                C.inv (G (Φ (g, g')))"
        using assms ΦG.map_simp_ide D.VV.ide_charSbC D.VV.arr_charSbC ΦG0_def by simp
      also have "... = ((ε g C ε g') C ε (G (F g) C G (F g'))) C
                       G (Φ (G (F g), G (F g')) D (η (F g) D η (F g'))) C
                       C.inv (G (Φ (g, g')))"
      proof -
        have "C.iso (ε (G (F g) C G (F g')) C G (Φ (G (F g), G (F g')) D
              (η (F g) D η (F g'))))"
          using assms 1
          by (intro C.isos_compose D.isos_compose G.preserves_iso) auto
        thus ?thesis
          using C.comp_assoc by simp
      qed
      also have "... = ε (g C g') C
                       (G (F (ε g C ε g')) C G (Φ (G (F g), G (F g')) D
                         (η (F g) D η (F g')))) C
                       C.inv (G (Φ (g, g')))"
      proof -
        have "(ε g C ε g') C ε (G (F g) C G (F g')) = ε (g C g') C G (F (ε g C ε g'))"
          using assms ε_naturality [of "ε g C ε g'"] by simp
        thus ?thesis
          using C.comp_assoc by simp
      qed
      also have "... = ε (g C g') C
                       G ((F (ε g C ε g') D Φ (G (F g), G (F g'))) D (η (F g) D η (F g'))) C
                       C.inv (G (Φ (g, g')))"
        using assms C.VV.ide_charSbC C.VV.arr_charSbC FF_def D.comp_assoc by auto
      also have "... = ε (g C g') C
                       G (Φ (g, g') D (F (ε g) D F (ε g')) D (η (F g) D η (F g'))) C
                       C.inv (G (Φ (g, g')))"
      proof -
        have "F (ε g C ε g') D Φ (G (F g), G (F g')) = Φ (g, g') D (F (ε g) D F (ε g'))"
          using assms C.VV.arr_charSbC D.VV.ide_charSbC D.VV.arr_charSbC FF_def
                Φ.naturality [of "(ε g, ε g')"] C.VV.dom_charSbC C.VV.cod_charSbC
          by simp
        thus ?thesis
          using D.comp_assoc by simp
      qed
      also have "... = ε (g C g') C
                       G (Φ (g, g') D (F (ε g) D η (F g) D F (ε g') D η (F g'))) C
                       C.inv (G (Φ (g, g')))"
        using assms D.interchange by simp
      also have "... = ε (g C g') C G (Φ (g, g') D (F g D F g')) C C.inv (G (Φ (g, g')))"
        using assms ηε.triangle_G Foε_ηoF.map_def [of g] Foε_ηoF.map_def [of g'] by simp
      also have "... = ε (g C g') C G (Φ (g, g')) C C.inv (G (Φ (g, g')))"
        using assms D.comp_arr_dom cmp_simps(1) cmp_simps(4) by auto
      also have "... = ε (g C g')"
        using assms D.comp_arr_dom C.comp_arr_inv' C.VV.ide_charSbC C.VV.arr_charSbC
              cmp_components_are_iso C.comp_arr_dom
        by simp
      finally show ?thesis by simp
    qed

    lemma G_preserves_hcomp:
    assumes "D.hseq ν ν'"
    shows "G (ν D ν') = ΦG.map (D.cod ν, D.cod ν') C (G ν C G ν') C
                         C.inv (ΦG.map (D.dom ν, D.dom ν'))"
    proof -
      have "G (ν D ν') C ΦG.map (D.dom ν, D.dom ν') =
            ΦG.map (D.cod ν, D.cod ν') C (G ν C G ν')"
        using assms ΦG.naturality D.VV.arr_charSbC D.VV.cod_charSbC G.FF_def D.VV.dom_charSbC by auto
      moreover have "C.seq (ΦG.map (D.cod ν, D.cod ν')) (G ν C G ν')"
      proof
        show "«G ν C G ν' :
                 G (D.dom ν) C G (D.dom ν') C G (D.cod ν) C G (D.cod ν')»"
          using assms G_in_hom(2) C.hcomp_in_vhom by auto
        show "«ΦG.map (D.cod ν, D.cod ν') :
                 G (D.cod ν) C G (D.cod ν') C G (D.cod ν D D.cod ν')»"
          using assms ΦG0_in_hom ΦG.map_simp_ide D.VV.ide_charSbC D.VV.arr_charSbC by auto
      qed
      moreover have "C.iso (ΦG.map (D.dom ν, D.dom ν'))"
        using assms ΦG.components_are_iso D.VV.ide_charSbC D.VV.arr_charSbC by auto
      ultimately show ?thesis
        using assms ΦG.components_are_iso C.comp_assoc
              C.invert_side_of_triangle(2)
                [of "ΦG.map (D.cod ν, D.cod ν') C (G ν C G ν')"
                    "G (ν D ν')" "ΦG.map (D.dom ν, D.dom ν')"]
        by simp
    qed

    lemma coherence_LHS:
    assumes "D.ide f" and "D.ide g" and "D.ide h"
    and "srcD f = trgD g" and "srcD g = trgD h"
    shows "F (G 𝖺D[f, g, h] C ΦG.map (f D g, h) C (ΦG.map (f, g) C G h))
             = (η (f D g D h) D (D.inv (η f) D D.inv (η g) D D.inv (η h))) D
               𝖺D[F (G f), F (G g), F (G h)] D
               (D.inv (Φ (G f, G g)) D F (G h)) D D.inv (Φ (G f C G g, G h))"
    proof -
      have "F (G 𝖺D[f, g, h] C ΦG.map (f D g, h) C (ΦG.map (f, g) C G h))
              = F (G 𝖺D[f, g, h] C
                (G (D.inv (η (f D g)) D D.inv (η h))) C
                G (D.inv (Φ (G (f D g), G h))) C
                C.inv (ε (G (f D g) C G h)) C
                (G (D.inv (η f) D D.inv (η g)) C
                G (D.inv (Φ (G f, G g))) C
                C.inv (ε (G f C G g)) C G h))"
        using assms ΦG_map_simp_ide C.comp_assoc by simp
      also have "... = F (G 𝖺D[f, g, h] C
                       (G (D.inv (η (f D g)) D D.inv (η h))) C
                       G (D.inv (Φ (G (f D g), G h))) C
                       C.inv (ε (G (f D g) C G h)) C
                       (G (D.inv (η f) D D.inv (η g)) C G h) C
                       (G (D.inv (Φ (G f, G g))) C G h) C
                       (C.inv (ε (G f C G g)) C G h))"
        using assms C.whisker_right D.comp_assoc by simp
      also have "... = F (G 𝖺D[f, g, h]) D
                       F (G (D.inv (η (f D g)) D D.inv (η h))) D
                       F (G (D.inv (Φ (G (f D g), G h)))) D
                       F (C.inv (ε (G (f D g) C G h))) D
                       F (G (D.inv (η f) D D.inv (η g)) C G h) D
                       F (G (D.inv (Φ (G f, G g))) C G h) D
                       F (C.inv (ε (G f C G g)) C G h)"
      proof -
        have "C.arr (G 𝖺D[f, g, h] C
                     (G (D.inv (η (f D g)) D D.inv (η h))) C
                     G (D.inv (Φ (G (f D g), G h))) C
                     C.inv (ε (G (f D g) C G h)) C
                     (G (D.inv (η f) D D.inv (η g)) C G h) C
                     (G (D.inv (Φ (G f, G g))) C G h) C
                     (C.inv (ε (G f C G g)) C G h))"
          using assms ΦG_simps G.FF_def G0_props by auto
        thus ?thesis
          by (metis C.seqE as_nat_trans.preserves_comp_2)
      qed
      also have "... = F (G 𝖺D[f, g, h]) D
                       F (G (D.inv (η (f D g)) D D.inv (η h))) D
                       F (G (D.inv (Φ (G (f D g), G h)))) D
                       F (C.inv (ε (G (f D g) C G h))) D
                       (Φ (G (f D g), G h) D
                       (F (G (D.inv (η f) D D.inv (η g))) D F (G h)) D
                       ((D.inv (Φ (G (F (G f) D F (G g)), G h))) D
                       (Φ (G (F (G f) D F (G g)), G h)) D
                       (F (G (D.inv (Φ (G f, G g)))) D F (G h))) D
                       ((D.inv (Φ (G (F (G f C G g)), G h))) D
                       (Φ (G (F (G f C G g)), G h)) D
                       (F (C.inv (ε (G f C G g))) D F (G h))) D
                       D.inv (Φ (G f C G g, G h)))"
        using assms G0_props preserves_hcomp FF_def D.comp_assoc by auto
      also have "... = F (G 𝖺D[f, g, h]) D
                       F (G (D.inv (η (f D g)) D D.inv (η h))) D
                       F (G (D.inv (Φ (G (f D g), G h)))) D
                       F (C.inv (ε (G (f D g) C G h))) D
                       (Φ (G (f D g), G h) D
                       (F (G (D.inv (η f) D D.inv (η g))) D F (G h)) D
                       (F (G (D.inv (Φ (G f, G g)))) D F (G h)) D
                       (F (C.inv (ε (G f C G g))) D F (G h)) D
                       D.inv (Φ (G f C G g, G h)))"
       proof -
        have "((D.inv (Φ (G (F (G f) D F (G g)), G h))) D
              (Φ (G (F (G f) D F (G g)), G h))) D
              (F (G (D.inv (Φ (G f, G g)))) D F (G h))
                = F (G (D.inv (Φ (G f, G g)))) D F (G h)"
          using assms Φ.components_are_iso C.VV.ide_charSbC C.VV.arr_charSbC G0_props FF_def
                D.comp_inv_arr' D.comp_cod_arr
          by simp
        moreover have "(D.inv (Φ (G (F (G f C G g)), G h)) D
                       Φ (G (F (G f C G g)), G h)) D
                       (F (C.inv (ε (G f C G g))) D F (G h)) =
                         F (C.inv (ε (G f C G g))) D F (G h)"
          using assms D.comp_cod_arr Φ.components_are_iso C.VV.ide_charSbC C.VV.arr_charSbC
                G0_props FF_def D.comp_inv_arr'
          by simp
        ultimately show ?thesis
          using D.comp_assoc by simp
      qed
      also have "... = η (f D g D h) D
                       𝖺D[f, g, h] D
                       ((D.inv (η ((f D g) D h))) D
                       (η ((f D g) D h)) D
                       (D.inv (η (f D g)) D D.inv (η h))) D
                       (((D.inv (η (F (G (f D g)) D F (G h)))) D
                       (η (F (G (f D g)) D F (G h)))) D
                       D.inv (Φ (G (f D g), G h))) D
                       D.inv (η (F (G (f D g) C G h))) D
                       F (C.inv (ε (G (f D g) C G h))) D
                       (Φ (G (f D g), G h) D
                       ((η (f D g) D
                       (D.inv (η f) D D.inv (η g)) D
                       D.inv (η (F (G f) D F (G g)))) D F (G h)) D
                       ((η (F (G f) D F (G g)) D
                       D.inv (Φ (G f, G g)) D
                       D.inv (η (F (G f C G g)))) D F (G h)) D
                       (F (C.inv (ε (G f C G g))) D F (G h)) D
                       D.inv (Φ (G f C G g, G h)))"
      proof -
        have "ν. D.arr ν  F (G ν) = η (D.cod ν) D ν D D.inv (η (D.dom ν))"
          by (metis (full_types) D.arr_dom D.comp_assoc D.ide_dom D.invert_side_of_triangle(2)
              η_naturality(1-2) η_simps(1,6))
        thus ?thesis
          using assms D.comp_assoc by simp
      qed
      also have "... = η (f D g D h) D
                       𝖺D[f, g, h] D
                       (D.inv (η (f D g)) D D.inv (η h)) D
                       D.inv (Φ (G (f D g), G h)) D
                       D.inv (η (F (G (f D g) C G h))) D
                       F (C.inv (ε (G (f D g) C G h))) D
                       (Φ (G (f D g), G h) D
                       ((η (f D g) D
                       (D.inv (η f) D D.inv (η g)) D
                       D.inv (η (F (G f) D F (G g))) D F (G h)) D
                       (η (F (G f) D F (G g)) D
                       D.inv (Φ (G f, G g)) D
                       D.inv (η (F (G f C G g))) D F (G h))) D
                       (F (C.inv (ε (G f C G g))) D F (G h)) D
                       D.inv (Φ (G f C G g, G h)))"
      proof -
        have "(D.inv (η ((f D g) D h)) D η ((f D g) D h)) D
              (D.inv (η (f D g)) D D.inv (η h))
                = D.inv (η (f D g)) D D.inv (η h)"
          using assms D.comp_inv_arr' D.comp_cod_arr η_in_hom by simp
        moreover have "((D.inv (η (F (G (f D g)) D F (G h)))) D
                       η (F (G (f D g)) D F (G h))) D
                       D.inv (Φ (G (f D g), G h))
                         = D.inv (Φ (G (f D g), G h))"
          using assms D.comp_inv_arr' D.comp_cod_arr by simp
        ultimately show ?thesis
          using D.comp_assoc by simp
      qed
      also have "... = η (f D g D h) D
                       𝖺D[f, g, h] D
                       ((D.inv (η (f D g)) D D.inv (η h))) D
                       (D.inv (Φ (G (f D g), G h)) D
                       D.inv (η (F (G (f D g) C G h)))) D
                       F (C.inv (ε (G (f D g) C G h))) D
                       (Φ (G (f D g), G h) D
                       (η (f D g) D F (G h)) D
                       ((D.inv (η f) D D.inv (η g)) D F (G h)) D
                       (D.inv (Φ (G f, G g)) D F (G h)) D
                       (((D.inv (η (F (G f C G g))) D F (G h)) D
                       (F (C.inv (ε (G f C G g))) D F (G h))) D
                       D.inv (Φ (G f C G g, G h))))"
      proof -
        have "(η (f D g) D
              (D.inv (η f) D D.inv (η g)) D
              D.inv (η (F (G f) D F (G g))) D F (G h)) D
              (η (F (G f) D F (G g)) D
              D.inv (Φ (G f, G g)) D
              D.inv (η (F (G f C G g))) D F (G h))
                = (η (f D g) D F (G h)) D
                  ((D.inv (η f) D D.inv (η g)) D F (G h)) D
                  (((D.inv (η (F (G f) D F (G g))) D F (G h)) D
                  (η (F (G f) D F (G g)) D F (G h))) D
                  (D.inv (Φ (G f, G g)) D F (G h))) D
                  (D.inv (η (F (G f C G g))) D F (G h))"
          using assms D.comp_assoc D.whisker_right by simp
        also have "... = (η (f D g) D F (G h)) D
                         ((D.inv (η f) D D.inv (η g)) D F (G h)) D
                         (D.inv (Φ (G f, G g)) D F (G h)) D
                         (D.inv (η (F (G f C G g))) D F (G h))"
        proof -
          have "((D.inv (η (F (G f) D F (G g))) D F (G h)) D
                (η (F (G f) D F (G g)) D F (G h))) D
                (D.inv (Φ (G f, G g)) D F (G h))
                  = D.inv (Φ (G f, G g)) D F (G h)"
            using assms D.comp_inv_arr' D.comp_cod_arr
                  D.interchange [of "D.inv (η (F (G f) D F (G g)))" "η (F (G f) D F (G g))"
                                    "F (G h)" "F (G h)"]
            by simp
          thus ?thesis
            using D.comp_assoc by simp
        qed
        finally show ?thesis
          using D.comp_assoc by simp
      qed
      also have "... = η (f D g D h) D
                       𝖺D[f, g, h] D
                       ((D.inv (η (f D g)) D D.inv (η h))) D
                       (D.inv (Φ (G (f D g), G h)) D
                       ((D.inv (η (F (G (f D g) C G h)))) D
                       F (C.inv (ε (G (f D g) C G h)))) D
                       Φ (G (f D g), G h)) D
                       (η (f D g) D F (G h)) D
                       ((D.inv (η f) D D.inv (η g)) D F (G h)) D
                       (D.inv (Φ (G f, G g)) D F (G h)) D
                       D.inv (Φ (G f C G g, G h))"
      proof -
        have "((D.inv (η (F (G f C G g))) D F (G h)) D
              (F (C.inv (ε (G f C G g))) D F (G h))) D
              D.inv (Φ (G f C G g, G h))
                = D.inv (Φ (G f C G g, G h))"
        proof -
          have "((D.inv (η (F (G f C G g))) D F (G h)) D
                (F (C.inv (ε (G f C G g))) D F (G h))) D
                D.inv (Φ (G f C G g, G h))
                  = ((D.inv (η (F (G f C G g))) D F (G h)) D
                    (D.inv (F (ε (G f C G g))) D F (G h))) D
                    D.inv (Φ (G f C G g, G h))"
            using assms by (simp add: preserves_inv)
          also have "... = (D.inv (η (F (G f C G g))) D
                           D.inv (F (ε (G f C G g))) D F (G h) D F (G h)) D
                           D.inv (Φ (G f C G g, G h))"
            using assms Φ.components_are_iso
                  D.interchange [of "D.inv (η (F (G f C G g)))" "D.inv (F (ε (G f C G g)))"
                                    "F (G h)" "F (G h)"]
            by simp
          also have "... = (D.inv (F (ε (G f C G g)) D η (F (G f C G g))) D F (G h)) D
                           D.inv (Φ (G f C G g, G h))"
          proof -
            have "D.iso (F (ε (G f C G g)) D η (F (G f C G g)))"
              using assms by auto
            hence "D.inv (η (F (G f C G g))) D D.inv (F (ε (G f C G g))) =
                  D.inv (F (ε (G f C G g)) D η (F (G f C G g)))"
              using assms by (simp add: D.inv_comp_right(1))
            thus ?thesis
              using assms by simp
          qed
          also have "... = (F (G f C G g) D F (G h)) D D.inv (Φ (G f C G g, G h))"
            using assms ηε.triangle_G Foε_ηoF.map_def [of "G f C G g"] by simp
          also have "... = D.inv (Φ (G f C G g, G h))"
            using assms D.comp_cod_arr by simp
          finally show ?thesis by simp
        qed
        thus ?thesis
          using D.comp_assoc by simp
      qed
      also have "... = η (f D g D h) D
                       𝖺D[f, g, h] D
                       ((D.inv (η (f D g)) D D.inv (η h))) D
                       ((D.inv (Φ (G (f D g), G h)) D
                       Φ (G (f D g), G h)) D
                       (η (f D g) D F (G h))) D
                       ((D.inv (η f) D D.inv (η g)) D F (G h)) D
                       (D.inv (Φ (G f, G g)) D F (G h)) D
                       D.inv (Φ (G f C G g, G h))"
      proof -
        have "((D.inv (η (F (G (f D g) C G h))) D F (C.inv (ε (G (f D g) C G h))))) D
              (Φ (G (f D g), G h))
                = Φ (G (f D g), G h)"
        proof -
          have "D.inv (η (F (G (f D g) C G h))) D F (C.inv (ε (G (f D g) C G h))) =
                D.inv (η (F (G (f D g) C G h))) D D.inv (F (ε (G (f D g) C G h)))"
            using assms preserves_inv by simp
          also have "... = D.inv (F (ε (G (f D g) C G h)) D η (F (G (f D g) C G h)))"
          proof -
            have "D.iso (F (ε (G (f D g) C G h)) D η (F (G (f D g) C G h)))"
              using assms by auto
            thus ?thesis
              using assms D.inv_comp_right(1) by simp
          qed
          also have "... = F (G (f D g) C G h)"
            using assms ηε.triangle_G Foε_ηoF.map_def [of "G (f D g) C G h"] by simp
          finally show ?thesis
            using assms D.comp_cod_arr [of "Φ (G (f D g), G h)" "F (G (f D g) C G h)"]
            by fastforce
        qed
        thus ?thesis
          using D.comp_assoc by simp
      qed
      also have "... = η (f D g D h) D
                       𝖺D[f, g, h] D
                       (((D.inv (η (f D g)) D D.inv (η h)) D
                       (η (f D g) D F (G h))) D
                       ((D.inv (η f) D D.inv (η g)) D F (G h))) D
                       (D.inv (Φ (G f, G g)) D F (G h)) D
                       D.inv (Φ (G f C G g, G h))"
      proof -
        have "((D.inv (Φ (G (f D g), G h)) D Φ (G (f D g), G h)) D
              (η (f D g) D F (G h)))
                = η (f D g) D F (G h)"
          using assms D.comp_inv_arr' Φ.components_are_iso C.VV.ide_charSbC C.VV.arr_charSbC
                G0_props FF_def D.comp_cod_arr
          by simp
        thus ?thesis
          using D.comp_assoc by simp
      qed
      also have "... = η (f D g D h) D
                       (𝖺D[f, g, h] D
                       ((D.inv (η f) D D.inv (η g)) D D.inv (η h))) D
                       (D.inv (Φ (G f, G g)) D F (G h)) D
                       D.inv (Φ (G f C G g, G h))"
      proof -
        have "((D.inv (η (f D g)) D D.inv (η h)) D (η (f D g) D F (G h))) D
              ((D.inv (η f) D D.inv (η g)) D F (G h))
                = (D.inv (η f) D D.inv (η g)) D D.inv (η h)"
        proof -
          have "(D.inv (η (f D g)) D D.inv (η h)) D (η (f D g) D F (G h))
                  = (f D g) D D.inv (η h)"
            using assms D.comp_inv_arr' D.comp_arr_dom
                  D.interchange [of "D.inv (η (f D g))" "η (f D g)" "D.inv (η h)" "F (G h)"]
                  D.invert_side_of_triangle(2)
            by simp
          moreover have "((f D g) D D.inv (η h)) D
                         ((D.inv (η f) D D.inv (η g)) D F (G h))
                           = (D.inv (η f) D D.inv (η g)) D D.inv (η h)"
            using assms D.comp_cod_arr D.comp_arr_dom
                  D.interchange [of "f D g" "D.inv (η f) D D.inv (η g)" "D.inv (η h)" "F (G h)"]
            by simp
          ultimately show ?thesis by simp
        qed
        thus ?thesis
          using D.comp_assoc by simp
      qed
      also have "... = (η (f D g D h) D
                       (D.inv (η f) D D.inv (η g) D D.inv (η h)) D
                       𝖺D[F (G f), F (G g), F (G h)]) D
                       (D.inv (Φ (G f, G g)) D F (G h)) D
                       D.inv (Φ (G f C G g, G h))"
        using assms D.assoc_naturality [of "D.inv (η f)" "D.inv (η g)" "D.inv (η h)"]
              D.comp_assoc
        by simp
      finally show ?thesis
        using D.comp_assoc by simp
    qed

    lemma coherence_RHS:
    assumes "D.ide f" and "D.ide g" and "D.ide h"
    and "srcD f = trgD g" and "srcD g = trgD h"
    shows "F (ΦG.map (f, g D h) C (G f C ΦG.map (g, h)))
             = (η (f D g D h) D (D.inv (η f) D D.inv (η g) D D.inv (η h))) D
               (F (G f) D D.inv (Φ (G g, G h))) D
               D.inv (Φ (G f, G g C G h))"
    proof -
      have "F (ΦG.map (f, g D h) C (G f C ΦG.map (g, h)))
              = F (G (D.inv (η f) D D.inv (η (g D h))) C
                G (D.inv (Φ (G f, G (g D h)))) C
                C.inv (ε (G f C G (g D h))) C
                (G f C G (D.inv (η g) D D.inv (η h)) C
                        G (D.inv (Φ (G g, G h))) C
                        C.inv (ε (G g C G h))))"
        using assms ΦG_map_simp_ide C.comp_assoc by simp
      also have "... = F (G (D.inv (η f) D D.inv (η (g D h))) C
                       G (D.inv (Φ (G f, G (g D h)))) C
                       C.inv (ε (G f C G (g D h))) C
                       (G f C G (D.inv (η g) D D.inv (η h))) C
                       (G f C G (D.inv (Φ (G g, G h)))) C
                       (G f C C.inv (ε (G g C G h))))"
        using assms C.whisker_left by simp
      also have "... = F (G (D.inv (η f) D D.inv (η (g D h)))) D
                       F (G (D.inv (Φ (G f, G (g D h))))) D
                       F (C.inv (ε (G f C G (g D h)))) D
                       F (G f C G (D.inv (η g) D D.inv (η h))) D
                       F (G f C G (D.inv (Φ (G g, G h)))) D
                       F (G f C C.inv (ε (G g C G h)))"
      proof -
        have "C.arr (G (D.inv (η f) D D.inv (η (g D h))) C
                     G (D.inv (Φ (G f, G (g D h)))) C
                     C.inv (ε (G f C G (g D h))) C
                     (G f C G (D.inv (η g) D D.inv (η h))) C
                     (G f C G (D.inv (Φ (G g, G h)))) C
                     (G f C C.inv (ε (G g C G h))))"
          using assms G0_props by auto
        thus ?thesis
          using assms by (metis C.seqE as_nat_trans.preserves_comp_2)
      qed
      also have "... = F (G (D.inv (η f) D D.inv (η (g D h)))) D
                       F (G (D.inv (Φ (G f, G (g D h))))) D
                       F (C.inv (ε (G f C G (g D h)))) D
                       (Φ (G f, G (g D h)) D
                       (F (G f) D F (G (D.inv (η g) D D.inv (η h)))) D
                       ((D.inv (Φ (G f, G (F (G g) D F (G h))))) D
                       (Φ (G f, G (F (G g) D F (G h)))) D
                       (F (G f) D F (G (D.inv (Φ (G g, G h)))))) D
                       ((D.inv (Φ (G f, G (F (G g C G h))))) D
                       (Φ (G f, G (F (G g C G h)))) D
                       (F (G f) D F (C.inv (ε (G g C G h))))) D
                       D.inv (Φ (G f, G g C G h)))"
        using assms preserves_hcomp G0_props D.comp_assoc by simp
      also have "... = F (G (D.inv (η f) D D.inv (η (g D h)))) D
                       F (G (D.inv (Φ (G f, G (g D h))))) D
                       F (C.inv (ε (G f C G (g D h)))) D
                       (Φ (G f, G (g D h)) D
                       (F (G f) D F (G (D.inv (η g) D D.inv (η h)))) D
                       (F (G f) D F (G (D.inv (Φ (G g, G h))))) D
                       (F (G f) D F (C.inv (ε (G g C G h)))) D
                       D.inv (Φ (G f, G g C G h)))"
      proof -
        have "D.inv (Φ (G f, G (F (G g) D F (G h)))) D
                     Φ (G f, G (F (G g) D F (G h))) D
                     (F (G f) D F (G (D.inv (Φ (G g, G h)))))
                = (F (G f) D F (G (D.inv (Φ (G g, G h)))))"
        proof -
          have "D.inv (Φ (G f, G (F (G g) D F (G h)))) D Φ (G f, G (F (G g) D F (G h)))
                  = F (G f) D F (G (F (G g) D F (G h)))"
            using assms D.comp_inv_arr D.inv_is_inverse G0_props FF_def by simp
          moreover have "... = D.cod (F (G f) D F (G (D.inv (Φ (G g, G h)))))"
            using assms G0_props by simp
          moreover have "D.hseq (F (G f)) (F (G (D.inv (Φ (G g, G h)))))"
            using assms G0_props by auto
          ultimately show ?thesis
            using assms D.comp_assoc
                  D.comp_cod_arr [of "F (G f) D F (G (D.inv (Φ (G g, G h))))"
                                     "F (G f) D F (G (F (G g) D F (G h)))"]
            by metis
        qed
        moreover have "D.inv (Φ (G f, G (F (G g C G h)))) D
                       (Φ (G f, G (F (G g C G h)))) D
                       (F (G f) D F (C.inv (ε (G g C G h))))
                         = (F (G f) D F (C.inv (ε (G g C G h))))"
        proof -
          have "D.inv (Φ (G f, G (F (G g C G h)))) D Φ (G f, G (F (G g C G h)))
                  = F (G f) D F (G (F (G g C G h)))"
            using assms D.comp_inv_arr D.inv_is_inverse G0_props by simp
          moreover have "... = D.cod (F (G f) D F (C.inv (ε (G g C G h))))"
            using assms by simp
          moreover have "D.hseq (F (G f)) (F (C.inv (ε (G g C G h))))"
            using assms by (intro D.hseqI') auto
          ultimately show ?thesis
            using assms D.comp_assoc
                  D.comp_cod_arr [of "F (G f) D F (C.inv (ε (G g C G h)))"
                                     "F (G f) D F (G (F (G g C G h)))"]
            by metis
        qed
        ultimately show ?thesis by simp
      qed
      also have "... = (η (f D g D h) D
                       (D.inv (η f) D D.inv (η (g D h))) D
                       ((D.inv (η (F (G f) D F (G (g D h))))) D
                       η (F (G f) D F (G (g D h)))) D
                       D.inv (Φ (G f, G (g D h)))) D
                       ((D.inv (η (F (G f C G (g D h)))) D
                       F (C.inv (ε (G f C G (g D h))))) D
                       Φ (G f, G (g D h))) D
                       (F (G f) D
                          η (g D h) D
                          (D.inv (η g) D D.inv (η h)) D
                          D.inv (η (F (G g) D F (G h)))) D
                       (F (G f) D
                          η (F (G g) D F (G h)) D
                          D.inv (Φ (G g, G h)) D
                          D.inv (η (F (G g C G h)))) D
                       (F (G f) D F (C.inv (ε (G g C G h)))) D
                       D.inv (Φ (G f, G g C G h))"
      proof -
        have "ν. D.arr ν  F (G ν) = η (D.cod ν) D ν D D.inv (η (D.dom ν))"
          by (metis (full_types) D.arr_dom D.comp_assoc D.ide_dom D.invert_side_of_triangle(2)
              η_naturality(1-2) η_simps(1,6))
        thus ?thesis
          using assms D.comp_assoc by simp
      qed
      also have "... = (η (f D g D h) D
                       (D.inv (η f) D D.inv (η (g D h))) D
                       D.inv (Φ (G f, G (g D h)))) D
                       Φ (G f, G (g D h)) D
                       (F (G f) D
                          η (g D h) D
                          (D.inv (η g) D D.inv (η h)) D
                          D.inv (η (F (G g) D F (G h)))) D
                       (F (G f) D
                          η (F (G g) D F (G h)) D
                          D.inv (Φ (G g, G h)) D
                          D.inv (η (F (G g C G h)))) D
                       (F (G f) D F (C.inv (ε (G g C G h)))) D
                       D.inv (Φ (G f, G g C G h))"
      proof -
        have "D.inv (η (F (G f) D F (G (g D h)))) D η (F (G f) D F (G (g D h))) D
              D.inv (Φ (G f, G (g D h)))
                = D.inv (Φ (G f, G (g D h)))"
        proof -
          have "D.inv (η (F (G f) D F (G (g D h)))) D η (F (G f) D F (G (g D h))) =
                F (G f) D F (G (g D h))"
            using assms D.comp_inv_arr D.inv_is_inverse by simp
          moreover have "... = D.cod (D.inv (Φ (G f, G (g D h))))"
            using assms by simp
          ultimately show ?thesis
            using assms D.comp_cod_arr [of "D.inv (Φ (G f, G (g D h)))"]
            by (metis D.arr_inv D.comp_assoc D.hcomp_simps(2) D.ideD(1) D.ide_hcomp
                G.preserves_ide G_simps(2) G_simps(3) cmp_components_are_iso)
        qed
        moreover have "(D.inv (η (F (G f C G (g D h)))) D
                          F (C.inv (ε (G f C G (g D h))))) D
                       Φ (G f, G (g D h))
                         = Φ (G f, G (g D h))"
        proof -
          have "D.inv (η (F (G f C G (g D h)))) D F (C.inv (ε (G f C G (g D h))))
                  = F (G f C G (g D h))"
          proof -
            have "D.inv (η (F (G f C G (g D h)))) D F (C.inv (ε (G f C G (g D h)))) =
                  D.inv (η (F (G f C G (g D h)))) D D.inv (F (ε (G f C G (g D h))))"
              using assms by (simp add: preserves_inv)
            also have "... = D.inv (F (ε (G f C G (g D h))) D η (F (G f C G (g D h))))"
            proof -
              have "D.iso (F (ε (G f C G (g D h))) D η (F (G f C G (g D h))))"
                using assms by auto
              thus ?thesis
                using assms by (simp add: D.inv_comp_right(1))
            qed
            also have "... = F (G f C G (g D h))"
              using assms Foε_ηoF.map_def [of "G f C G (g D h)"] ηε.triangle_G by simp
            finally show ?thesis by blast
          qed
          moreover have "... = D.cod (Φ (G f, G (g D h)))"
            using assms by simp
          moreover have "D.arr (Φ (G f, G (g D h)))"
            using assms by auto
          ultimately show ?thesis
            using D.comp_cod_arr by simp
        qed
        ultimately show ?thesis
          using D.comp_assoc by simp
      qed
      also have "... = (η (f D g D h) D
                       (D.inv (η f) D D.inv (η (g D h))) D
                       ((D.inv (Φ (G f, G (g D h)))) D
                       Φ (G f, G (g D h))) D
                       (F (G f) D η (g D h))) D
                       (F (G f) D D.inv (η g) D D.inv (η h)) D
                       (((F (G f) D D.inv (η (F (G g) D F (G h)))) D
                       (F (G f) D η (F (G g) D F (G h)))) D
                       (F (G f) D D.inv (Φ (G g, G h)))) D
                       ((F (G f) D D.inv (η (F (G g C G h)))) D
                       (F (G f) D F (C.inv (ε (G g C G h))))) D
                       D.inv (Φ (G f, G g C G h))"
        using assms D.whisker_left D.comp_assoc by simp
      also have "... = (η (f D g D h) D
                       (((D.inv (η f) D D.inv (η (g D h))) D
                       (F (G f) D η (g D h)))) D
                       (F (G f) D D.inv (η g) D D.inv (η h))) D
                       (F (G f) D D.inv (Φ (G g, G h))) D
                       D.inv (Φ (G f, G g C G h))"
      proof -
        have "(D.inv (Φ (G f, G (g D h))) D Φ (G f, G (g D h))) D (F (G f) D η (g D h))
                = (F (G f) D η (g D h))"
          using assms D.comp_inv_arr' G0_props D.comp_cod_arr by simp
        moreover have "((F (G f) D D.inv (η (F (G g) D F (G h)))) D
                       (F (G f) D η (F (G g) D F (G h)))) D
                       (F (G f) D D.inv (Φ (G g, G h)))
                         = F (G f) D D.inv (Φ (G g, G h))"
        proof -
          have "(F (G f) D D.inv (η (F (G g) D F (G h)))) D
                (F (G f) D η (F (G g) D F (G h)))
                  = F (G f) D F (G g) D F (G h)"
          proof -
            have "(F (G f) D D.inv (η (F (G g) D F (G h)))) D
                  (F (G f) D η (F (G g) D F (G h)))
                    = F (G f) D D.inv (η (F (G g) D F (G h))) D η (F (G g) D F (G h))"
              using assms D.interchange [of "F (G f)" "F (G f)"] by simp
            also have "... = F (G f) D F (G g) D F (G h)"
              using assms D.comp_inv_arr' by simp
            finally show ?thesis by blast
          qed
          moreover have "... = D.cod (F (G f) D D.inv (Φ (G g, G h)))"
            using assms by simp
          moreover have "D.arr (F (G f) D D.inv (Φ (G g, G h)))"
            using assms by simp
          ultimately show ?thesis
            using D.comp_cod_arr by simp
        qed
        moreover have "((F (G f) D D.inv (η (F (G g C G h)))) D
                       (F (G f) D F (C.inv (ε (G g C G h))))) D
                       D.inv (Φ (G f, G g C G h))
                         = D.inv (Φ (G f, G g C G h))"
        proof -
          have "(F (G f) D D.inv (η (F (G g C G h)))) D
                (F (G f) D F (C.inv (ε (G g C G h))))
                  = F (G f) D F (G g C G h)"
          proof -
            have "(F (G f) D D.inv (η (F (G g C G h)))) D
                  (F (G f) D F (C.inv (ε (G g C G h))))
                    = F (G f) D D.inv (η (F (G g C G h))) D F (C.inv (ε (G g C G h)))"
              using assms D.interchange [of "F (G f)" "F (G f)"] by simp
            also have "... = F (G f) D D.inv (η (F (G g C G h))) D
                                        D.inv (F (ε (G g C G h)))"
              using assms preserves_inv by simp
            also have "... = F (G f) D D.inv (F (ε (G g C G h)) D η (F (G g C G h)))"
            proof -
              have "D.iso (F (ε (G g C G h)) D η (F (G g C G h)))"
                using assms by auto
              thus ?thesis
                using assms by (simp add: D.inv_comp_right(1))
            qed
            also have "... = F (G f) D F (G g C G h)"
              using assms ηε.triangle_G Foε_ηoF.map_def [of "G g C G h"] by simp
            finally show ?thesis by blast
          qed
          thus ?thesis using assms D.comp_cod_arr by simp
        qed
        ultimately show ?thesis
          using D.comp_assoc by simp
      qed
      also have "... = (η (f D g D h) D
                       (D.inv (η f) D D.inv (η g) D D.inv (η h))) D
                       (F (G f) D D.inv (Φ (G g, G h))) D
                       D.inv (Φ (G f, G g C G h))"
      proof -
        have "((D.inv (η f) D D.inv (η (g D h))) D
              (F (G f) D η (g D h))) D
              (F (G f) D D.inv (η g) D D.inv (η h))
                = D.inv (η f) D D.inv (η g) D D.inv (η h)"
          using assms D.comp_cod_arr D.comp_arr_dom D.interchange
        proof -
          have "((D.inv (η f) D D.inv (η (g D h))) D
                (F (G f) D η (g D h))) D
                (F (G f) D D.inv (η g) D D.inv (η h))
                  = (D.inv (η f) D F (G f) D D.inv (η (g D h)) D η (g D h)) D
                    (F (G f) D D.inv (η g) D D.inv (η h))"
            using assms D.interchange by simp
          also have "... = (D.inv (η f) D g D h) D (F (G f) D D.inv (η g) D D.inv (η h))"
            using assms D.comp_inv_arr' D.comp_arr_dom by simp
          also have "... = D.inv (η f) D F (G f) D (g D h) D (D.inv (η g) D D.inv (η h))"
            using assms D.interchange by simp
          also have "... = D.inv (η f) D D.inv (η g) D D.inv (η h)"
            using assms D.comp_arr_dom D.comp_cod_arr by simp
          finally show ?thesis by simp
        qed
        thus ?thesis
          using D.comp_assoc by simp
      qed
      finally show ?thesis by blast
    qed

    interpretation G: pseudofunctor
                        VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC G ΦG.map
    proof
      show "f g h. D.ide f; D.ide g; D.ide h; srcD f = trgD g; srcD g = trgD h
                       G 𝖺D[f, g, h] C ΦG.map (f D g, h) C (ΦG.map (f, g) C G h) =
                          ΦG.map (f, g D h) C (G f C ΦG.map (g, h)) C 𝖺C[G f, G g, G h]"
      proof -
        fix f g h
        assume f: "D.ide f" and g: "D.ide g" and h: "D.ide h"
        assume fg: "srcD f = trgD g" and gh: "srcD g = trgD h"
        have "F (G 𝖺D[f, g, h] C ΦG.map (f D g, h) C (ΦG.map (f, g) C G h)) =
              F (ΦG.map (f, g D h) C (G f C ΦG.map (g, h)) C 𝖺C[G f, G g, G h])"
        proof -
          have "F (G 𝖺D[f, g, h] C ΦG.map (f D g, h) C (ΦG.map (f, g) C G h))
                  = (η (f D g D h) D
                    (D.inv (η f) D D.inv (η g) D D.inv (η h))) D
                    𝖺D[F (G f), F (G g), F (G h)] D
                    (D.inv (Φ (G f, G g)) D F (G h)) D
                    D.inv (Φ (G f C G g, G h))"
            using f g h fg gh coherence_LHS by simp
          also have "... = (η (f D g D h) D
                           (D.inv (η f) D D.inv (η g) D D.inv (η h))) D
                           ((F (G f) D D.inv (Φ (G g, G h))) D
                           D.inv (Φ (G f, G g C G h))) D
                           F 𝖺C[G f, G g, G h]"
          proof -
            have "𝖺D[F (G f), F (G g), F (G h)] D (D.inv (Φ (G f, G g)) D F (G h)) D
                  D.inv (Φ (G f C G g, G h))
                    = 𝖺D[F (G f), F (G g), F (G h)] D D.inv (Φ (G f C G g, G h) D
                      (Φ (G f, G g) D F (G h)))"
            proof -
              have "D.iso (Φ (G f C G g, G h) D (Φ (G f, G g) D F (G h)))"
                using f g h fg gh Φ.components_are_iso C.VV.ide_charSbC C.VV.arr_charSbC FF_def
                by (intro D.iso_hcomp D.isos_compose D.seqI) auto
              thus ?thesis
                using f g h fg gh
                by (simp add: C.VV.arr_charSbC D.inv_comp_right(1))
            qed
            also have "... = D.inv (Φ (G f, G g C G h) D (F (G f) D Φ (G g, G h))) D
                             F 𝖺C[G f, G g, G h]"
              using f g h fg gh cmp_simps assoc_coherence D.comp_assoc D.isos_compose
                    Φ.components_are_iso C.VV.ide_charSbC C.VV.arr_charSbC FF_def
                    D.invert_opposite_sides_of_square
              by simp
            also have "... = ((F (G f) D D.inv (Φ (G g, G h))) D
                                          D.inv (Φ (G f, G g C G h))) D
                             F 𝖺C[G f, G g, G h]"
            proof -
              have "D.iso (Φ (G f, G g C G h) D (F (G f) D Φ (G g, G h)))"
                using f g h fg gh Φ.components_are_iso C.VV.ide_charSbC C.VV.arr_charSbC FF_def
                by auto
              thus ?thesis
                using f g h fg gh Φ.components_are_iso C.VV.ide_charSbC C.VV.arr_charSbC FF_def
                by (simp add: C.VV.arr_charSbC D.comp_assoc D.inv_comp_right(1))
            qed
            finally show ?thesis
              using D.comp_assoc by simp
          qed
          also have "... = F (ΦG.map (f, g D h) C (G f C ΦG.map (g, h))) D
                           F 𝖺C[G f, G g, G h]"
            using f g h fg gh coherence_RHS D.comp_assoc by simp
          also have "... = F ((ΦG.map (f, g D h) C (G f C ΦG.map (g, h))) C
                           𝖺C[G f, G g, G h])"
            using f g h fg gh ΦG_simps by auto
          also have "... = F (ΦG.map (f, g D h) C (G f C ΦG.map (g, h)) C
                           𝖺C[G f, G g, G h])"
            using C.comp_assoc by simp
          finally show ?thesis by simp
        qed
        moreover have "C.par (G 𝖺D[f, g, h] C ΦG.map (f D g, h) C
                               (ΦG.map (f, g) C G h))
                             (ΦG.map (f, g D h) C (G f C ΦG.map (g, h)) C 𝖺C[G f, G g, G h])"
          using f g h fg gh ΦG_simps by auto
        ultimately show "G 𝖺D[f, g, h] C ΦG.map (f D g, h) C (ΦG.map (f, g) C G h) =
                         ΦG.map (f, g D h) C (G f C ΦG.map (g, h)) C 𝖺C[G f, G g, G h]"
          using is_faithful by blast
      qed
    qed

    interpretation GF: composite_pseudofunctor VC HC 𝖺C 𝗂C srcC trgC
                         VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC F Φ G ΦG.map
      ..
    interpretation FG: composite_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD
                         VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD G ΦG.map F Φ
      ..
    interpretation IC: identity_pseudofunctor VC HC 𝖺C 𝗂C srcC trgC ..
    interpretation ID: identity_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD ..

    lemma η_equals_FG_unit:
    assumes "D.obj a"
    shows "η a = FG.unit a"
    proof (intro FG.unit_eqI)
      show "D.obj a" by fact
      show "D.iso (η a)"
        using assms by auto
      show "«η a : FG.map0 a D FG.map a»"
        using assms η_in_hom [of a] FG.map0_def G0_props D.obj_def
        by (metis D.ideD(2) D.ideD(3) D.objE D.vconn_implies_hpar(3) η.preserves_cod
            η_simps(5))
      show "η a D 𝗂D[FG.map0 a] = (FG.map 𝗂D[a] D FG.cmp (a, a)) D (η a D η a)"
      proof -
        have "(FG.map 𝗂D[a] D FG.cmp (a, a)) D (η a D η a) =
              F (G 𝗂D[a]) D FG.cmp (a, a) D (η a D η a)"
          using assms D.comp_assoc by simp
        also have "... = F (G 𝗂D[a]) D F (G (ID.cmp (a, a))) D F (ΦG.map (a, a)) D
                         Φ (G a, G a) D (η a D η a)"
          using assms FG.cmp_def D.comp_assoc D.VV.arr_charSbC D.VV.dom_charSbC D.VV.cod_charSbC
          by auto
        also have "... = F (G 𝗂D[a]) D F (G (ID.cmp (a, a))) D η (a D a)"
          using assms η_hcomp by auto
        also have "... = F (G 𝗂D[a]) D η (a D a)"
          using assms D.comp_cod_arr by auto
        also have "... = η a D 𝗂D[a]"
          using assms η_naturality [of "𝗂D[a]"] by simp
        also have "... = η a D 𝗂D[FG.map0 a]"
          using assms «η a : FG.map0 a D FG.map a» by fastforce
        finally show ?thesis by simp
      qed
    qed

    lemma ε_hcomp':
    assumes "C.ide g" and "C.ide f" and "srcC g = trgC f"
    shows "ε (g C f) C GF.cmp (g, f) = ε g C ε f"
    proof -
      have "ε (g C f) C GF.cmp (g, f)
              = (ε g C ε f) C C.inv (ΦG.map (F g, F f)) C C.inv (G (Φ (g, f))) C
                G (F (g C f)) C G (Φ (g, f)) C ΦG.map (F g, F f)"
        using assms ε_hcomp GF.cmp_def C.VV.arr_charSbC C.comp_cod_arr
              C.comp_inv_arr' C.comp_assoc C.VV.dom_charSbC C.VV.cod_charSbC
        by simp
      also have "... = (ε g C ε f) C C.inv (ΦG.map (F g, F f)) C (C.inv (G (Φ (g, f))) C
                       G (F (g C f)) C G (Φ (g, f))) C ΦG.map (F g, F f)"
        using C.comp_assoc by simp
      also have "... = (ε g C ε f) C C.inv (ΦG.map (F g, F f)) C (C.inv (G (Φ (g, f))) C
                       G (Φ (g, f))) C ΦG.map (F g, F f)"
        using assms C.comp_ide_arr [of "G (F (g C f))" "G (Φ (g, f))"]
              C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC
        by simp
      also have "... = (ε g C ε f) C C.inv (ΦG.map (F g, F f)) C ΦG.map (F g, F f)"
      proof -
        have "C.inv (G (Φ (g, f))) C G (Φ (g, f)) = G (F g D F f)"
          using assms C.comp_inv_arr' cmp_components_are_iso C.inv_is_inverse
                C.VV.arr_charSbC C.VV.ide_charSbC cmp_simps(4)
          by auto
        moreover have "... = C.cod (ΦG.map (F g, F f))"
          using assms by simp
        ultimately have "(C.inv (G (Φ (g, f))) C G (Φ (g, f))) C ΦG.map (F g, F f) =
                         ΦG.map (F g, F f)"
          using assms C.comp_cod_arr [of "ΦG.map (F g, F f)" "G (F g D F f)"]
                C.ideD(1) G.cmp_simps(1) preserves_ide preserves_src preserves_trg
          by presburger
        thus ?thesis by simp
      qed
      also have "... = ε g C ε f"
        using assms C.comp_inv_arr' C.comp_arr_dom [of "ε g C ε f" "G (F g) C G (F f)"]
        by simp
      finally show ?thesis by simp
    qed

    lemma ε_inverts_GF_unit:
    assumes "C.obj a"
    shows "ε a C GF.unit a = a"
    proof -
      have "ε a C GF.unit a = IC.unit a"
      proof (intro IC.unit_eqI)
        show "C.obj a" by fact
        show 1: "«ε a C GF.unit a : IC.map0 a C IC.map a»"
        proof -
          have "srcC (G (F a)) = srcC (IC.map a)"
            using assms G0_props C.obj_def' by simp
          thus ?thesis
            using assms IC.map0_def GF.map0_def GF.unit_in_hom
            by (intro C.comp_in_homI') auto
        qed
        show "C.iso (ε a C GF.unit a)"
          using assms ε_in_hom GF.unit_char(2)
          by (intro C.isos_compose) auto
        show "(ε a C GF.unit a) C 𝗂C[IC.map0 a]
                = (IC.map 𝗂C[a] C IC.cmp (a, a)) C (ε a C GF.unit a C ε a C GF.unit a)"
        proof -
          have "(IC.map 𝗂C[a] C IC.cmp (a, a)) C (ε a C GF.unit a C ε a C GF.unit a) =
                𝗂C[a] C (a C a) C (ε a C GF.unit a C ε a C GF.unit a)"
            using assms C.comp_assoc by simp
          also have "... = 𝗂C[a] C (ε a C GF.unit a C ε a C GF.unit a)"
          proof -
            have "C.hseq (ε a C GF.unit a) (ε a C GF.unit a)"
              using assms GF.unit_simps C.iso_is_arr C.iso (ε a C GF.unit a)
              by (intro C.hseqI') auto
            moreover have "a C a = C.cod (ε a C GF.unit a C ε a C GF.unit a)"
            proof -
              have "C.cod (ε a C GF.unit a) = a"
                using assms 1 C.obj_simps by auto
              moreover have "C.hseq (ε a C GF.unit a) (ε a C GF.unit a)"
                using assms 1 C.src_dom [of "ε a C GF.unit a"] C.trg_dom [of "ε a C GF.unit a"]
                apply (intro C.hseqI')
                by auto force
              ultimately show ?thesis by auto
            qed
            ultimately show ?thesis
              using C.comp_cod_arr by simp
          qed
          also have "... = 𝗂C[a] C (ε a C ε a) C (GF.unit a C GF.unit a)"
            using assms C.interchange [of "ε a" "GF.unit a" "ε a" "GF.unit a"]
            by (simp add: C.iso_is_arr C.iso (ε a C GF.unit a))
          also have "... = 𝗂C[a] C (ε (a C a) C GF.cmp (a, a)) C (GF.unit a C GF.unit a)"
            using assms ε_hcomp' [of a a] by auto
          also have "... = (𝗂C[a] C ε (a C a)) C GF.cmp (a, a) C (GF.unit a C GF.unit a)"
            using C.comp_assoc by simp
          also have "... = (ε a C G (F 𝗂C[a])) C GF.cmp (a, a) C (GF.unit a C GF.unit a)"
            using assms ε_naturality [of "𝗂C[a]"] by simp
          also have "... = ε a C (G (F 𝗂C[a]) C GF.cmp (a, a)) C (GF.unit a C GF.unit a)"
            using C.comp_assoc by simp
          also have "... = ε a C GF.unit a C 𝗂C[GF.map0 a]"
            using assms GF.unit_char [of a] by simp
          also have "... = (ε a C GF.unit a) C 𝗂C[IC.map0 a]"
          proof -
            have "GF.map0 a = IC.map0 a"
              using assms G0_props(2) [of a] GF.map0_def by auto
            thus ?thesis
              using assms GF.unit_char [of a] C.comp_assoc by simp
          qed
          finally show ?thesis
            using C.comp_assoc by simp
        qed
      qed
      thus ?thesis
        using assms IC.unit_char' by simp
    qed

    lemma η_respects_comp:
    assumes "D.ide f" and "D.ide g" and "srcD g = trgD f"
    shows "(𝗅D-1[F (G (g D f))] D η (g D f) D 𝗋D[g D f]) D ((g D f) D srcD f)
             = (trgD g D FG.cmp (g, f)) D 𝗅D-1[F (G g) D F (G f)] D (η g D η f) D
               𝗋D[g D f]"
    and "(trgD g D FG.cmp (g, f)) D 𝖺D[trgD g, F (G g), F (G f)] D
         (𝗅D-1[F (G g)] D η g D 𝗋D[g] D F (G f)) D D.inv 𝖺D[g, srcD g, F (G f)] D
         (g D 𝗅D-1[F (G f)] D η f D 𝗋D[f]) D 𝖺D[g, f, srcD f]
           = (trgD g D FG.cmp (g, f)) D 𝗅D-1[F (G g) D F (G f)] D (η g D η f) D
             𝗋D[g D f]"
    proof -
      show "(𝗅D-1[F (G (g D f))] D η (g D f) D 𝗋D[g D f]) D ((g D f) D srcD f)
              = (trgD g D FG.cmp (g, f)) D 𝗅D-1[F (G g) D F (G f)] D (η g D η f) D
                𝗋D[g D f]"
      proof -
        have "(𝗅D-1[F (G (g D f))] D η (g D f) D 𝗋D[g D f]) D ((g D f) D srcD f)
                = 𝗅D-1[F (G (g D f))] D η (g D f) D 𝗋D[g D f]"
          using assms D.comp_assoc D.comp_arr_dom by simp
        also have 1: "... = (𝗅D-1[F (G (g D f))] D F (ΦG.map (g, f))) D
                            Φ (G g, G f) D (η g D η f) D 𝗋D[g D f]"
          using assms η_hcomp D.comp_assoc by simp
        also have "... = (trgD g D F (ΦG.map (g, f))) D (𝗅D-1[F (G g C G f)] D
                         Φ (G g, G f)) D (η g D η f) D 𝗋D[g D f]"
        proof -
          have "𝗅D-1[F (G (g D f))] D F (ΦG.map (g, f)) =
                (trgD g D F (ΦG.map (g, f))) D 𝗅D-1[F (G g C G f)]"
            using assms G0_props D.lunit'_naturality [of "F (ΦG.map (g, f))"] ΦG_in_hom [of g f]
            by auto
          thus ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = ((trgD g D F (ΦG.map (g, f))) D (trgD g D Φ (G g, G f))) D
                         𝗅D-1[F (G g) D F (G f)] D (η g D η f) D 𝗋D[g D f]"
        proof -
          have "𝗅D-1[F (G g C G f)] D Φ (G g, G f)
                  = (trgD g D Φ (G g, G f)) D 𝗅D-1[F (G g) D F (G f)]"
            using assms D.lunit'_naturality [of "Φ (G g, G f)"] G0_props by fastforce
          thus ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = (trgD g D F (ΦG.map (g, f)) D Φ (G g, G f)) D
                         𝗅D-1[F (G g) D F (G f)] D (η g D η f) D 𝗋D[g D f]"
          using assms 1 D.whisker_left [of "trgD g" "F (ΦG.map (g, f))" "Φ (G g, G f)"]
          by force
        also have "... = (trgD g D FG.cmp (g, f)) D
                         𝗅D-1[F (G g) D F (G f)] D (η g D η f) D 𝗋D[g D f]"
        proof -
          have 1: "FG.cmp (g, f) = (F (G (g D f)) D F (ΦG.map (g, f))) D Φ (G g, G f)"
            using assms FG.cmp_def D.VV.arr_charSbC D.VV.dom_charSbC D.comp_assoc by simp
          also have "... = F (ΦG.map (g, f)) D Φ (G g, G f)"
          proof -
            have "D.cod (F (ΦG (g, f))) = F (G (g D f))"
              using assms 1
              by (metis (mono_tags, lifting) D.cod_eqI D.ideD(1) D.ide_hcomp D.seqE
                  FG.cmp_simps'(1) G.preserves_ide preserves_ide)
            thus ?thesis
              using assms D.comp_cod_arr [of "F (ΦG.map (g, f))" "F (G (g D f))"]
              by fastforce
          qed
          finally have "FG.cmp (g, f) = F (ΦG.map (g, f)) D Φ (G g, G f)" by simp
          thus ?thesis by simp
        qed
        finally show ?thesis by simp
      qed
      show "(trgD g D FG.cmp (g, f)) D
            𝖺D[trgD g, F (G g), F (G f)] D
            (𝗅D-1[F (G g)] D η g D 𝗋D[g] D F (G f)) D
            D.inv 𝖺D[g, srcD g, F (G f)] D
            (g D 𝗅D-1[F (G f)] D η f D 𝗋D[f]) D
            𝖺D[g, f, srcD f]
              = (trgD g D FG.cmp (g, f)) D
                𝗅D-1[F (G g) D F (G f)] D
                (η g D η f) D
                𝗋D[g D f]"
      proof -
        have "(trgD g D FG.cmp (g, f)) D
              𝖺D[trgD g, F (G g), F (G f)] D
              (𝗅D-1[F (G g)] D η g D 𝗋D[g] D F (G f)) D
              D.inv 𝖺D[g, srcD g, F (G f)] D
              (g D 𝗅D-1[F (G f)] D η f D 𝗋D[f]) D
              𝖺D[g, f, srcD f]
                = (trgD g D FG.cmp (g, f)) D
                  𝖺D[trgD g, F (G g), F (G f)] D
                  (𝗅D-1[F (G g)] D F (G f)) D
                  (η g D F (G f)) D
                  (𝗋D[g] D F (G f)) D
                  D.inv 𝖺D[g, srcD g, F (G f)] D
                  (g D 𝗅D-1[F (G f)]) D
                  (g D η f) D
                  (g D 𝗋D[f]) D
                  𝖺D[g, f, srcD f]"
          using assms D.comp_assoc D.whisker_right D.whisker_left by simp
        also have "... = (trgD g D FG.cmp (g, f)) D
                         (𝖺D[trgD g, F (G g), F (G f)] D
                         (𝗅D-1[F (G g)] D F (G f))) D
                         (η g D F (G f)) D
                         (𝗋D[g] D F (G f)) D
                         D.inv 𝖺D[g, srcD g, F (G f)] D
                         (g D 𝗅D-1[F (G f)]) D
                         (g D η f) D
                         ((g D 𝗋D[f]) D
                         𝖺D[g, f, srcD f])"
          using assms D.comp_cod_arr D.VV.ide_charSbC D.VV.arr_charSbC D.VV.dom_charSbC
                FG.FF_def G0_props D.comp_assoc
          by presburger
        also have "... = (trgD g D FG.cmp (g, f)) D
                         𝗅D-1[F (G g) D F (G f)] D
                         (η g D F (G f)) D
                         ((𝗋D[g] D F (G f)) D
                         D.inv 𝖺D[g, srcD g, F (G f)]) D
                         (g D 𝗅D-1[F (G f)]) D
                         (g D η f) D
                         𝗋D[g D f]"
        proof -
          have "(g D 𝗋D[f]) D 𝖺D[g, f, srcD f] = 𝗋D[g D f]"
            using assms D.runit_hcomp by simp
          moreover have "𝖺D[trgD g, F (G g), F (G f)] D (𝗅D-1[F (G g)] D F (G f)) =
                         𝗅D-1[F (G g) D F (G f)]"
            using assms D.lunit_hcomp [of "F (G g)" "F (G f)"] G0_props by simp
          ultimately show ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = (trgD g D FG.cmp (g, f)) D
                         𝗅D-1[F (G g) D F (G f)] D
                         (η g D F (G f)) D
                         (((g D 𝗅D[F (G f)]) D
                         (g D 𝗅D-1[F (G f)])) D
                         (g D η f)) D
                         𝗋D[g D f]"
        proof -
          have "(𝗋D[g] D F (G f)) D D.inv 𝖺D[g, srcD g, F (G f)] = g D 𝗅D[F (G f)]"
            using assms D.triangle' [of g "F (G f)"] G0_props by simp
          thus ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = (trgD g D FG.cmp (g, f)) D
                         𝗅D-1[F (G g) D F (G f)] D
                         ((η g D F (G f)) D
                         (g D η f)) D
                         𝗋D[g D f]"
        proof -
          have "((g D 𝗅D[F (G f)]) D (g D 𝗅D-1[F (G f)])) D (g D η f) = g D η f"
            using assms D.interchange [of g g "𝗅D[F (G f)]" "𝗅D-1[F (G f)]"]
                  D.comp_arr_inv' D.comp_cod_arr
            by simp
          thus ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = (trgD g D FG.cmp (g, f)) D 𝗅D-1[F (G g) D F (G f)] D
                           (η g D η f) D 𝗋D[g D f]"
          using assms D.interchange [of "η g" g "F (G f)" "η f"] D.comp_arr_dom D.comp_cod_arr
          by simp
        finally show ?thesis by simp
      qed
    qed

    lemma η_respects_unit:
    assumes "D.obj a"
    shows "(a D FG.unit a) D 𝗋D-1[a] D 𝗅D[a] =
           (𝗅D-1[FG.map (D.cod a)] D η a D 𝗋D[D.dom a]) D (ID.unit a D a)"
    proof -
      have "(𝗅D-1[FG.map (D.cod a)] D η a D 𝗋D[D.dom a]) D (ID.unit a D a) =
            (𝗅D-1[F (G a)] D η a) D 𝗋D[a]"
        using assms ID.lunit_coherence ID.unit_char' D.comp_arr_dom D.comp_assoc by auto
      also have "... = ((a D η a) D 𝗅D-1[a]) D 𝗋D[a]"
        using assms D.lunit'_naturality [of "η a"] by auto
      also have "... = (a D η a) D 𝗅D-1[a] D 𝗋D[a]"
        using D.comp_assoc by simp
      also have "... = (a D η a) D 𝗋D-1[a] D 𝗅D[a]"
        using assms D.unitor_coincidence by simp
      also have "... = (a D FG.unit a) D 𝗋D-1[a] D 𝗅D[a]"
        using assms η_equals_FG_unit by simp
      finally show ?thesis by simp
    qed

    lemma ε_respects_comp:
    assumes "C.ide f" and "C.ide g" and "srcC g = trgC f"
    shows "(trgC g C g C f) C 𝖺C[trgC g, g, f] C (𝗅C-1[g] C ε g C 𝗋C[G (F g)] C f) C
           C.inv 𝖺C[G (F g), srcC g, f] C (G (F g) C 𝗅C-1[f] C ε f C 𝗋C[G (F f)]) C
           𝖺C[G (F g), G (F f), srcC f]
             = 𝗅C-1[g C f] C (ε g C ε f) C 𝗋C[G (F g) C G (F f)]"
    and "(𝗅C-1[g C f] C ε (g C f) C 𝗋C[G (F (g C f))]) C (GF.cmp (g, f) C srcC f)
           = 𝗅C-1[g C f] C (ε g C ε f) C 𝗋C[G (F g) C G (F f)]"
    proof -
      have "(trgC g C g C f) C
            𝖺C[trgC g, g, f] C
            (𝗅C-1[g] C ε g C 𝗋C[G (F g)] C f) C
            C.inv 𝖺C[G (F g), srcC g, f] C
            (G (F g) C 𝗅C-1[f] C ε f C 𝗋C[G (F f)]) C
            𝖺C[G (F g), G (F f), srcC f]
              = ((trgC g C g C f) C
                𝖺C[trgC g, g, f]) C
                (𝗅C-1[g] C ε g C 𝗋C[G (F g)] C f) C
                C.inv 𝖺C[G (F g), srcC g, f] C
                (G (F g) C 𝗅C-1[f] C ε f C 𝗋C[G (F f)]) C
                𝖺C[G (F g), G (F f), srcC f]"
        using assms C.comp_assoc by simp
      also have "... = 𝖺C[trgC g, g, f] C
                       (𝗅C-1[g] C ε g C 𝗋C[G (F g)] C f) C
                       C.inv 𝖺C[G (F g), srcC g, f] C
                       (G (F g) C 𝗅C-1[f] C ε f C 𝗋C[G (F f)]) C
                       𝖺C[G (F g), G (F f), srcC f]"
        using assms C.comp_cod_arr by simp
      also have "... = (𝖺C[trgC g, g, f] C
                       (𝗅C-1[g] C f)) C
                       (ε g C f) C
                       (𝗋C[G (F g)] C f) C
                       C.inv 𝖺C[G (F g), srcC g, f] C
                       (G (F g) C 𝗅C-1[f]) C
                       (G (F g) C ε f) C
                       ((G (F g) C 𝗋C[G (F f)]) C
                       𝖺C[G (F g), G (F f), srcC f])"
        using assms C.whisker_left C.whisker_right C.comp_assoc by simp
      also have "... = 𝗅C-1[g C f] C
                       (ε g C f) C
                       ((𝗋C[G (F g)] C f) C
                       C.inv 𝖺C[G (F g), srcC g, f]) C
                       (G (F g) C 𝗅C-1[f]) C
                       (G (F g) C ε f) C
                       𝗋C[G (F g) C G (F f)]"
        using assms G0_props C.lunit_hcomp C.runit_hcomp C.comp_assoc by simp
      also have "... = 𝗅C-1[g C f] C
                       (ε g C f) C
                       (((G (F g) C 𝗅C[f]) C
                       (G (F g) C 𝗅C-1[f])) C
                       (G (F g) C ε f)) C
                       𝗋C[G (F g) C G (F f)]"
        using assms G0_props C.triangle' C.comp_assoc by simp
      also have "... = 𝗅C-1[g C f] C
                       ((ε g C f) C
                       (G (F g) C ε f)) C
                       𝗋C[G (F g) C G (F f)]"
      proof -
        have "(G (F g) C 𝗅C[f]) C (G (F g) C 𝗅C-1[f]) = G (F g) C f"
          using assms C.whisker_left [of "G (F g)" "𝗅C[f]" "𝗅C-1[f]"] C.comp_arr_inv'
          by simp
        moreover have "... = C.cod (G (F g) C ε f)"
          using assms G0_props by auto
        moreover have "C.hseq (G (F g)) (ε f)"
          using assms G0_props by simp
        ultimately have "((G (F g) C 𝗅C[f]) C (G (F g) C 𝗅C-1[f])) C (G (F g) C ε f) =
                         (G (F g) C ε f)"
          using assms G0_props C.comp_cod_arr by presburger
        thus ?thesis
          using C.comp_assoc by simp
      qed
      also have "... = 𝗅C-1[g C f] C (ε g C ε f) C 𝗋C[G (F g) C G (F f)]"
        using assms C.interchange [of "ε g" "G (F g)" f "ε f"] C.comp_cod_arr C.comp_arr_dom
        by simp
      finally show "(trgC g C g C f) C
                    𝖺C[trgC g, g, f] C
                    (𝗅C-1[g] C ε g C 𝗋C[G (F g)] C f) C
                    C.inv 𝖺C[G (F g), srcC g, f] C
                    (G (F g) C 𝗅C-1[f] C ε f C 𝗋C[G (F f)]) C
                    𝖺C[G (F g), G (F f), srcC f]
                      = 𝗅C-1[g C f] C (ε g C ε f) C 𝗋C[G (F g) C G (F f)]"
        by simp
      show "(𝗅C-1[g C f] C ε (g C f) C 𝗋C[G (F (g C f))]) C (GF.cmp (g, f) C srcC f)
              = ..."
      proof -
        have "(𝗅C-1[g C f] C ε (g C f) C 𝗋C[G (F (g C f))]) C (GF.cmp (g, f) C srcC f) =
              𝗅C-1[g C f] C ε (g C f) C 𝗋C[G (F (g C f))] C (GF.cmp (g, f) C srcC f)"
          using assms C.comp_assoc by simp
        also have "... = 𝗅C-1[g C f] C (ε (g C f) C GF.cmp (g, f)) C 𝗋C[G (F g) C G (F f)]"
        proof -
          have "srcC (GF.cmp (g, f)) = srcC f"
            using assms G0_props by simp
          hence "𝗋C[G (F (g C f))] C (GF.cmp (g, f) C srcC f) =
                GF.cmp (g, f) C 𝗋C[G (F g) C G (F f)]"
            using assms C.runit_naturality [of "GF.cmp (g, f)"] C.VV.arr_charSbC C.VV.cod_charSbC
                  GF.cmp_simps'(1,4-5)
            by simp
          thus ?thesis
            using C.comp_assoc by simp
        qed
        also have "... = 𝗅C-1[g C f] C (ε g C ε f) C 𝗋C[G (F g) C G (F f)]"
          using assms ε_hcomp' by simp
        ultimately show ?thesis by simp
      qed
    qed

    lemma ε_respects_unit:
    assumes "C.obj a"
    shows "(a C IC.unit a) C 𝗋C-1[a] C 𝗅C[a] =
           (𝗅C-1[C.cod a] C ε a C 𝗋C[GF.map (C.dom a)]) C (GF.unit a C a)"
    proof -
      have "(𝗅C-1[C.cod a] C ε a C 𝗋C[GF.map (C.dom a)]) C (GF.unit a C a) =
            𝗅C-1[a] C ε a C 𝗋C[G (F a)] C (GF.unit a C a)"
        using assms C.comp_assoc by auto
      also have "... = (𝗅C-1[a] C ε a) C GF.unit a C 𝗋C[a]"
      proof -
        have "srcC (GF.unit a) = a"
          using assms GF.unit_simps(2) GF.map0_def [of a] G0_props
          by (simp add: C.obj_simps(1-2))
        thus ?thesis
          using assms C.runit_naturality [of "GF.unit a"] C.comp_assoc by simp
      qed
      also have "... = (a C ε a) C (𝗅C-1[G (F a)] C GF.unit a) C 𝗋C[a]"
      proof -
        have "𝗅C-1[a] C ε a = (a C ε a) C 𝗅C-1[G (F a)]"
          using assms C.lunit'_naturality [of "ε a"] by auto
        thus ?thesis
          using C.comp_assoc by simp
      qed
      also have "... = ((a C ε a) C (a C GF.unit a)) C 𝗅C-1[a] C 𝗋C[a]"
      proof -
        have "𝗅C-1[G (F a)] C GF.unit a = (a C GF.unit a) C 𝗅C-1[a]"
          using assms C.lunit'_naturality [of "GF.unit a"] G0_props C.obj_simps
          by (simp add: GF.map0_def)
        thus ?thesis
          using C.comp_assoc by simp
      qed
      also have "... = (a C ε a C GF.unit a) C 𝗅C-1[a] C 𝗋C[a]"
        using assms C.interchange [of a a "ε a" "GF.unit a"] by force
      also have "... = (a C ε a C GF.unit a) C 𝗋C-1[a] C 𝗅C[a]"
        using assms C.unitor_coincidence by simp
      also have "... = (a C IC.unit a) C 𝗋C-1[a] C 𝗅C[a]"
        using assms ε_inverts_GF_unit IC.unit_char' by simp
      finally show ?thesis by simp
    qed

    (*
     * The following clash with definitions made later in the context of
     * equivalence_pseudofunctor.  The clash is only flagged when -o export_theory
     * is invoked with "isabelle_build", as was pointed out by Fabian Huch in an email
     * of 3/31/2022.  The reason for the clash is not clear.
     * I have distinguished the constants here with a prime because it is unlikely
     * that they will be useful outside of the present theory, so the name is not
     * so important.
     *)

    abbreviation counit0'
    where "counit0'  λb. b"

    abbreviation counit1'
    where "counit1'  λg. 𝗅D-1[F (G g)] D η g D 𝗋D[g]"

    interpretation ε: pseudonatural_equivalence
                        VD HD 𝖺D 𝗂D srcD trgD VD HD 𝖺D 𝗂D srcD trgD
                        FG.map FG.cmp ID.map ID.cmp counit0' counit1'
    proof
      show "a. D.obj a  D.ide a"
        by auto
      show "f. D.ide f  D.iso (𝗅D-1[F (G f)] D η f D 𝗋D[f])"
        using D.iso_lunit' D.iso_runit η_in_hom
        by (intro D.isos_compose D.seqI) auto
      show "a. D.obj a  «a : srcD (FG.map a) D srcD (ID.map a)»"
        using D.obj_def G0_props(1) by auto
      show "f. D.ide f  «𝗅D-1[F (G f)] D η f D 𝗋D[f] :
                                ID.map f D srcD f D trgD f D FG.map f»"
        using G0_props
        by (intro D.comp_in_homI') auto
      show "μ. D.arr μ 
                (𝗅D-1[F (G (D.cod μ))] D η (D.cod μ) D 𝗋D[D.cod μ]) D (ID.map μ D srcD μ)
                  = (trgD μ D FG.map μ) D 𝗅D-1[F (G (D.dom μ))] D η (D.dom μ) D 𝗋D[D.dom μ]"
      proof -
        fix μ
        assume μ: "D.arr μ"
        have "(𝗅D-1[FG.map (D.cod μ)] D η (D.cod μ) D 𝗋D[D.cod μ]) D
              (ID.map μ D srcD μ)
                = 𝗅D-1[FG.map (D.cod μ)] D η μ D 𝗋D[D.dom μ]"
        proof -
          have "(𝗅D-1[FG.map (D.cod μ)] D η (D.cod μ) D 𝗋D[D.cod μ]) D
                (ID.map μ D srcD μ)
                  = 𝗅D-1[FG.map (D.cod μ)] D (η (D.cod μ) D μ) D 𝗋D[D.dom μ]"
            using μ D.runit_naturality D.comp_assoc by simp
          also have "... = 𝗅D-1[FG.map (D.cod μ)] D η μ D 𝗋D[D.dom μ]"
            using μ η_naturality(1) by simp
          finally show ?thesis by blast
        qed
        also have "... = (trgD μ D FG.map μ) D 𝗅D-1[FG.map (D.dom μ)] D η (D.dom μ) D
                         𝗋D[D.dom μ]"
        proof -
          have "𝗅D-1[FG.map (D.cod μ)] D η μ D 𝗋D[D.dom μ] =
                𝗅D-1[FG.map (D.cod μ)] D (FG.map μ D η (D.dom μ)) D 𝗋D[D.dom μ]"
            using μ η_naturality(2) D.comp_assoc by simp
          also have "... = (𝗅D-1[FG.map (D.cod μ)] D FG.map μ) D η (D.dom μ) D
                           𝗋D[D.dom μ]"
            using D.comp_assoc by simp
          also have "... = ((trgD μ D FG.map μ) D 𝗅D-1[FG.map (D.dom μ)]) D
                           η (D.dom μ) D 𝗋D[D.dom μ]"
            using μ D.lunit'_naturality [of "FG.map μ"] G0_props by simp
          also have "... = (trgD μ D FG.map μ) D 𝗅D-1[FG.map (D.dom μ)] D
                           η (D.dom μ) D 𝗋D[D.dom μ]"
            using D.comp_assoc by simp
          finally show ?thesis by blast
        qed
        finally show "(𝗅D-1[F (G (D.cod μ))] D η (D.cod μ) D 𝗋D[D.cod μ]) D
                      (ID.map μ D srcD μ)
                        = (trgD μ D FG.map μ) D 𝗅D-1[F (G (D.dom μ))] D η (D.dom μ) D
                          𝗋D[D.dom μ]"
          by simp
      qed
      show "f g. D.ide f; D.ide g; srcD g = trgD f
                     (trgD g D FG.cmp (g, f)) D
                        𝖺D[trgD g, FG.map g, FG.map f] D
                        (𝗅D-1[F (G g)] D η g D 𝗋D[g] D FG.map f) D
                        D.inv 𝖺D[ID.map g, srcD g, FG.map f] D
                        (ID.map g D 𝗅D-1[F (G f)] D η f D 𝗋D[f]) D
                        𝖺D[ID.map g, ID.map f, srcD f]
                          = (𝗅D-1[F (G (g D f))] D η (g D f) D
                            𝗋D[g D f]) D
                            (ID.cmp (g, f) D srcD f)"
        using η_respects_comp by simp
      show "a. D.obj a  (a D FG.unit a) D 𝗋D-1[a] D 𝗅D[a] =
                            (𝗅D-1[F (G a)] D η a D 𝗋D[a]) D (ID.unit a D a)"
        using η_respects_unit by auto
      show "a. D.obj a  D.equivalence_map a"
        using D.obj_is_equivalence_map by simp
    qed

    abbreviation unit0'
    where "unit0'  λa. a"

    abbreviation unit1'
    where "unit1'  λf. 𝗅C-1[f] C ε f C 𝗋C[G (F f)]"

    interpretation η: pseudonatural_equivalence
                           VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC
                           IC.map IC.cmp GF.map GF.cmp unit0' unit1'
    proof
      show "a. C.obj a  C.ide a"
        by auto
      show "f. C.ide f  C.iso (𝗅C-1[f] C ε f C 𝗋C[G (F f)])"
        using C.iso_runit C.iso_lunit'
        by (intro C.isos_compose) auto
      show "a. C.obj a  «a : srcC (IC.map a) C srcC (GF.map a)»"
        by (simp_all add: C.obj_simps(1-3) G0_props(2))
      show "f. C.ide f  «𝗅C-1[f] C ε f C 𝗋C[G (F f)] :
                               GF.map f C srcC f C trgC f C IC.map f»"
        using G0_props by auto
      show "μ. C.arr μ 
                  (𝗅C-1[C.cod μ] C ε (C.cod μ) C 𝗋C[G (F (C.cod μ))]) C
                  (GF.map μ C srcC μ)
                    = (trgC μ C IC.map μ) C 𝗅C-1[C.dom μ] C ε (C.dom μ) C
                      𝗋C[G (F (C.dom μ))]"
      proof -
        fix μ
        assume μ: "C.arr μ"
        have "(𝗅C-1[C.cod μ] C ε (C.cod μ) C 𝗋C[GF.map (C.cod μ)]) C
              (GF.map μ C srcC μ)
                = 𝗅C-1[C.cod μ] C ε μ C 𝗋C[GF.map (C.dom μ)]"
        proof -
          have "(𝗅C-1[C.cod μ] C ε (C.cod μ) C 𝗋C[GF.map (C.cod μ)]) C
                (GF.map μ C srcC μ)
                  = 𝗅C-1[C.cod μ] C ε (C.cod μ) C 𝗋C[GF.map (C.cod μ)] C
                    (GF.map μ C srcC μ)"
            using C.comp_assoc by simp
          also have "... = 𝗅C-1[C.cod μ] C (ε (C.cod μ) C GF.map μ) C 𝗋C[GF.map (C.dom μ)]"
            using μ C.runit_naturality [of "GF.map μ"] G0_props C.comp_assoc by simp
          also have "... = 𝗅C-1[C.cod μ] C ε μ C 𝗋C[GF.map (C.dom μ)]"
            using μ ε_naturality(1) [of μ] by simp
          finally show ?thesis by blast
        qed
        also have "... = (trgC μ C IC.map μ) C 𝗅C-1[C.dom μ] C ε (C.dom μ) C
                         𝗋C[GF.map (C.dom μ)]"
        proof -
          have "... = 𝗅C-1[C.cod μ] C (μ C ε (C.dom μ)) C 𝗋C[GF.map (C.dom μ)]"
            using μ ε_naturality(2) [of μ] by simp
          also have "... = (𝗅C-1[C.cod μ] C μ) C ε (C.dom μ) C 𝗋C[GF.map (C.dom μ)]"
            using C.comp_assoc by simp
          also have "... = ((trgC μ C IC.map μ) C 𝗅C-1[C.dom μ]) C ε (C.dom μ) C
                           𝗋C[GF.map (C.dom μ)]"
            using μ C.lunit'_naturality [of μ] by simp
          also have "... = (trgC μ C IC.map μ) C 𝗅C-1[C.dom μ] C ε (C.dom μ) C
                           𝗋C[GF.map (C.dom μ)]"
            using C.comp_assoc by simp
          finally show ?thesis by blast
        qed
        finally show "(𝗅C-1[C.cod μ] C ε (C.cod μ) C 𝗋C[G (F (C.cod μ))]) C
                      (GF.map μ C srcC μ)
                        = (trgC μ C IC.map μ) C 𝗅C-1[C.dom μ] C ε (C.dom μ) C
                          𝗋C[G (F (C.dom μ))]"
          by simp
      qed
      show "f g. C.ide f; C.ide g; srcC g = trgC f 
                    (trgC g C IC.cmp (g, f)) C
                    𝖺C[trgC g, IC.map g, IC.map f] C
                    (𝗅C-1[g] C ε g C 𝗋C[G (F g)] C IC.map f) C
                    C.inv 𝖺C[GF.map g, srcC g, IC.map f] C
                    (GF.map g C 𝗅C-1[f] C ε f C 𝗋C[G (F f)]) C
                    𝖺C[GF.map g, GF.map f, srcC f]
                      = (𝗅C-1[g C f] C ε (g C f) C 𝗋C[G (F (g C f))]) C
                        (GF.cmp (g, f) C srcC f)"
        using ε_respects_comp by simp
      show "a. C.obj a  (a C IC.unit a) C 𝗋C-1[a] C 𝗅C[a] =
                             (𝗅C-1[a] C ε a C 𝗋C[G (F a)]) C (GF.unit a C a)"
        using ε_respects_unit by auto
      show "a. C.obj a  C.equivalence_map a"
        using C.obj_is_equivalence_map by simp
    qed

    interpretation EQ: equivalence_of_bicategories
                         VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                         F Φ G ΦG.map unit0' unit1' counit0' counit1'
      ..

    lemma extends_to_equivalence_of_bicategories:
    shows "equivalence_of_bicategories VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
             F Φ G ΦG.map unit0' unit1' counit0' counit1'"
      ..

  end

  subsection "Equivalence Pseudofunctors Extend to Equivalences of Bicategories"

  text ‹
    Now we put the pieces together and prove that an arbitrary equivalence pseudofunctor extends
    to an equivalence of bicategories.
  ›

  context equivalence_pseudofunctor
  begin

    text ‹
      Define a set of objects U› of C› by choosing a representative of each equivalence
      class of objects having the same image under the object map of the given equivalence
      pseudofunctor.  Then U› is obviously dense, because every object of C› belongs to
      such an equivalence class.
    ›

    definition U
    where "U = {a. C.obj a  a = (SOME a'. C.obj a'  map0 a' = map0 a)}"

    lemma U_dense:
    assumes "C.obj a"
    shows "a'  U. C.equivalent_objects a a'"
    proof -
      let ?a' = "SOME a'. C.obj a'  map0 a' = map0 a"
      have "a'. C.obj a'  map0 a' = map0 a"
        using assms by auto
      hence 1: "?a'  U  map0 ?a' = map0 a"
        using assms U_def someI_ex [of "λa'. C.obj a'  map0 a' = map0 a"] by simp
      hence "C.equivalent_objects ?a' a"
        using D.equivalent_objects_reflexive reflects_equivalent_objects [of ?a' a]
        by (simp add: U_def assms)
      thus ?thesis
        using 1 C.equivalent_objects_symmetric by auto
    qed

    text ‹
      Take V› to be the collection of images of all objects of C› under the given equivalence
      pseudofunctor.  Since equivalence pseudofunctors are biessentially surjective on objects,
      V› is dense.  Moreover, by construction, the object map of the given equivalence
      pseudofunctor is a bijection from U› to V›.
    ›

    definition V
    where "V = map0 ` Collect C.obj"

    lemma V_dense:
    assumes "D.obj b"
    shows "b'. b'  map0 ` Collect C.obj  D.equivalent_objects b b'"
      using assms biessentially_surjective_on_objects D.equivalent_objects_symmetric
      by blast

    lemma bij_betw_U_V:
    shows "bij_betw map0 U V"
    proof -
      have "inj_on map0 U"
        using U_def by (intro inj_onI) simp
      moreover have "map0 ` U = V"
      proof
        show "map0 ` U  V"
          using U_def V_def by auto
        show "V  map0 ` U"
        proof
          fix b
          assume b: "b  V"
          obtain a where a: "C.obj a  map0 a = b"
            using b V_def by auto
          let ?a' = "SOME a'. C.obj a'  map0 a' = map0 a"
          have 1: "C.obj ?a'  map0 ?a' = b"
            using a someI_ex [of "λa'. C.obj a'  map0 a' = map0 a"] by auto
          moreover have "?a' = (SOME a''. C.obj a''  map0 a'' = map0 ?a')"
            using a 1 by simp
          ultimately have "?a'  U"
            unfolding U_def by simp
          thus "b  map0 ` U"
             using a 1 U_def
             by (metis (no_types, lifting) image_eqI)
        qed
      qed
      ultimately show ?thesis
        using bij_betw_def [of map0 U V] by simp
    qed

    abbreviation (input) ArrU
    where "ArrU  λμ. C.arr μ  srcC μ  U  trgC μ  U"

    interpretation CU: subbicategory VC HC 𝖺C 𝗂C srcC trgC ArrU
      using C.𝔩_ide_simp C.𝔯_ide_simp
      apply unfold_locales
                  apply auto
       apply (metis C.comp_ide_self C.ide_src C.src_cod C.src_dom)
      by (metis C.trg.as_nat_trans.is_natural_2 C.trg.as_nat_trans.naturality C.trg_cod)

    interpretation CU: dense_subbicategory VC HC 𝖺C 𝗂C srcC trgC U  (* 15 sec *)
    proof
      show "a. C.obj a  a'. a'  U  C.equivalent_objects a' a"
        using U_dense C.equivalent_objects_symmetric by blast
      (* TODO: The above inference is trivial, but qed consumes 15 seconds! *)
    qed (* 25 sec *)

    abbreviation (input) ArrV
    where "ArrV  λμ. D.arr μ  srcD μ  V  trgD μ  V"

    interpretation DV: subbicategory VD HD 𝖺D 𝗂D srcD trgD ArrV
      using D.𝔩_ide_simp D.𝔯_ide_simp
      apply unfold_locales
                  apply auto
       apply (metis D.comp_ide_self D.ide_src D.src_cod D.src_dom)
      by (metis D.trg.as_nat_trans.is_natural_2 D.trg.as_nat_trans.naturality D.trg_cod)

    interpretation DV: dense_subbicategory VD HD 𝖺D 𝗂D srcD trgD V  (* 25 sec *)
      using V_dense D.equivalent_objects_def D.equivalent_objects_symmetric V_def
      by unfold_locales metis (* 35 sec -- time is similar doing it this way. *)

    interpretation FU: restricted_pseudofunctor
                         VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F Φ
                         λμ. C.arr μ  srcC μ  U  trgC μ  U
      ..

    interpretation FUV: corestricted_pseudofunctor
                         CU.comp CU.hcomp CU.𝖺 𝗂C CU.src CU.trg
                         VD HD 𝖺D 𝗂D srcD trgD FU.map FU.cmp λa. a  V
    proof
      show "μ. CU.arr μ  DV.arr (FU.map μ)"
      proof -
        fix μ
        assume μ: "CU.arr μ"
        have "srcC μ  U  trgC μ  U"
          using μ CU.arr_charSbC by simp
        moreover have "a. a  U  FU.map0 a = map0 a"
        proof -
          fix a :: 'a
          assume a: "a  U"
          hence 1: "C.obj a"
            using U_def by blast
          have "srcC a = a"
            using a U_def by blast
          thus "FU.map0 a = map0 a"
            using a 1 C.obj_simps(1,3) CU.arr_charSbC FU.map0_def map0_def by presburger
        qed
        ultimately have "FU.map0 (srcC μ)  V  FU.map0 (trgC μ)  V"
          using μ bij_betw_U_V bij_betw_def
          by (metis (no_types, lifting) image_eqI)
        hence "srcD (FU.map μ)  V  trgD (FU.map μ)  V"
          using μ FU.map0_def CU.src_def CU.trg_def FU.preserves_src FU.preserves_trg
          by auto
        thus "DV.arr (FU.map μ)"
          using μ DV.arr_charSbC [of "FU.map μ"] by blast
      qed
    qed

    interpretation FUV: equivalence_pseudofunctor
                         CU.comp CU.hcomp CU.𝖺 𝗂C CU.src CU.trg
                         DV.comp DV.hcomp DV.𝖺 𝗂D DV.src DV.trg
                         FUV.map FUV.cmp
    proof
      show "f f'. CU.par f f'; FU.map f = FU.map f'  f = f'"
        using CU.arr_charSbC CU.dom_charSbC CU.cod_charSbC
        by (metis (no_types, lifting) is_faithful)
      show "a b g. CU.obj a; CU.obj b; DV.in_hhom g (FUV.map0 a) (FUV.map0 b);
                      DV.ide g
                        f. CU.in_hhom f a b  CU.ide f  DV.isomorphic (FU.map f) g"
      proof -
        fix a b g
        assume a: "CU.obj a" and b: "CU.obj b"
        assume g: "DV.in_hhom g (FUV.map0 a) (FUV.map0 b)" and ide_g: "DV.ide g"
        have 1: "f. «f : a C b»  C.ide f  D.isomorphic (F f) g"
        proof -
          have "C.obj a  C.obj b"
            using a b CU.obj_char by simp
          moreover have "«g : map0 a D map0 b»"
            using a b g DV.in_hhom_def DV.arr_charSbC DV.src_def DV.trg_def
            by (intro D.in_hhomI) auto
          moreover have "D.ide g"
            using ide_g DV.ide_charSbC DV.arr_charSbC by simp
          ultimately show ?thesis
            using locally_essentially_surjective by simp
        qed
        obtain f where f: "«f : a C b»  C.ide f  D.isomorphic (F f) g"
          using 1 by blast
        have 2: "CU.in_hhom f a b"
          using f a b CU.arr_charSbC CU.obj_char CU.src_def CU.trg_def by fastforce
        moreover have "CU.ide f"
          using f 2 CU.ide_charSbC CU.arr_charSbC by auto
        moreover have "DV.isomorphic (FUV.map f) g"
        proof -
          obtain φ where φ: "«φ : F f D g»  D.iso φ"
            using f D.isomorphic_def by auto
          have 3: "DV.in_hom φ (FUV.map f) g"
          proof -
            have "DV.arr φ"
              using f g φ 2 DV.arr_charSbC
              by (metis (no_types, lifting) D.arrI D.vconn_implies_hpar(1-4) DV.ideD(1) ide_g)
            thus ?thesis
              using φ 2 DV.dom_charSbC DV.cod_charSbC
              by (intro DV.in_homI) auto
          qed
          moreover have "DV.iso φ"
            using φ DV.iso_charSbC DV.arr_charSbC 3 by fastforce
          ultimately show ?thesis
            using DV.isomorphic_def by auto
        qed
        ultimately show "f. CU.in_hhom f a b  CU.ide f  DV.isomorphic (FUV.map f) g"
          by auto
      qed
      show "f f' ν. CU.ide f; CU.ide f'; CU.src f = CU.src f'; CU.trg f = CU.trg f';
                      DV.in_hom ν (FU.map f) (FU.map f')
                           μ. CU.in_hom μ f f'  FU.map μ = ν"
      proof -
        fix f f' ν
        assume f: "CU.ide f" and f': "CU.ide f'"
        and eq_src: "CU.src f = CU.src f'" and eq_trg: "CU.trg f = CU.trg f'"
        and ν: "DV.in_hom ν (FU.map f) (FU.map f')"
        have 1: "C.ide f  C.ide f'"
          using f f' CU.ide_charSbC CU.arr_charSbC by simp
        have 2: "μ. «μ : f C f'»  F μ = ν"
        proof -
          have "srcC f = srcC f'  trgC f = trgC f'"
            using f f' 1 eq_src eq_trg CU.src_def CU.trg_def by simp
          moreover have "«ν : F f D F f'»"
            using ν DV.in_hom_charSbC DV.arr_charSbC by auto
          ultimately show ?thesis
            using 1 locally_full by simp
        qed
        obtain μ where μ: "«μ : f C f'»  F μ = ν"
          using 2 by auto
        have 3: "CU.in_hom μ f f'"
          using μ f f' CU.in_hom_charSbC CU.ide_charSbC CU.arr_charSbC by auto
        moreover have "FU.map μ = ν"
          using μ 3 by auto
        ultimately show "μ. CU.in_hom μ f f'  FU.map μ = ν"
          by auto
      qed
      show "b. DV.obj b  a. CU.obj a  DV.equivalent_objects (FUV.map0 a) b"
      proof -
        fix b
        assume b: "DV.obj b"
        obtain a where a: "C.obj a  map0 a = b"
          using b DV.obj_char DV.arr_charSbC V_def by auto
        have 1: "D.obj b  b  V"
          using a b DV.obj_char V_def by auto
        obtain a' where a': "a'  U  C.equivalent_objects a' a"
          using a U_dense C.equivalent_objects_symmetric by blast
        have obj_a': "CU.obj a'"
          using a' U_def CU.obj_char CU.arr_charSbC by fastforce
        moreover have "DV.equivalent_objects (FUV.map0 a') b"
        proof -
          have "D.equivalent_objects (map0 a') (map0 a)"
            using a' preserves_equivalent_objects by simp
          hence 2: "D.equivalent_objects (map0 a') b"
            using a D.equivalent_objects_transitive by simp
          obtain e where e: "«e : map0 a' D b»  D.equivalence_map e"
            using 2 D.equivalent_objects_def by auto
          have 3: "D.obj (map0 a')  map0 a'  V"
            using a' e U_def V_def by simp
          moreover have e_in_hhom: "DV.in_hhom e (FUV.map0 a') b"
          proof
            show 4: "DV.arr e"
              using 1 3 a e b DV.arr_charSbC DV.obj_char V_def
              by (metis (no_types, lifting) D.in_hhomE)
            show "DV.src e = FUV.map0 a'"
              using e 4 DV.src_def FUV.map0_def obj_a' map0_def by auto
            show "DV.trg e = b"
              using e 4 DV.trg_def by auto
          qed
          moreover have "DV.equivalence_map e"
          proof -
            obtain d η ε where d: "equivalence_in_bicategory VD HD 𝖺D 𝗂D srcD trgD e d η ε"
              using e D.equivalence_map_def by auto
            interpret e: equivalence_in_bicategory VD HD 𝖺D 𝗂D srcD trgD e d η ε
              using d by simp
            have "equivalence_in_bicategory DV.comp DV.hcomp DV.𝖺 𝗂D DV.src DV.trg e d η ε"
            proof
              show ide_e: "DV.ide e"
                using e e_in_hhom DV.ide_charSbC by auto
              show ide_d: "DV.ide d"
                using d ide_e e.antipar DV.ide_charSbC DV.arr_charSbC by simp
              show 4: "DV.in_hom η (DV.src e) (DV.hcomp d e)"
              proof -
                have "DV.hseq d e"
                  using ide_d ide_e e.antipar DV.src_def DV.trg_def by simp
                thus ?thesis
                  using ide_d ide_e e.unit_in_hom DV.arr_charSbC DV.ide_charSbC
                        DV.dom_charSbC DV.cod_charSbC DV.src_def DV.trg_def
                  by (intro DV.in_homI) auto
              qed
              show 5: "DV.in_hom ε (DV.hcomp e d) (DV.src d)"
              proof -
                have "DV.hseq e d"
                  using ide_d ide_e e.antipar DV.src_def DV.trg_def by simp
                thus ?thesis
                  using ide_d ide_e e.antipar e.counit_in_hom DV.arr_charSbC DV.ide_charSbC
                        DV.dom_charSbC DV.cod_charSbC DV.src_def DV.trg_def
                  by (intro DV.in_homI) auto
              qed
              show "DV.iso η"
                using 4 DV.iso_charSbC DV.arr_charSbC by fastforce
              show "DV.iso ε"
                using 5 DV.iso_charSbC DV.arr_charSbC by fastforce
            qed
            thus ?thesis
              using DV.equivalence_map_def by auto
          qed
          ultimately show ?thesis
            using DV.equivalent_objects_def by auto
        qed
        ultimately show "a. CU.obj a  DV.equivalent_objects (FUV.map0 a) b" by auto
      qed
    qed

    interpretation FUV: equivalence_pseudofunctor_bij_on_obj
                         CU.comp CU.hcomp CU.𝖺 𝗂C CU.src CU.trg
                         DV.comp DV.hcomp DV.𝖺 𝗂D DV.src DV.trg
                         FUV.map FUV.cmp
    proof
      have "Collect CU.obj = U"
        using CU.obj_char CU.arr_charSbC U_def by fastforce
      moreover have "Collect DV.obj = V"
        using DV.obj_char DV.arr_charSbC V_def D.obj_def' map0_simps(1) by auto
      moreover have "a. a  Collect CU.obj  FUV.map0 a = map0 a"
        using FUV.map0_def CU.obj_char CU.arr_charSbC DV.src_def
              FU.map0_simp FUV.map0_simp by auto
      ultimately show "bij_betw FUV.map0 (Collect CU.obj) (Collect DV.obj)"
        using bij_betw_U_V
        by (simp add: bij_betw_U_V bij_betw_cong)
    qed

    interpretation EQUV: equivalence_of_bicategories  (* 25 sec *)
                          DV.comp DV.hcomp DV.𝖺 𝗂D DV.src DV.trg
                          CU.comp CU.hcomp CU.𝖺 𝗂C CU.src CU.trg
                          FUV.map FUV.cmp FUV.G FUVG
                          FUV.unit0' FUV.unit1' FUV.counit0' FUV.counit1'
      using FUV.extends_to_equivalence_of_bicategories by blast  (* 18 sec, mostly "by" *)

    text ‹
      Now compose EQUV with the equivalence from DV to D› and the converse of the equivalence
      from CU to C›.  The result is an equivalence of bicategories from C› to D›.
    ›

    interpretation EQC: equivalence_of_bicategories
                           VC HC 𝖺C 𝗂C srcC trgC CU.comp CU.hcomp CU.𝖺 𝗂C CU.src CU.trg
                           CU.E CUE CU.P CUP
                           CU.unit0 CU.unit1 CU.counit0 CU.counit1
      using CU.induces_equivalence by simp

    interpretation EQC': converse_equivalence_of_bicategories
                           VC HC 𝖺C 𝗂C srcC trgC CU.comp CU.hcomp CU.𝖺 𝗂C CU.src CU.trg
                           CU.E CUE CU.P CUP
                           CU.unit0 CU.unit1 CU.counit0 CU.counit1
      ..

    interpretation EQD: equivalence_of_bicategories
                           VD HD 𝖺D 𝗂D srcD trgD DV.comp DV.hcomp DV.𝖺 𝗂D DV.src DV.trg
                           DV.E DVE DV.P DVP
                           DV.unit0 DV.unit1 DV.counit0 DV.counit1
      using DV.induces_equivalence by simp

    interpretation EQUVoEQC': composite_equivalence_of_bicategories  (* 35 sec *)
                               DV.comp DV.hcomp DV.𝖺 𝗂D DV.src DV.trg
                               CU.comp CU.hcomp CU.𝖺 𝗂C CU.src CU.trg
                               VC HC 𝖺C 𝗂C srcC trgC
                               FUV.map FUV.cmp FUV.G FUVG
                               CU.P CUP CU.E CUE
                               FUV.unit0' FUV.unit1' FUV.counit0' FUV.counit1'
                               EQC'.unit0 EQC'.unit1 EQC'.counit0 EQC'.counit1
      ..  (* 55 sec *)

    interpretation EQDoEQUVoEQC': composite_equivalence_of_bicategories  (* 30 sec *)
                                   VD HD 𝖺D 𝗂D srcD trgD
                                   DV.comp DV.hcomp DV.𝖺 𝗂D DV.src DV.trg
                                   VC HC 𝖺C 𝗂C srcC trgC
                                   DV.E DVE DV.P DVP
                                   EQUVoEQC'.left_map EQUVoEQC'.left_cmp
                                   EQUVoEQC'.right_map EQUVoEQC'.right_cmp
                                   DV.unit0 DV.unit1 DV.counit0 DV.counit1
                                   EQUVoEQC'.unit0 EQUVoEQC'.unit1
                                   EQUVoEQC'.counit0 EQUVoEQC'.counit1
      ..  (* 50 sec *)

    lemma induces_equivalence_of_bicategories:
    shows "equivalence_of_bicategories VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
             EQDoEQUVoEQC'.left_map EQDoEQUVoEQC'.left_cmp
             EQDoEQUVoEQC'.right_map EQDoEQUVoEQC'.right_cmp
             EQDoEQUVoEQC'.unit0 EQDoEQUVoEQC'.unit1
             EQDoEQUVoEQC'.counit0 EQDoEQUVoEQC'.counit1"
      ..

    lemma left_map_simp:
    assumes "C.arr μ"
    shows "EQDoEQUVoEQC'.left_map μ = DV.E (F (CU.P μ))"
      using assms by simp

    lemma right_map_simp:
    assumes "D.arr ν"
    shows "EQDoEQUVoEQC'.right_map ν = CU.E (FUV.G (DV.P ν))"
      using assms by simp

    lemma unit0_simp:
    assumes "C.obj a"
    shows "EQDoEQUVoEQC'.unit0 a =
           CU.E (FUV.G (DV0 (DV.src (F (CU.P a))))) C CU.E (CU.P0 (srcC a))
             C EQC'.unit0 a"
      using assms EQDoEQUVoEQC'.unit0_simp EQUVoEQC'.unit0_simp
            EQUVoEQC'.FH.map0_def CU.src_def
      by auto

    (*
     * TODO: Expansions of the other parts of the equivalence are not comprehensible
     * without interpreting a bunch of other locales.  I don't know that there is any particular
     * value to doing that.  The above show that the pseudofunctors between C and D and the
     * components of the unit of the equivalence are as expected.
     *)

    text ‹
      We've now got an equivalence of bicategories between C› and D›, but it involves
      EQDoEQUVoEQC'.left_map› and not the originally given equivalence pseudofunctor F›.
      However, we can patch things up by showing that EQDoEQUVoEQC'.left_map› is pseudonaturally
      equivalent to F›.  From this, we may conclude, using the fact that equivalences of
      bicategories respect pseudonatural equivalence, that there is an equivalence of bicategories
      between C› and D› that involves F› and EQDoEQUVoEQC'.right_map›, rather than
      EQDoEQUVoEQC'.left_map› and EQDoEQUVoEQC'.right_map›.
    ›

    abbreviation τ0
    where "τ0 a  F (CU0 a)"

    abbreviation τ1
    where "τ1 f  D.inv (Φ (CU0 (trgC f), CU.P f)) D F (CU1 f) D Φ (f, CU0 (srcC f))"

    lemma τ0_in_hom [intro]:
    assumes "C.obj a"
    shows "«τ0 a : map0 (CU.P0 a) D map0 a»"
    and "«τ0 a : τ0 a D τ0 a»"
    proof -
      show "«τ0 a : map0 (CU.P0 a) D map0 a»"
      proof
        show "D.arr (τ0 a)"
          using assms by simp
        show "srcD (τ0 a) = map0 (CU.P0 a)"
          using assms map0_def CU.counit.ide_map0_obj CU.equivalence_data_simpsB(7) C.ideD(1)
                preserves_src
          by presburger
        show "trgD (τ0 a) = map0 a"
          using assms by auto
      qed
      show "«τ0 a : τ0 a D τ0 a»"
        using assms by auto
    qed

    lemma τ0_simps [simp]:
    assumes "C.obj a"
    shows "D.ide (τ0 a)"
    and "srcD (τ0 a) = map0 (CU.P0 a)" and "trgD (τ0 a) = map0 a"
      using assms τ0_in_hom(1) [of a] by auto

    lemma τ1_in_hom [intro]:
    assumes "C.ide f"
    shows "«τ1 f : map0 (CU.P0 (srcC f)) D map0 (trgC f)»"
    and "«τ1 f : F f D τ0 (srcC f) D τ0 (trgC f) D F (CU.P f)»"
    proof -
      show 1: "«τ1 f : F f D τ0 (srcC f) D τ0 (trgC f) D F (CU.P f)»"
      proof (intro D.comp_in_homI)
        show "«Φ (f, CU.d (srcC f)) : F f D τ0 (srcC f) D F (f C CU0 (srcC f))»"
          using assms by auto
        show "«F (CU1 f) : F (f C CU0 (srcC f)) D F (CU0 (trgC f) C CU.P f)»"
          using assms CU.counit1_in_hom [of f] CU.P_def by auto
        show "«D.inv (Φ (CU0 (trgC f), CU.P f)) :
                   F (CU0 (trgC f) C CU.P f) D τ0 (trgC f) D F (CU.P f)»"
          using assms C.VV.arr_charSbC C.VV.ide_charSbC Φ.components_are_iso
          by (metis (no_types, lifting) CU.P_def CU.counit.ide_map0_obj CU.counit1_simps(1,5)
              CU.equivalence_data_simpsB(2) C.hseqE C.ide_hcomp C.obj_trg D.arr_cod D.inv_in_homI
              cmp_components_are_iso cmp_in_hom(2) preserves_cod preserves_reflects_arr)
      qed
      show "«τ1 f : map0 (CU.P0 (srcC f)) D map0 (trgC f)»"
        using assms 1 map0_def [of "CU.P0 (srcC f)"] CU.emb.map0_simp CU.P0_props
        apply (intro D.in_hhomI)
          apply auto
        using D.src_dom [of "τ1 f"]
         apply (elim D.in_homE)
         apply auto
        using D.trg_dom [of "τ1 f"]
        apply (elim D.in_homE)
        by auto
    qed

    lemma τ1_simps [simp]:
    assumes "C.ide f"
    shows "D.arr (τ1 f)"
    and "srcD (τ1 f) = map0 (CU.P0 (srcC f))" and "trgD (τ1 f) = map0 (trgC f)"
    and "D.dom (τ1 f) = F f D τ0 (srcC f)" and "D.cod (τ1 f) = τ0 (trgC f) D F (CU.P f)"
      using assms τ1_in_hom by blast+

    lemma iso_τ1:
    assumes "C.ide f"
    shows "D.iso (τ1 f)"
    proof -
      have "C.VV.ide (CU0 (trgC f), CU.P f)"
      proof -
        have "C.VV.arr (CU0 (trgC f), CU.P f)"
          using assms CU.equivalence_data_simpsB(7) C.ideD(1) by auto
        moreover have "C.VxV.ide (CU0 (trgC f), CU.P f)"
          using assms CU.prj.preserves_ide [of f] CU.ide_charSbC by simp
        ultimately show ?thesis
          using assms C.VV.ide_charSbC [of "(CU0 (trgC f), CU.P f)"] by blast
      qed
      hence "D.iso (Φ (CU.d (trgC f), CU.P f))"
        by simp
      thus ?thesis
        using assms C.VV.ide_charSbC C.VV.arr_charSbC Φ.components_are_iso
        by (intro D.isos_compose) auto
    qed

    interpretation τ: pseudonatural_equivalence
                          VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                          EQDoEQUVoEQC'.left_map EQDoEQUVoEQC'.left_cmp F Φ
                          τ0 τ1
    proof
      show "a. C.obj a  D.ide (τ0 a)"
        by simp
      show "a. C.obj a  D.equivalence_map (τ0 a)"
        using CU.counit.components_are_equivalences preserves_equivalence_maps by blast
      show "a. C.obj a  «τ0 a : srcD (EQDoEQUVoEQC'.left_map a) D srcD (F a)»"
      proof -
        fix a
        assume a: "C.obj a"
        show "«τ0 a : srcD (EQDoEQUVoEQC'.left_map a) D srcD (F a)»"
        proof
          show "D.arr (τ0 a)"
            using a by simp
          show "srcD (τ0 a) = srcD (EQDoEQUVoEQC'.left_map a)"
          proof -
            have "DV.arr (F (CU.P a))"
            proof -
              have "srcD (F (CU.P a))  V"
              proof -
                have "srcD (F (CU.P a)) = map0 (CU.prj.map0 a)"
                  using a by auto
                moreover have "C.obj (CU.prj.map0 a)"
                  using a CU.obj_char [of "CU.prj.map0 a"] CU.prj.map0_simps(1) by auto
                ultimately show ?thesis
                  using V_def by simp
              qed
              moreover have "trgD (F (CU.P a))  V"
              proof -
                have "trgD (F (CU.P a)) = map0 (CU.prj.map0 a)"
                  using a by auto
                moreover have "C.obj (CU.prj.map0 a)"
                  using a CU.obj_char [of "CU.prj.map0 a"] CU.prj.map0_simps(1) by auto
                ultimately show ?thesis
                  using V_def by simp
              qed
              ultimately show ?thesis
                using a DV.arr_charSbC by fastforce
            qed
            thus ?thesis
              using a DV.emb.map0_def DV.emb.map_def DV.src_def C.obj_simps(1-2) CU.P_simpsB(2)
              by (simp add: CU.P0_props(1))
          qed
          show "trgD (F (CU.d a)) = srcD (F a)"
            using a by auto
        qed
      qed
      show "f. C.ide f 
                   «τ1 f : F f D τ0 (srcC f) D τ0 (trgC f) D EQDoEQUVoEQC'.left_map f»"
      proof -
        fix f
        assume f: "C.ide f"
        show "«τ1 f : F f D τ0 (srcC f) D τ0 (trgC f) D EQDoEQUVoEQC'.left_map f»"
        proof -
          have "DV.arr (F (CU.P f))"
            using f FUV.preserves_arr by auto
          hence "EQDoEQUVoEQC'.left_map f = F (CU.P f)"
            using f EQUVoEQC'.FH.map0_def DV.emb.map_def by simp
          thus ?thesis
            using f by auto
        qed
      qed
      show "f. C.ide f  D.iso (τ1 f)"
        using iso_τ1 by simp
      show "μ. C.arr μ 
                   τ1 (C.cod μ) D (F μ D τ0 (srcC μ)) =
                   (τ0 (trgC μ) D EQDoEQUVoEQC'.left_map μ) D τ1 (C.dom μ)"
      proof -
        fix μ
        assume μ: "C.arr μ"
        show "τ1 (C.cod μ) D (F μ D τ0 (srcC μ))
                = (τ0 (trgC μ) D EQDoEQUVoEQC'.left_map μ) D τ1 (C.dom μ)"
        proof -
          have "τ1 (C.cod μ) D (F μ D τ0 (srcC μ))
                  = D.inv (Φ (CU0 (trgC μ), CU.P (C.cod μ))) D
                    F (CU1 (C.cod μ)) D
                    Φ (C.cod μ, CU0 (srcC μ)) D
                    (F μ D τ0 (srcC μ))"
            using μ D.comp_assoc by simp
          also have "... = D.inv (Φ (CU0 (trgC μ), CU.P (C.cod μ))) D
                           (F (CU1 (C.cod μ)) D
                           F (μ C CU0 (srcC μ))) D
                           Φ (C.dom μ, CU0 (srcC μ))"
            using μ Φ.naturality [of "(μ, CU0 (srcC μ))"] D.comp_assoc
                  C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC FF_def
            by simp
          also have "... = (D.inv (Φ (CU0 (trgC μ), CU.P (C.cod μ))) D
                           F (CU0 (trgC μ) C CU.EoP.map μ)) D
                           F (CU1 (C.dom μ)) D
                           Φ (C.dom μ, CU0 (srcC μ))"
          proof -
            have "F (CU1 (C.cod μ)) D F (μ C CU0 (srcC μ))
                    = F (CU1 (C.cod μ) C (μ C CU0 (srcC μ)))"
              using μ by simp
            also have "... = F ((CU0 (trgC μ) C CU.EoP.map μ) C CU1 (C.dom μ))"
              using μ CU.counit.naturality [of μ] by simp
            also have "... = F (CU0 (trgC μ) C CU.EoP.map μ) D
                             F (CU1 (C.dom μ))"
              using μ by simp
            finally have "F (CU1 (C.cod μ)) D F (μ C CU0 (srcC μ))
                            = F (CU0 (trgC μ) C CU.EoP.map μ) D F (CU1 (C.dom μ))"
              by blast
            thus ?thesis
              using D.comp_assoc by simp
          qed
          also have "... = (F (CU0 (trgC μ)) D F (CU.EoP.map μ)) D
                           D.inv (Φ (CU0 (trgC μ), CU.P (C.dom μ))) D
                           F (CU1 (C.dom μ)) D
                           Φ (C.dom μ, CU0 (srcC μ))"
          proof -
            have "D.inv (Φ (CU0 (trgC μ), CU.P (C.cod μ))) D
                  F (CU0 (trgC μ) C CU.EoP.map μ)
                    = (F (CU0 (trgC μ)) D F (CU.EoP.map μ)) D
                      D.inv (Φ (CU0 (trgC μ), CU.P (C.dom μ)))"
            proof -
              have "C.VV.arr (CU0 (trgC μ), CU.P μ)"
              proof (unfold C.VV.arr_charSbC, intro conjI)
                show "C.arr (fst (CU.d (trgC μ), CU.P μ))"
                  using μ by simp
                show "C.arr (snd (CU.d (trgC μ), CU.P μ))"
                  using μ by simp
                show "srcC (fst (CU.d (trgC μ), CU.P μ)) = trgC (snd (CU.d (trgC μ), CU.P μ))"
                  using μ CU.emb.map0_def CU.emb.map_def apply simp
                  using CU.P0_props(1) CU.P_simpsB(3) CU.src_def by fastforce
              qed
              moreover have "CU.EoP.map μ = CU.P μ"
                using μ by (simp add: CU.emb.map_def)
              moreover have "CU.emb.map0 (CU.P0 (trgC μ)) = CU.P0 (trgC μ)"
                using μ CU.P0_props(1) CU.emb.map0_simp by blast
              ultimately show ?thesis
                using μ C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC C.VV.ide_charSbC FF_def
                      Φ.inv_naturality [of "(CU.counit0 (trgC μ), CU.P μ)"]
                by simp
            qed
            thus ?thesis
              using D.comp_assoc by simp
          qed
          also have "... = (τ0 (trgC μ) D EQDoEQUVoEQC'.left_map μ) D τ1 (C.dom μ)"
            using μ FUV.preserves_arr D.comp_assoc DV.emb.map_def CU.emb.map_def by simp
          finally show "τ1 (C.cod μ) D (F μ D τ0 (srcC μ))
                          = (τ0 (trgC μ) D EQDoEQUVoEQC'.left_map μ) D τ1 (C.dom μ)"
            by blast
        qed
      qed
      show "a. C.obj a 
                  (τ0 a D EQDoEQUVoEQC'.FH.unit a) D 𝗋D-1[τ0 a] D 𝗅D[τ0 a] =
                  τ1 a D (unit a D τ0 a)"
      proof -
        fix a
        assume a: "C.obj a"
        have 1: "CU.obj (CU.prj.map0 a)"
          using a CU.prj.map0_simps(1) by simp
        have 2: "«CU.prj.unit a : CU.prj.map0 a C CU.P a»"
          using a CU.in_hom_charSbC by blast
        have 3: "«CU.prj.unit a : CU.prj.map0 a C CU.prj.map0 a»"
        proof (intro C.in_hhomI)
          show 4: "C.arr (CU.prj.unit a)"
            using 2 by auto
          show "srcC (CU.prj.unit a) = CU.prj.map0 a"
          proof -
            have "srcC (CU.prj.unit a) = srcC (CU.P a)"
              using 2 4 C.src_cod [of "CU.prj.unit a"] C.vconn_implies_hpar(1,3) by auto
            also have "... = CU.prj.map0 a"
              using a C.obj_simps(1) CU.prj.map0_def [of a] CU.src_def [of "CU.P a"]
              by simp
            finally show ?thesis by blast
          qed
          show "trgC (CU.prj.unit a) = CU.prj.map0 a"
          proof -
            have "trgC (CU.prj.unit a) = trgC (CU.P a)"
              using 2 4 C.trg_cod [of "CU.prj.unit a"] C.vconn_implies_hpar(2,4) by auto
            also have "... = CU.prj.map0 a"
              using a C.obj_simps(1,3) CU.prj.map0_def [of a] CU.trg_def [of "CU.P a"]
              by simp
            finally show ?thesis by blast
          qed
        qed
        have [simp]: "C.dom (CU.prj.unit a) = CU.prj.map0 a"
          using 2 by auto
        have [simp]: "C.arr (CU.prj.unit a)"
          using 2 by auto
        have [simp]: "C.cod (CU.prj.unit a) = CU.P a"
          using 2 by auto
        have [simp]: "srcC (CU.prj.unit a) = CU.prj.map0 a"
          using 3 by auto
        have [simp]: "trgC (CU.prj.unit a) = CU.prj.map0 a"
          using 3 by auto
        have [simp]: "C.arr (CU.d a)"
          using a by simp
        have [simp]: "C.ide (CU.d a)"
          using a by simp
        have [simp]: "C.dom (CU.d a) = CU.d a"
          using a by auto
        have [simp]: "C.cod (CU.d a) = CU.d a"
          using a by auto
        have [simp]: "srcC (CU.d a) = CU.prj.map0 a"
          using a CU.equivalence_data_simpsB(7) by auto
        have [simp]: "trgC (CU.d a) = a"
          using a CU.equivalence_data_simpsB(7) by auto

        have seq: "D.seq (F (CU.prj.unit a)) (FUV.unit (CU.prj.map0 a))"
          using a
          apply (intro D.seqI)
          using CU.prj.map0_simps(1) DV.arr_charSbC FUV.unit_simps(1) apply blast
          using CU.prj.unit_simps(1) CU.arr_charSbC apply blast
        proof -
          have "D.dom (F (CU.prj.unit a)) = F (C.dom (CU.prj.unit a))"
            using a CU.prj.unit_simps(1) CU.arr_charSbC by simp
          also have "... = F (CU.prj.map0 a)"
            using a CU.prj.unit_simps [of a] CU.dom_charSbC by simp
          also have "... = D.cod (FUV.unit (CU.prj.map0 a))"
            using a 1 CU.prj.unit_simps [of a] FUV.unit_simps(5) [of "CU.prj.map0 a"] DV.cod_charSbC
            by simp
          finally show "D.dom (F (CU.prj.unit a)) = D.cod (FUV.unit (CU.prj.map0 a))"
            by blast
        qed

        have 4: "EQDoEQUVoEQC'.FH.unit a = F (CU.prj.unit a) D unit (CU.prj.map0 a)"
        proof -
          have "EQUVoEQC'.FH.map0 a = map0 (CU.P0 a)"
            using a EQUVoEQC'.FH.map0_def DV.src_def C.obj_def FUV.preserves_arr by auto
          have seq': "DV.arr (F (CU.prj.unit a) D FUV.unit (CU.prj.map0 a))"
          proof -
            have 5: "D.arr (F (CU.prj.unit a) D FUV.unit (CU.prj.map0 a))"
              using seq by simp
            moreover have "srcD (F (CU.prj.unit a) D FUV.unit (CU.prj.map0 a))  V"
            proof -
              have "srcD (F (CU.prj.unit a) D FUV.unit (CU.prj.map0 a))
                      = map0 (CU.prj.map0 a)"
              proof -
                have "srcD (F (CU.prj.unit a) D FUV.unit (CU.prj.map0 a))
                        = srcD (FUV.unit (CU.prj.map0 a))"
                  using 5 D.src_vcomp D.vseq_implies_hpar by simp
                also have "... = map0 (CU.prj.map0 a)"
                  using a FUV.unit_simps(1-2) [of "CU.prj.map0 a"] DV.src_def FUV.map0_def
                  apply simp
                  using a CU.prj.map0_simps(1) FUV.preserves_arr map0_def by force
                finally show ?thesis by blast
              qed
              moreover have "C.obj (CU.prj.map0 a)"
                using a CU.prj.map0_simps(1) CU.obj_char by blast
              ultimately show ?thesis
                using V_def by blast
            qed
            moreover have "trgD (F (CU.prj.unit a) D FUV.unit (CU.prj.map0 a))  V"
            proof -
              have "trgD (F (CU.prj.unit a) D FUV.unit (CU.prj.map0 a))
                      = map0 (CU.prj.map0 a)"
              proof -
                have "trgD (F (CU.prj.unit a) D FUV.unit (CU.prj.map0 a)) =
                      trgD (FUV.unit (CU.prj.map0 a))"
                  using 5 D.trg_vcomp D.vseq_implies_hpar by simp
                also have "... = map0 (CU.prj.map0 a)"
                  using a FUV.unit_simps(1,3) [of "CU.prj.map0 a"] DV.src_def DV.trg_def
                        FUV.map0_def
                  apply simp
                  using a CU.prj.map0_simps(1) FUV.preserves_arr map0_def by force
                finally show ?thesis by blast
              qed
              moreover have "C.obj (CU.prj.map0 a)"
                using a CU.prj.map0_simps(1) CU.obj_char by blast
              ultimately show ?thesis
                using V_def by blast
            qed
            ultimately show ?thesis
              using a DV.arr_charSbC by simp
          qed

          have "EQDoEQUVoEQC'.FH.unit a =
                F (CU.prj.unit a) D FUV.unit (CU.prj.map0 a) D map0 (CU.P0 a)"
          proof -
            have 5: "DV.obj (srcD (F (CU.P0 a)))"
              using a CU.P0_props [of a]
              by (metis (no_types, lifting) EQUVoEQC'.FH.map0_simps(1)
                  EQUVoEQC'.FH.map0 a = map0 (CU.P0 a) map0_def)
            have 6: "DV.emb.unit (map0 (CU.P0 a)) = map0 (CU.P0 a)"
              using a 5 DV.emb.unit_char' EQUVoEQC'.FH.map0_simps(1)
                    EQUVoEQC'.FH.map0 a = map0 (CU.P0 a) map0_def
              by simp
            moreover have "DV.emb.unit (map0 (CU.P0 a)) = map0 (CU.prj.map0 a)"
              using 6 a C.obj_simps CU.equivalence_data_simpsB(7) by simp
            moreover have "DV.arr (F (CU.prj.unit a) D FUV.unit (CU.prj.map0 a))"
              using a C.obj_simps seq' by blast
            moreover have "DV.arr (FUV.unit (CU.P0 a))"
              using a 1 by simp
            moreover have "DV.arr (F (CU.prj.unit a))"
              using a FUV.preserves_arr by auto
            moreover have "DV.obj (FUV.map0 (CU.P0 a))"
              using a 1 FUV.map0_simps(1) by auto
            moreover have "FUV.map0 (CU.P0 a) = map0 (CU.P0 a)"
              using a EQUVoEQC'.FH.map0 a = map0 (CU.P0 a) map0_def FUV.map0_simps
              by auto
            ultimately show ?thesis
              using a seq EQDoEQUVoEQC'.FH.unit_char' [of a] EQUVoEQC'.FH.unit_char' [of a]
                    DV.emb.map_def DV.seq_charSbC DV.comp_char D.comp_assoc
              by simp
          qed
          also have "... = F (CU.prj.unit a) D FUV.unit (CU.prj.map0 a)"
          proof -
            have "D.dom (FUV.unit (CU.prj.map0 a)) = map0 (CU.P0 a)"
              using a 1 FUV.unit_simps(4) [of "CU.prj.map0 a"] DV.dom_charSbC FUV.map0_def
                    DV.src_def map0_def C.obj_def FUV.preserves_arr
              by simp
            moreover have "D.arr (FUV.unit (CU.prj.map0 a))"
              using a D.seq (F (CU.prj.unit a)) (FUV.unit (CU.prj.map0 a)) by blast
            ultimately show ?thesis
              using a D.comp_arr_dom by auto
          qed
          also have "... = F (CU.prj.unit a) D unit (CU.prj.map0 a)"
            using a 1 FUV.unit_char' FU.unit_char' by simp
          finally show ?thesis by blast
        qed

        have "(τ0 a D EQDoEQUVoEQC'.FH.unit a) D 𝗋D-1[τ0 a] D 𝗅D[τ0 a] =
              (τ0 a D F (CU.prj.unit a) D unit (CU.prj.map0 a)) D 𝗋D-1[τ0 a] D 𝗅D[τ0 a]"
          using 4 by simp
        also have "(τ0 a D F (CU.prj.unit a) D unit (CU.prj.map0 a)) D
                   𝗋D-1[τ0 a] D 𝗅D[τ0 a]
                     = (τ0 a D F (CU.prj.unit a) D
                       unit (CU.prj.map0 a)) D
                       (τ0 a D D.inv (unit (srcC (CU.d a)))) D
                       D.inv (Φ (CU.d a, srcC (CU.d a))) D
                       F 𝗋C-1[CU.d a] D
                       F 𝗅C[CU.d a] D
                       Φ (trgC (CU.d a), CU.d a) D
                       (unit (trgC (CU.d a)) D τ0 a)"
        proof -
          have "𝗅D[τ0 a] = F 𝗅C[CU.d a] D Φ (trgC (CU.d a), CU.d a) D
                          (unit (trgC (CU.d a)) D τ0 a)"
            using a preserves_lunit [of "CU0 a"] CU.counit.ide_map0_obj lunit_coherence
            by blast
          moreover
          have "𝗋D-1[τ0 a] = (τ0 a D D.inv (unit (srcC (CU.d a)))) D
                            D.inv (Φ (CU.d a, srcC (CU.d a))) D F 𝗋C-1[CU.d a]"
          proof -
            have "F 𝗋C-1[CU.d a] =
                  Φ (CU.d a, srcC (CU.d a)) D (τ0 a D unit (srcC (CU.d a))) D 𝗋D-1[τ0 a]"
              using preserves_runit(2) [of "CU0 a"] by simp
            moreover have 1: "D.iso (Φ (CU.d a, srcC (CU.d a)))"
              using a C.VV.ide_charSbC C.VV.arr_charSbC C.obj_src [of "CU.d a"]
                    Φ.components_are_iso [of "(CU.d a, srcC (CU.d a))"]
              by auto
            ultimately have "D.inv (Φ (CU.d a, srcC (CU.d a))) D F 𝗋C-1[CU.d a] =
                             (τ0 a D unit (srcC (CU.d a))) D 𝗋D-1[τ0 a]"
              using D.invert_side_of_triangle(1)
                      [of "F 𝗋C-1[CU.d a]" "Φ (CU.d a, srcC (CU.d a))"
                          "(τ0 a D unit (srcC (CU.d a))) D 𝗋D-1[τ0 a]"]
              by fastforce
            thus ?thesis
              using D.invert_side_of_triangle(1)
                      [of "D.inv (Φ (CU.d a, srcC (CU.d a))) D F 𝗋C-1[CU.d a]"
                          "τ0 a D unit (srcC (CU.d a))" "𝗋D-1[τ0 a]"]
              using C.obj_src [of "CU.d a"] unit_char(2) by auto
          qed
          ultimately show ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = ((τ0 a D F (CU.prj.unit a)) D
                         ((τ0 a D unit (CU.prj.map0 a)) D
                         (τ0 a D D.inv (unit (CU.prj.map0 a))))) D
                         D.inv (Φ (CU.d a, CU.prj.map0 a)) D
                         F 𝗋C-1[CU.d a] D F 𝗅C[CU.d a] D
                         Φ (a, CU.d a) D
                         (unit a D τ0 a)"
        proof -
          have "τ0 a D F (CU.prj.unit a) D unit (CU.prj.map0 a) =
                (τ0 a D F (CU.prj.unit a)) D (τ0 a D unit (CU.prj.map0 a))"
            using 1 seq D.whisker_left [of "τ0 a"]
            by (simp add: FU.unit_char' FUV.unit_char')
          thus ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = ((τ0 a D F (CU.prj.unit a)) D
                         D.inv (Φ (CU.d a, CU.prj.map0 a))) D
                         F 𝗋C-1[CU.d a] D F 𝗅C[CU.d a] D
                         Φ (a, CU.d a) D
                         (unit a D τ0 a)"
        proof -
          have "(τ0 a D F (CU.prj.unit a)) D
                ((τ0 a D unit (CU.prj.map0 a)) D
                (τ0 a D D.inv (unit (CU.prj.map0 a))))
                  = (τ0 a D F (CU.prj.unit a)) D
                    (τ0 a D unit (CU.prj.map0 a) D
                    D.inv (unit (CU.prj.map0 a)))"
          proof -
            have "D.seq (unit (CU.prj.map0 a)) (D.inv (unit (CU.prj.map0 a)))"
              using a 1 unit_char(1-2) CU.obj_char
              by (intro D.seqI) auto
            thus ?thesis
              using a D.whisker_left [of "τ0 a"] by simp
          qed
          also have "... = (τ0 a D F (CU.prj.unit a)) D (τ0 a D F (CU.prj.map0 a))"
            using a 1 unit_char(1-2) CU.obj_char unit_simps [of "CU.prj.map0 a"]
            by (simp add: D.comp_arr_inv')
          also have "... = τ0 a D F (CU.prj.unit a) D F (CU.prj.map0 a)"
            using seq D.whisker_left by auto
          also have "... = τ0 a D F (CU.prj.unit a)"
            using D.comp_arr_dom by simp
          finally have "(τ0 a D F (CU.prj.unit a)) D
                        ((τ0 a D unit (CU.prj.map0 a)) D
                        (τ0 a D D.inv (unit (CU.prj.map0 a))))
                          = τ0 a D F (CU.prj.unit a)"
            by blast
          thus ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = D.inv (Φ (CU.d a, CU.P a)) D
                         (F (CU.d a C CU.prj.unit a) D
                         F 𝗋C-1[CU.d a] D F 𝗅C[CU.d a]) D
                         Φ (a, CU.d a) D
                         (unit a D τ0 a)"
        proof -
          have "(τ0 a D F (CU.prj.unit a)) D D.inv (Φ (CU.d a, CU.prj.map0 a)) =
                D.inv (Φ (CU.d a, CU.P a)) D F (CU.d a C CU.prj.unit a)"
          proof -
            have "C.VV.arr (CU.d a, CU.prj.unit a)"
              using a C.VV.arr_charSbC by simp
            thus ?thesis
              using a Φ.inv_naturality [of "(CU.d a, CU.prj.unit a)"] CU.prj.unit_simps [of a]
                    C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC FF_def
              by simp
          qed
          thus ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = (D.inv (Φ (CU.d a, CU.P a)) D
                         F (CU.counit1 a C (a C CU.d a)) D
                         Φ (a, CU.d a)) D
                         (unit a D τ0 a)"
        proof -
          have "F (CU.d a C CU.prj.unit a) D F 𝗋C-1[CU.d a] D F 𝗅C[CU.d a] =
                F ((CU.d a C CU.prj.unit a) C 𝗋C-1[CU.d a] C 𝗅C[CU.d a])"
            using a by simp
          also have "... = F (CU.counit1 a C (a C CU.d a))"
          proof -
            have "CU.EoP.unit a = CU.prj.unit a"
              using a 1 CU.emb.map_def CU.EoP.unit_char' CU.emb.unit_char' C.comp_arr_dom
              by simp
            thus ?thesis
              using a CU.counit.respects_unit [of a] CU.IC.unit_char' [of a] by simp
          qed
          finally have "F (CU.d a C CU.prj.unit a) D F 𝗋C-1[CU.d a] D F 𝗅C[CU.d a] =
                        F (CU.counit1 a C (a C CU.d a))"
            by blast
          thus ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = τ1 a D (unit a D τ0 a)"
        proof -
          have "D.inv (Φ (CU.d a, CU.P a)) D F (CU1 a C (a C CU.d a)) D Φ (a, CU.d a)
                  = D.inv (Φ (CU.d a, CU.P a)) D F (CU1 a) D Φ (a, CU.d a)"
            using a C.obj_def' C.comp_arr_dom
            by (metis CU.counit1_simps(1) CU.counit1_simps(4) C.objE)
          also have "... = τ1 a"
            using a by auto
          finally have "D.inv (Φ (CU.d a, CU.P a)) D F (CU1 a C (a C CU.d a)) D Φ (a, CU.d a)
                          = τ1 a"
            by blast
          thus ?thesis
            using D.comp_assoc by simp
        qed
        finally show "(τ0 a D EQDoEQUVoEQC'.FH.unit a) D 𝗋D-1[τ0 a] D 𝗅D[τ0 a]
                        = τ1 a D (unit a D τ0 a)"
          by blast
      qed
      show "f g. C.ide f; C.ide g; srcC g = trgC f 
                     (τ0 (trgC g) D EQDoEQUVoEQC'.left_cmp (g, f)) D
                     𝖺D[τ0 (trgC g), EQDoEQUVoEQC'.left_map g,
                        EQDoEQUVoEQC'.left_map f] D
                     (τ1 g D EQDoEQUVoEQC'.left_map f) D
                     D.inv 𝖺D[F g, τ0 (srcC g), EQDoEQUVoEQC'.left_map f] D
                     (F g D τ1 f) D 𝖺D[F g, F f, τ0 (srcC f)]
                       = τ1 (g C f) D (Φ (g, f) D τ0 (srcC f))"
      proof -
        fix f g
        assume f: "C.ide f" and g: "C.ide g" and fg: "srcC g = trgC f"
        have 1: "C.ide f  C.ide g  C.ide (CU.P f)  C.ide (CU.P g) 
                 C.ide (CU.d (trgC f))  C.ide (CU.d (trgC g)) 
                 srcC (CU.d (trgC f)) = trgC (CU.P f) 
                 srcC (CU.d (trgC g)) = trgC (CU.P g) 
                 trgC (CU.d (trgC f)) = srcC g"
          using f g fg CU.emb.map0_def CU.emb.map_def CU.obj_char CU.P0_props(1) C.obj_simps(2)
                CU.prj.preserves_ide CU.ide_charSbC
          by auto
        have "τ1 (g C f) D (Φ (g, f) D τ0 (srcC f))
                = (D.inv (Φ (CU.d (trgC g), CU.P (g C f))) D
                  F ((CU.d (trgC g) C CU.EoP.cmp (g, f)) C
                  𝖺C[CU.d (trgC g), CU.P g, CU.P f] C
                  (CU1 g C CU.P f) C
                  C.inv 𝖺C[g, CU.d (srcC g), CU.P f] C
                  (g C CU1 f) C
                  𝖺C[g, f, CU.d (srcC f)]) D
                  Φ (g C f, CU.d (srcC f))) D
                  (Φ (g, f) D τ0 (srcC f))"
          using f g fg CU.emb.map_def CU.counit.respects_hcomp [of f g] D.comp_arr_dom
          by simp
        also have "... = ((D.inv (Φ (CU.d (trgC g), CU.P (g C f))) D
                         Φ (CU.d (trgC g), CU.P (g C f))) D
                         (τ0 (trgC g) D F (CU.EoP.cmp (g, f)))) D
                         (τ0 (trgC g) D Φ (CU.P g, CU.P f)) D
                         𝖺D[τ0 (trgC g), F (CU.P g), F (CU.P f)] D
                         (τ1 g D F (CU.P f)) D
                         D.inv 𝖺D[F g, τ0 (trgC f), F (CU.P f)] D
                         (F g D τ1 f) D
                         𝖺D[F g, F f, τ0 (srcC f)] D
                         (D.inv (Φ (g, f)) D τ0 (srcC f)) D
                         (D.inv (Φ (g C f, CU.d (srcC f))) D
                         Φ (g C f, CU.d (srcC f))) D
                         (Φ (g, f) D τ0 (srcC f))"
        proof -
          have "F ((CU.d (trgC g) C CU.EoP.cmp (g, f)) C
                   𝖺C[CU.d (trgC g), CU.P g, CU.P f] C
                   (CU1 g C CU.P f) C
                   C.inv 𝖺C[g, CU.d (srcC g), CU.P f] C
                   (g C CU1 f) C
                   𝖺C[g, f, CU.d (srcC f)])
                  = F (CU.d (trgC g) C CU.EoP.cmp (g, f)) D
                    F 𝖺C[CU.d (trgC g), CU.P g, CU.P f] D
                    F (CU1 g C CU.P f) D
                    F (C.inv 𝖺C[g, CU.d (srcC g), CU.P f]) D
                    F (g C CU1 f) D
                    F 𝖺C[g, f, CU.d (srcC f)]"
          proof -
            have "C.arr ((CU.d (trgC g) C CU.EoP.cmp (g, f)) C
                         𝖺C[CU.d (trgC g), CU.P g, CU.P f] C
                         (CU1 g C CU.P f) C
                         C.inv 𝖺C[g, CU.d (srcC g), CU.P f] C
                         (g C CU1 f) C
                         𝖺C[g, f, CU.d (srcC f)])"
             using f g fg 1 CU.emb.map_def C.VV.arr_charSbC C.VV.dom_charSbC CU.EoP.FF_def
             by (intro C.seqI C.hseqI') auto
           thus ?thesis
             using f g fg by fastforce
          qed
          also have "... = Φ (CU.d (trgC g), CU.P (g C f)) D
                           (τ0 (trgC g) D F (CU.EoP.cmp (g, f))) D
                           ((D.inv (Φ (CU.d (trgC g), CU.P g C CU.P f)) D
                           Φ (CU.d (trgC g), CU.P g C CU.P f)) D
                           (τ0 (trgC g) D Φ (CU.P g, CU.P f))) D
                           𝖺D[τ0 (trgC g), F (CU.P g), F (CU.P f)] D
                           (D.inv (Φ (CU.d (trgC g), CU.P g)) D F (CU.P f)) D
                           ((D.inv (Φ (CU.d (trgC g) C CU.P g, CU.P f)) D
                           Φ (CU.d (trgC g) C CU.P g, CU.P f)) D
                           (F (CU1 g) D F (CU.P f))) D
                           ((D.inv (Φ (g C CU.d (trgC f), CU.P f)) D
                           Φ (g C CU.d (trgC f), CU.P f)) D
                           (Φ (g, CU.d (trgC f)) D F (CU.P f))) D
                           D.inv 𝖺D[F g, τ0 (trgC f), F (CU.P f)] D
                           (F g D D.inv (Φ (CU.d (trgC f), CU.P f))) D
                           ((D.inv (Φ (g, CU.d (trgC f) C CU.P f)) D
                           Φ (g, CU.d (trgC f) C CU.P f)) D
                           (F g D F (CU1 f))) D
                           ((D.inv (Φ (g, f C CU.d (srcC f))) D
                           Φ (g, f C CU.d (srcC f))) D
                           (F g D Φ (f, CU.d (srcC f)))) D
                           𝖺D[F g, F f, τ0 (srcC f)] D
                           (D.inv (Φ (g, f)) D τ0 (srcC f)) D
                           D.inv (Φ (g C f, CU.d (srcC f)))"
            using 1 f g fg preserves_hcomp preserves_assoc CU.emb.map_def
                  C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def D.comp_assoc
            by simp
          also have "... = Φ (CU.d (trgC g), CU.P (g C f)) D
                           (τ0 (trgC g) D F (CU.EoP.cmp (g, f))) D
                           (τ0 (trgC g) D Φ (CU.P g, CU.P f)) D
                           𝖺D[τ0 (trgC g), F (CU.P g), F (CU.P f)] D
                           ((D.inv (Φ (CU.d (trgC g), CU.P g)) D F (CU.P f)) D
                           (F (CU1 g) D F (CU.P f)) D
                           (Φ (g, CU.d (trgC f)) D F (CU.P f))) D
                           D.inv 𝖺D[F g, τ0 (trgC f), F (CU.P f)] D
                           ((F g D D.inv (Φ (CU.d (trgC f), CU.P f))) D
                           (F g D F (CU1 f)) D
                           (F g D Φ (f, CU.d (srcC f)))) D
                           𝖺D[F g, F f, τ0 (srcC f)] D
                           (D.inv (Φ (g, f)) D τ0 (srcC f)) D
                           D.inv (Φ (g C f, CU.d (srcC f)))"
          proof -
            have "(D.inv (Φ (CU.d (trgC g), CU.P g C CU.P f)) D
                          Φ (CU.d (trgC g), CU.P g C CU.P f)) D
                          (τ0 (trgC g) D Φ (CU.P g, CU.P f))
                    = τ0 (trgC g) D Φ (CU.P g, CU.P f)"
              using f g fg 1 C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
                    Φ.components_are_iso C.VV.ide_charSbC FF_def
              by simp
            moreover have "(D.inv (Φ (CU.d (trgC g) C CU.P g, CU.P f)) D
                           Φ (CU.d (trgC g) C CU.P g, CU.P f)) D
                           (F (CU1 g) D F (CU.P f))
                             = F (CU1 g) D F (CU.P f)"
            proof -
              have "(D.inv (Φ (CU.d (trgC g) C CU.P g, CU.P f)) D
                            Φ (CU.d (trgC g) C CU.P g, CU.P f)) D
                            (F (CU1 g) D F (CU.P f))
                      = (F (CU.d (trgC g) C CU.P g) D F (CU.P f)) D
                        (F (CU1 g) D F (CU.P f))"
                using f g fg 1 C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
                    Φ.components_are_iso C.VV.ide_charSbC FF_def
                by simp
              also have "... = F (CU.counit1 g) D F (CU.P f)"
                using g D.comp_cod_arr f g fg CU.P0_props(1) CU.emb.map_def by simp
              finally show ?thesis by blast
            qed
            moreover have "(D.inv (Φ (g C CU.d (trgC f), CU.P f)) D
                           (Φ (g C CU.d (trgC f), CU.P f))) D
                           (Φ (g, CU.d (trgC f)) D F (CU.P f))
                             = (Φ (g, CU.d (trgC f)) D F (CU.P f))"
              using f g fg 1 C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
                    Φ.components_are_iso C.VV.ide_charSbC FF_def
              by simp
            moreover have "(D.inv (Φ (g, CU.d (trgC f) C CU.P f)) D
                           Φ (g, CU.d (trgC f) C CU.P f)) D
                           (F g D F (CU.counit1 f))
                             = F g D F (CU.counit1 f)"
              using f g fg 1 C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
                    Φ.components_are_iso C.VV.ide_charSbC FF_def
              by simp
            moreover have "(D.inv (Φ (g, f C CU.d (srcC f))) D
                           Φ (g, f C CU.d (srcC f))) D
                           (F g D Φ (f, CU.d (srcC f)))
                             = F g D Φ (f, CU.d (srcC f))"
              using f g fg C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
                    Φ.components_are_iso C.VV.ide_charSbC FF_def
              by simp
            ultimately show ?thesis
              using D.comp_assoc by simp
          qed
          also have "... = Φ (CU.d (trgC g), CU.P (g C f)) D
                           (τ0 (trgC g) D F (CU.EoP.cmp (g, f))) D
                           (τ0 (trgC g) D Φ (CU.P g, CU.P f)) D
                           𝖺D[τ0 (trgC g), F (CU.P g), F (CU.P f)] D
                           (τ1 g D F (CU.P f)) D
                           D.inv 𝖺D[F g, τ0 (trgC f), F (CU.P f)] D
                           (F g D τ1 f) D
                           𝖺D[F g, F f, τ0 (srcC f)] D
                           (D.inv (Φ (g, f)) D τ0 (srcC f)) D
                           D.inv (Φ (g C f, CU.d (srcC f)))"
          proof -
            have "(F g D D.inv (Φ (CU.d (trgC f), CU.P f))) D
                  (F g D F (CU.counit1 f)) D
                  (F g D Φ (f, CU.d (srcC f)))
                    = F g D τ1 f"
              using f g fg 1 D.whisker_left C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC
              by simp
            moreover have "(D.inv (Φ (CU.d (trgC g), CU.P g)) D F (CU.P f)) D
                           (F (CU.counit1 g) D F (CU.P f)) D
                           (Φ (g, CU.d (trgC f)) D F (CU.P f))
                             = τ1 g D F (CU.P f)"
              using f g fg 1 D.whisker_right C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC
                    C.VV.ide_charSbC Φ.components_are_iso CU.emb.map_def
              by simp
            ultimately show ?thesis
              using D.comp_assoc by simp
          qed
          finally have "F ((CU.d (trgC g) C CU.EoP.cmp (g, f)) C
                        𝖺C[CU.d (trgC g), CU.P g, CU.P f] C
                        (CU1 g C CU.P f) C
                        C.inv 𝖺C[g, CU.d (srcC g), CU.P f] C
                        (g C CU1 f) C
                        𝖺C[g, f, CU.d (srcC f)])
                          = Φ (CU.d (trgC g), CU.P (g C f)) D
                            (τ0 (trgC g) D F (CU.EoP.cmp (g, f))) D
                            (τ0 (trgC g) D Φ (CU.P g, CU.P f)) D
                            𝖺D[τ0 (trgC g), F (CU.P g), F (CU.P f)] D
                            (τ1 g D F (CU.P f)) D
                            D.inv 𝖺D[F g, τ0 (trgC f), F (CU.P f)] D
                            (F g D τ1 f) D
                            𝖺D[F g, F f, τ0 (srcC f)] D
                            (D.inv (Φ (g, f)) D τ0 (srcC f)) D
                            D.inv (Φ (g C f, CU.d (srcC f)))"
            by blast
          thus ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = ((τ0 (trgC g) D F (CU.EoP.cmp (g, f))) D
                         (τ0 (trgC g) D Φ (CU.P g, CU.P f))) D
                         𝖺D[τ0 (trgC g), F (CU.P g), F (CU.P f)] D
                         (τ1 g D F (CU.P f)) D
                         D.inv 𝖺D[F g, τ0 (trgC f), F (CU.P f)] D
                         (F g D τ1 f) D
                         𝖺D[F g, F f, τ0 (srcC f)]"
        proof -
          have "(D.inv (Φ (CU.d (trgC g), CU.P (g C f))) D
                Φ (CU.d (trgC g), CU.P (g C f))) D
                (τ0 (trgC g) D F (CU.EoP.cmp (g, f)))
                   = τ0 (trgC g) D F (CU.EoP.cmp (g, f))"
          proof -
            have "(D.inv (Φ (CU.d (trgC g), CU.P (g C f))) D
                  Φ (CU.d (trgC g), CU.P (g C f))) D
                  (τ0 (trgC g) D F (CU.EoP.cmp (g, f)))
                    = (τ0 (trgC g) D F (CU.P (g C f))) D
                      (τ0 (trgC g) D F (CU.EoP.cmp (g, f)))"
            proof -
              have "D.iso (Φ (CU.d (trgC g), CU.P (g C f)))"
                using f g fg Φ.components_are_iso C.VV.ide_charSbC C.VV.arr_charSbC
                by (metis (no_types, lifting) CU.prj.preserves_ide CU.P_simpsB(3) CU.ide_charSbC
                    CU.counit.ide_map0_obj CU.equivalence_data_simpsB(7) C.hcomp_simps(2)
                    C.ideD(1) C.ide_hcomp C.obj_trg cmp_components_are_iso)
              moreover have "D.dom (Φ (CU.d (trgC g), CU.P (g C f))) =
                             τ0 (trgC g) D F (CU.P (g C f))"
                using f g fg CU.ide_charSbC CU.equivalence_data_simpsB(7) CU.prj.preserves_ide
                      cmp_simps(4)
                by simp
              ultimately show ?thesis
                using D.comp_inv_arr' by auto
            qed
            also have "... = τ0 (trgC g) D F (CU.P (g C f)) D F (CU.EoP.cmp (g, f))"
              using f g fg C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def D.whisker_left [of "τ0 (trgC g)"]
              by simp
            also have "... = τ0 (trgC g) D F (CU.EoP.cmp (g, f))"
              using f g fg C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def D.comp_cod_arr
              by simp
            finally show ?thesis by blast
          qed
          moreover have "𝖺D[F g, F f, τ0 (srcC f)] D
                         (D.inv (Φ (g, f)) D τ0 (srcC f)) D
                         (D.inv (Φ (g C f, CU.d (srcC f))) D
                         Φ (g C f, CU.d (srcC f))) D
                         (Φ (g, f) D τ0 (srcC f))
                           = 𝖺D[F g, F f, τ0 (srcC f)]"
          proof -
            have "D.inv (Φ (g C f, CU.d (srcC f))) D Φ (g C f, CU.d (srcC f)) =
                    F (g C f) D F (CU.d (srcC f))"
              using f g fg C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
                    Φ.components_are_iso C.VV.ide_charSbC FF_def
              by simp
            moreover have "(F (g C f) D F (CU.d (srcC f))) D (Φ (g, f) D τ0 (srcC f)) =
                           Φ (g, f) D τ0 (srcC f)"
              using f g fg C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
                    Φ.components_are_iso C.VV.ide_charSbC FF_def
              by simp
            moreover have "(D.inv (Φ (g, f)) D τ0 (srcC f)) D (Φ (g, f) D τ0 (srcC f)) =
                           (F g D F f) D τ0 (srcC f)"
              using f g fg C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
                    Φ.components_are_iso C.VV.ide_charSbC FF_def
                    D.whisker_right [of "τ0 (srcC f)" "D.inv (Φ (g, f))" "Φ (g, f)"]
              by simp
            moreover have "𝖺D[F g, F f, τ0 (srcC f)] D ((F g D F f) D τ0 (srcC f)) =
                           𝖺D[F g, F f, τ0 (srcC f)]"
              using f g fg C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr
              by simp
            ultimately show ?thesis by argo
          qed
          ultimately show ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = (τ0 (trgC g) D F (CU.EoP.cmp (g, f)) D
                         Φ (CU.P g, CU.P f)) D
                         𝖺D[τ0 (trgC g), F (CU.P g), F (CU.P f)] D
                         (τ1 g D F (CU.P f)) D
                         D.inv 𝖺D[F g, τ0 (trgC f), F (CU.P f)] D
                         (F g D τ1 f) D
                         𝖺D[F g, F f, τ0 (srcC f)]"
        proof -
          have "(τ0 (trgC g) D F (CU.EoP.cmp (g, f))) D
                (τ0 (trgC g) D Φ (CU.P g, CU.P f))
                  = τ0 (trgC g) D F (CU.EoP.cmp (g, f)) D Φ (CU.P g, CU.P f)"
          proof -
            have "D.arr (F (CU.EoP.cmp (g, f)) D Φ (CU.P g, CU.P f))"
              using f g fg C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC CU.EoP.FF_def
                    CU.emb.map_def
              by (intro D.seqI) auto
            thus ?thesis
              using f g fg C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC
                    D.whisker_left [of "τ0 (trgC g)"]
              by simp
          qed
          thus ?thesis
            using D.comp_assoc by simp
        qed
        also have "... = (τ0 (trgC g) D EQDoEQUVoEQC'.left_cmp (g, f)) D
                         𝖺D[τ0 (trgC g), EQDoEQUVoEQC'.left_map g,
                            EQDoEQUVoEQC'.left_map f] D
                         (τ1 g D EQDoEQUVoEQC'.left_map f) D
                         D.inv 𝖺D[F g, τ0 (srcC g), EQDoEQUVoEQC'.left_map f] D
                         (F g D τ1 f) D
                         𝖺D[F g, F f, τ0 (srcC f)]"
        proof -
          have "EQDoEQUVoEQC'.left_cmp (g, f) =
                F (CU.EoP.cmp (g, f)) D Φ (CU.P g, CU.P f)"
          proof -
            have "EQDoEQUVoEQC'.left_cmp (g, f) =
                    F (CU.P (g C f)) D ((F (CU.P (g C f)) D F (CUP (g, f))) D
                    Φ (CU.P g, CU.P f)) D DVE (F (CU.P g), F (CU.P f))"
            proof -
              have 3: "D.arr (F (CU.P (g C f)) D F (CUP (g, f)) D Φ (CU.P g, CU.P f))"
                using f g fg C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC FF_def CU.prj.FF_def
                      CUP_in_hom [of g f] CU.hcomp_def CU.in_hom_charSbC
                by (intro D.seqI) auto
              have 4: "DV.arr (Φ (CU.P g, CU.P f))"
                using f g fg 1 cmp_in_hom(1) [of "CU.P g" "CU.P f"] CU.prj.map0_simps(1)
                      CU.obj_char DV.arr_charSbC V_def
                by auto
              moreover have 5: "DV.arr (F (CUP (g, f)))"
                using f g fg 1 CUP_simps [of g f] FUV.preserves_arr by presburger
              moreover have 6: "DV.arr (F (CU.P (g C f)))"
                using f g fg CU.P_simps(1) [of "g C f"] FUV.preserves_arr by simp
              moreover have 7: "DV.arr (F (CUP (g, f)) D Φ (CU.P g, CU.P f))"
                using 3 4 5 DV.seq_charSbC [of "F (CUP (g, f))" "Φ (CU.P g, CU.P f)"]
                      DV.comp_char [of "F (CUP (g, f))" "Φ (CU.P g, CU.P f)"]
                by auto
              moreover have "D.seq (F (CUP (g, f))) (Φ (CU.P g, CU.P f))"
                using 3 by blast
              moreover have "D.seq (F (CU.P (g C f)))
                                   (F (CUP (g, f)) D Φ (CU.P g, CU.P f))"
                using 3 by blast
              moreover have "DV.arr (F (CU.P (g C f)) D F (CUP (g, f)) D
                                     Φ (CU.P g, CU.P f))"
                using f g fg 3 6 7 D.vseq_implies_hpar V_def
                by (metis (no_types, lifting) DV.comp_char DV.seq_charSbC)
              ultimately show ?thesis
                using f g fg EQDoEQUVoEQC'.FH.cmp_def EQUVoEQC'.FH.cmp_def
                      C.VV.arr_charSbC C.VV.dom_charSbC CU.VV.arr_charSbC DV.comp_char
                      DV.emb.map_def D.comp_assoc
                by simp
            qed
            also have "... = F (CU.P (g C f)) D F (CUP (g, f)) D
                             Φ (CU.P g, CU.P f) D DVE (F (CU.P g), F (CU.P f))"
            proof -
              have "F (CU.P (g C f)) D F (CUP (g, f)) = F (CUP (g, f))"
              proof -
                have 1: "C.arr (CUP (g, f))"
                  using f g fg C.VV.arr_charSbC
                  by (metis (no_types, lifting) CU.prj.preserves_hcomp CU.prj.preserves_ide
                      CU.ide_charSbC CU.seq_charSbC C.ideD(1) C.ideD(3) C.ide_hcomp C.seqE)
                moreover have "C.cod (CUP (g, f)) = CU.P (g C f)"
                  using f g fg C.VV.arr_charSbC CUP_in_hom(2) [of g f]
                        CU.hcomp_def CU.in_hom_charSbC by auto
                ultimately have "D.cod (F (CUP (g, f))) = F (CU.P (g C f))"
                  using f g fg C.VV.arr_charSbC preserves_cod [of "CUP (g, f)"]
                  by simp
                thus ?thesis
                  using 1 f g fg D.comp_cod_arr [of "F (CUP (g, f))" "F (CU.P (g C f))"]
                  by simp
              qed
              thus ?thesis
                using D.comp_assoc by simp
            qed
            also have "... = F (CU.P (g C f)) D F (CUP (g, f)) D
                             F (CUE (CU.P g, CU.P f)) D Φ (CU.P g, CU.P f)"
            proof -
              have "DVE (F (CU.P g), F (CU.P f)) = F (CU.P g) D F (CU.P f)"
                using f g fg DV.emb.cmp_def DV.VV.arr_charSbC DV.src_def DV.trg_def
                      FUV.preserves_arr
                by auto
              moreover have "CUE (CU.P g, CU.P f) = CU.P g C CU.P f"
                using f g fg CU.VV.arr_charSbC CU.emb.cmp_def by simp
              ultimately
              have "Φ (CU.P g, CU.P f) D DVE (F (CU.P g), F (CU.P f)) =
                      F (CUE (CU.P g, CU.P f)) D Φ (CU.P g, CU.P f)"
                using f g fg C.VV.arr_charSbC C.VV.dom_charSbC C.VV.cod_charSbC FF_def
                      Φ.naturality [of "(CU.P g, CU.P f)"]
                by simp
              thus ?thesis
                using D.comp_assoc by simp
            qed
            also have "... = F (CU.P (g C f) C CUP (g, f) C CUE (CU.P g, CU.P f)) D
                             Φ (CU.P g, CU.P f)"
            proof -
              have "C.arr (CU.P (g C f) C CUP (g, f) C CUE (CU.P g, CU.P f))"
              proof (intro C.seqI)
                show "C.arr (CUE (CU.P g, CU.P f))"
                  using f g fg by auto
                show 1: "C.arr (CUP (g, f))"
                  using f g fg CU.arr_charSbC CUP_simps(1) [of g f] by auto
                show "C.dom (CUP (g, f)) = C.cod (CUE (CU.P g, CU.P f))"
                proof -
                  have "C.dom (CUP (g, f)) = CU.hcomp (CU.P g) (CU.P f)"
                    using f g fg CU.VV.arr_charSbC CUP_simps(4) [of g f]
                          CU.dom_charSbC [of "CUP (g, f)"]
                    by fastforce
                  also have "... = CU.P g C (CU.P f)"
                    using f g fg by auto
                  also have "... = C.cod (CUE (CU.P g, CU.P f))"
                    using f g fg CU.emb.cmp_def [of "(CU.P g, CU.P f)"] by auto
                  finally show ?thesis by blast
                qed
                show "C.arr (CU.P (g C f))"
                  using f g fg by simp
                show "C.dom (CU.P (g C f)) =
                      C.cod (CUP (g, f) C CUE (CU.P g, CU.P f))"
                proof -
                  have "C.dom (CU.P (g C f)) = CU.P (g C f)"
                    using f g fg by simp
                  also have "... = C.cod (CUP (g, f))"
                    using f g fg CUP_simps(5) [of g f] CU.cod_charSbC [of "CUP (g, f)"]
                    by fastforce
                  also have "... = C.cod (CUP (g, f) C CUE (CU.P g, CU.P f))"
                  proof -
                    have 2: "CU.hcomp (CU.P g) (CU.P f) = CU.P g C CU.P f"
                      using f g fg by auto
                    have "CU.arr (CU.P g C CU.P f)"
                    proof -
                      have "CU.arr (CU.hcomp (CU.P g) (CU.P f))"
                        using f g fg by simp
                      thus ?thesis
                        using 2 by simp
                    qed
                    moreover have "C.dom (CUP (g, f)) = CU.P g C CU.P f"
                      using f g fg 2 CU.VV.arr_charSbC CUP_simps(4) [of g f]
                            CU.dom_charSbC [of "CUP (g, f)"]
                      by fastforce
                    ultimately have "C.arr (CUP (g, f) C CUE (CU.P g, CU.P f))"
                      using f g fg 1 CU.VV.arr_charSbC CU.VV.cod_charSbC CU.hcomp_char
                            CU.emb.map_def
                      by auto
                    thus ?thesis by simp
                  qed
                  finally show ?thesis by blast
                qed
              qed
              thus ?thesis
                using D.comp_assoc by auto
            qed
            also have "... = F (CU.EoP.cmp (g, f)) D Φ (CU.P g, CU.P f)"
              using f g fg CU.EoP.cmp_def C.VV.arr_charSbC C.VV.dom_charSbC CU.emb.map_def
              by simp
            finally show ?thesis by blast
          qed
          moreover have "EQDoEQUVoEQC'.left_map g = F (CU.P g)"
          proof -
            have "DV.arr (F (CU.P g))"
              using g DV.arr_charSbC CU.P_simps(1) C.ideD(1) FUV.preserves_arr by presburger
            thus ?thesis
              using g DV.emb.map_def by simp
          qed
          moreover have "EQDoEQUVoEQC'.left_map f = F (CU.P f)"
          proof -
            have "DV.arr (F (CU.P f))"
              using f DV.arr_charSbC CU.P_simps(1) C.ideD(1) FUV.preserves_arr by presburger
            thus ?thesis
              using f DV.emb.map_def by simp
          qed
          ultimately show ?thesis
            using fg by argo
        qed
        finally show "(τ0 (trgC g) D EQDoEQUVoEQC'.left_cmp (g, f)) D
                      𝖺D[τ0 (trgC g), EQDoEQUVoEQC'.left_map g,
                         EQDoEQUVoEQC'.left_map f] D
                      (τ1 g D EQDoEQUVoEQC'.left_map f) D
                      D.inv 𝖺D[F g, τ0 (srcC g), EQDoEQUVoEQC'.left_map f] D
                      (F g D τ1 f) D 𝖺D[F g, F f, τ0 (srcC f)]
                        = τ1 (g C f) D (Φ (g, f) D τ0 (srcC f))"
          by argo
      qed
    qed

    interpretation EQ: equivalence_of_bicategories_and_pseudonatural_equivalence_left  (* 17 sec *)
                         VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                         EQDoEQUVoEQC'.left_map EQDoEQUVoEQC'.left_cmp
                         EQDoEQUVoEQC'.right_map EQDoEQUVoEQC'.right_cmp
                         EQDoEQUVoEQC'.unit0 EQDoEQUVoEQC'.unit1
                         EQDoEQUVoEQC'.counit0 EQDoEQUVoEQC'.counit1
                         F Φ τ0 τ1
    proof -
      interpret E: equivalence_of_bicategories
                      VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                         EQDoEQUVoEQC'.left_map EQDoEQUVoEQC'.left_cmp
                         EQDoEQUVoEQC'.right_map EQDoEQUVoEQC'.right_cmp
                         EQDoEQUVoEQC'.unit0 EQDoEQUVoEQC'.unit1
                         EQDoEQUVoEQC'.counit0 EQDoEQUVoEQC'.counit1
        using induces_equivalence_of_bicategories by blast
      interpret τ: pseudonatural_equivalence
                      VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD
                      EQDoEQUVoEQC'.left_map EQDoEQUVoEQC'.left_cmp F Φ τ0 τ1
        ..
      show "equivalence_of_bicategories_and_pseudonatural_equivalence_left
                         VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                         EQDoEQUVoEQC'.left_map EQDoEQUVoEQC'.left_cmp
                         EQDoEQUVoEQC'.right_map EQDoEQUVoEQC'.right_cmp
                         EQDoEQUVoEQC'.unit0 EQDoEQUVoEQC'.unit1
                         EQDoEQUVoEQC'.counit0 EQDoEQUVoEQC'.counit1
                         F Φ τ0 τ1"
        ..
    qed

    definition right_map
    where "right_map  EQDoEQUVoEQC'.right_map"

    definition right_cmp
    where "right_cmp  EQDoEQUVoEQC'.right_cmp"

    definition unit0
    where "unit0  EQ.unit.map0"

    definition unit1
    where "unit1  EQ.unit.map1"

    definition counit0
    where "counit0  EQ.counit.map0"

    definition counit1
    where "counit1  EQ.counit.map1"

    theorem extends_to_equivalence_of_bicategories:
    shows "equivalence_of_bicategories VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
             F Φ right_map right_cmp unit0 unit1 counit0 counit1"
      unfolding right_map_def right_cmp_def unit0_def unit1_def counit0_def counit1_def
      ..

  end

  locale converse_equivalence_pseudofunctor =  (* 10 sec *)
    C: bicategory VC HC 𝖺C 𝗂C srcC trgC +
    D: bicategory VD HD 𝖺D 𝗂D srcD trgD +
    F: equivalence_pseudofunctor VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F ΦF
  for VC :: "'c comp"                    (infixr "C" 55)
  and HC :: "'c comp"                   (infixr "C" 53)
  and 𝖺C :: "'c  'c  'c  'c"       ("𝖺C[_, _, _]")
  and 𝗂C :: "'c  'c"                   ("𝗂C[_]")
  and srcC :: "'c  'c"
  and trgC :: "'c  'c"
  and VD :: "'d comp"                    (infixr "D" 55)
  and HD :: "'d comp"                   (infixr "D" 53)
  and 𝖺D :: "'d  'd  'd  'd"       ("𝖺D[_, _, _]")
  and 𝗂D :: "'d  'd"                   ("𝗂D[_]")
  and srcD :: "'d  'd"
  and trgD :: "'d  'd"
  and F :: "'c  'd"
  and ΦF :: "'c * 'c  'd"
  begin

    interpretation E: equivalence_of_bicategories
                        VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                        F ΦF F.right_map F.right_cmp F.unit0 F.unit1 F.counit0 F.counit1
      using F.extends_to_equivalence_of_bicategories by simp
    interpretation E': converse_equivalence_of_bicategories
                         VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                         F ΦF F.right_map F.right_cmp F.unit0 F.unit1 F.counit0 F.counit1
      ..

    sublocale equivalence_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                F.right_map F.right_cmp
      using E'.equivalence_pseudofunctor_left by simp

    lemma is_equivalence_pseudofunctor:
    shows "equivalence_pseudofunctor VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC
                F.right_map F.right_cmp"
      ..

  end

  (*
   * TODO: Change the following definition to be based on the existence of an equivalence of
   * bicategories, rather than an equivalence pseudofunctor.  This will require a few changes
   * elsewhere.
   *)

  definition equivalent_bicategories
  where "equivalent_bicategories VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC 
         F Φ. equivalence_pseudofunctor
                 VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC F Φ"

  lemma equivalent_bicategories_reflexive:
  assumes "bicategory VC HC 𝖺C 𝗂C srcC trgC"
  shows "equivalent_bicategories VC HC 𝖺C 𝗂C srcC trgC VC HC 𝖺C 𝗂C srcC trgC"
  proof -
    interpret bicategory VC HC 𝖺C 𝗂C srcC trgC
      using assms by simp
    interpret I: identity_pseudofunctor VC HC 𝖺C 𝗂C srcC trgC ..
    show ?thesis
      using I.equivalence_pseudofunctor_axioms equivalent_bicategories_def by blast
  qed

  lemma equivalent_bicategories_symmetric:
  assumes "equivalent_bicategories VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD"
  shows "equivalent_bicategories VD HD 𝖺D 𝗂D srcD trgD VC HC 𝖺C 𝗂C srcC trgC"
  proof -
    obtain F ΦF where F: "equivalence_pseudofunctor
                            VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F ΦF"
      using assms equivalent_bicategories_def by blast
    interpret F: equivalence_pseudofunctor
                   VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F ΦF
      using F by simp
    interpret G: converse_equivalence_pseudofunctor
                   VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F ΦF
      ..
    show ?thesis
      using G.is_equivalence_pseudofunctor equivalent_bicategories_def by blast
  qed

  lemma equivalent_bicategories_transitive:
  assumes "equivalent_bicategories VB HB 𝖺B 𝗂B srcB trgB VC HC 𝖺C 𝗂C srcC trgC"
  and "equivalent_bicategories VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD"
  shows "equivalent_bicategories VB HB 𝖺B 𝗂B srcB trgB VD HD 𝖺D 𝗂D srcD trgD"
  proof -
    obtain F ΦF where F: "equivalence_pseudofunctor
                            VB HB 𝖺B 𝗂B srcB trgB VC HC 𝖺C 𝗂C srcC trgC F ΦF"
      using assms(1) equivalent_bicategories_def by blast
    obtain G ΦG where G: "equivalence_pseudofunctor
                            VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD G ΦG"
      using assms(2) equivalent_bicategories_def by blast
    interpret F: equivalence_pseudofunctor
                   VB HB 𝖺B 𝗂B srcB trgB VC HC 𝖺C 𝗂C srcC trgC F ΦF
      using F by simp
    interpret G: equivalence_pseudofunctor
                   VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD G ΦG
      using G by simp
    interpret GF: composite_equivalence_pseudofunctor VB HB 𝖺B 𝗂B srcB trgB
                    VC HC 𝖺C 𝗂C srcC trgC VD HD 𝖺D 𝗂D srcD trgD F ΦF G ΦG ..
    show "equivalent_bicategories VB HB 𝖺B 𝗂B srcB trgB VD HD 𝖺D 𝗂D srcD trgD"
      using equivalent_bicategories_def GF.equivalence_pseudofunctor_axioms by blast
  qed

end