diff --git a/thys/Bicategory/EquivalenceOfBicategories.thy b/thys/Bicategory/EquivalenceOfBicategories.thy --- a/thys/Bicategory/EquivalenceOfBicategories.thy +++ b/thys/Bicategory/EquivalenceOfBicategories.thy @@ -1,10427 +1,10421 @@ (* Title: EquivalenceOfBicategories Author: Eugene W. Stark , 2020 Maintainer: Eugene W. Stark *) 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 \\: I\<^sub>D \ GF\ and \\: FG \ I\<^sub>C\. \ - locale equivalence_of_bicategories = (* 35 sec *) + locale equivalence_of_bicategories = (* 25 sec *) C: bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + D: bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + F: pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \\<^sub>F + G: pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D G \\<^sub>G + FG: composite_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C G \\<^sub>G F \\<^sub>F + GF: composite_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \\<^sub>F G \\<^sub>G + I\<^sub>C: identity_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + I\<^sub>D: identity_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + \: pseudonatural_equivalence V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D I\<^sub>D.map I\<^sub>D.cmp GF.map GF.cmp \\<^sub>0 \\<^sub>1 + \: pseudonatural_equivalence V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C FG.map FG.cmp I\<^sub>C.map I\<^sub>C.cmp \\<^sub>0 \\<^sub>1 for V\<^sub>C :: "'c comp" (infixr "\\<^sub>C" 55) and H\<^sub>C :: "'c comp" (infixr "\\<^sub>C" 53) and \\<^sub>C :: "'c \ 'c \ 'c \ 'c" ("\\<^sub>C[_, _, _]") and \\<^sub>C :: "'c \ 'c" ("\\<^sub>C[_]") and src\<^sub>C :: "'c \ 'c" and trg\<^sub>C :: "'c \ 'c" and V\<^sub>D :: "'d comp" (infixr "\\<^sub>D" 55) and H\<^sub>D :: "'d comp" (infixr "\\<^sub>D" 53) and \\<^sub>D :: "'d \ 'd \ 'd \ 'd" ("\\<^sub>D[_, _, _]") and \\<^sub>D :: "'d \ 'd" ("\\<^sub>D[_]") and src\<^sub>D :: "'d \ 'd" and trg\<^sub>D :: "'d \ 'd" and F :: "'d \ 'c" and \\<^sub>F :: "'d * 'd \ 'c" and G :: "'c \ 'd" and \\<^sub>G :: "'c * 'c \ 'd" and \\<^sub>0 :: "'d \ 'd" and \\<^sub>1 :: "'d \ 'd" and \\<^sub>0 :: "'c \ 'c" and \\<^sub>1 :: "'c \ 'c" begin notation C.isomorphic (infix "\\<^sub>C" 50) notation D.isomorphic (infix "\\<^sub>D" 50) notation C.iso_in_hom ("\_ : _ \\<^sub>C _\") notation D.iso_in_hom ("\_ : _ \\<^sub>D _\") notation C.some_quasi_inverse ("_\<^sup>~\<^sup>C" [1000] 1000) notation D.some_quasi_inverse ("_\<^sup>~\<^sup>D" [1000] 1000) interpretation \': converse_pseudonatural_equivalence V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D I\<^sub>D.map I\<^sub>D.cmp GF.map GF.cmp \\<^sub>0 \\<^sub>1 .. interpretation \': converse_pseudonatural_equivalence V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C FG.map FG.cmp I\<^sub>C.map I\<^sub>C.cmp \\<^sub>0 \\<^sub>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 V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D GF.map GF.cmp proof - interpret GF: pseudofunctor_pseudonaturally_equivalent_to_equivalence_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D I\<^sub>D.map I\<^sub>D.cmp GF.map GF.cmp \\<^sub>0 \\<^sub>1 .. show "equivalence_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D GF.map GF.cmp" using GF.is_equivalence_pseudofunctor by simp qed interpretation FG: equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C FG.map FG.cmp proof - interpret FG: pseudofunctor_pseudonaturally_equivalent_to_equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C I\<^sub>C.map I\<^sub>C.cmp FG.map FG.cmp \'.map\<^sub>0 \'.map\<^sub>1 .. show "equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C 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 V\<^sub>D V\<^sub>C F proof fix \ \' assume par: "D.par \ \'" and eq: "F \ = F \'" have "src\<^sub>D \ = src\<^sub>D \' \ trg\<^sub>D \ = trg\<^sub>D \'" using par D.src_dom D.trg_cod by (metis D.src_cod D.trg_cod) hence "\\<^sub>0 (trg\<^sub>D \) \\<^sub>D \ = \\<^sub>0 (trg\<^sub>D \) \\<^sub>D \'" using par eq \.iso_map\<^sub>1_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_map\<^sub>0_obj \.components_are_equivalences D.equivalence_cancel_left [of "\\<^sub>0 (trg\<^sub>D \)" \ \'] by simp qed interpretation G: faithful_functor V\<^sub>C V\<^sub>D G proof show "\\ \'. \C.par \ \'; G \ = G \'\ \ \ = \'" proof - fix \ \' assume par: "C.par \ \'" and eq: "G \ = G \'" have "src\<^sub>C \ = src\<^sub>C \' \ trg\<^sub>C \ = trg\<^sub>C \'" using par by (metis C.src_cod C.trg_cod) hence "\ \\<^sub>C \\<^sub>0 (src\<^sub>C \) = \' \\<^sub>C \\<^sub>0 (src\<^sub>C \)" using par eq \.iso_map\<^sub>1_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_map\<^sub>0_obj \.components_are_equivalences C.equivalence_cancel_right [of "\\<^sub>0 (src\<^sub>C \)" \ \'] 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' \ \\<^sub>C[f'] \\<^sub>C (f' \\<^sub>C \'.counit (src\<^sub>C f)) \\<^sub>C \\<^sub>C[f', \\<^sub>0 (src\<^sub>C f), \'.map\<^sub>0 (src\<^sub>C f)] \\<^sub>C (C.inv (\\<^sub>1 f') \\<^sub>C \'.map\<^sub>0 (src\<^sub>C f)) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C f), F (G f'), \'.map\<^sub>0 (src\<^sub>C f)] \\<^sub>C (\\<^sub>0 (trg\<^sub>C f) \\<^sub>C F \ \\<^sub>C \'.map\<^sub>0 (src\<^sub>C f)) \\<^sub>C (\\<^sub>0 (trg\<^sub>C f) \\<^sub>C C.inv (\'.map\<^sub>1 f)) \\<^sub>C \\<^sub>C[\\<^sub>0 (trg\<^sub>C f), \'.map\<^sub>0 (trg\<^sub>C f), f] \\<^sub>C (C.inv (\'.counit (trg\<^sub>C f)) \\<^sub>C f) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f]" lemma G_reflects_isomorphic: assumes "C.ide f" and "C.ide f'" and "src\<^sub>C f = src\<^sub>C f'" and "trg\<^sub>C f = trg\<^sub>C f'" and "\\ : G f \\<^sub>D G f'\" shows "\\ \ f f' : f \\<^sub>C f'\" proof - let ?a = "src\<^sub>C f" and ?a' = "trg\<^sub>C f" have f: "\f : ?a \\<^sub>C ?a'\ \ C.ide f" using assms by simp have f': "\f' : ?a \\<^sub>C ?a'\ \ C.ide f'" using assms by simp have \: "\\ : G.map\<^sub>0 ?a \\<^sub>D G.map\<^sub>0 ?a'\" using assms D.vconn_implies_hpar(1-2) by (elim D.iso_in_homE) auto hence F\: "\F \ : F.map\<^sub>0 (G.map\<^sub>0 ?a) \\<^sub>C F.map\<^sub>0 (G.map\<^sub>0 ?a')\" by auto show "\\ \ f f' : f \\<^sub>C f'\" proof (intro C.comp_iso_in_hom) show "\\\<^sub>C\<^sup>-\<^sup>1[f] : f \\<^sub>C ?a' \\<^sub>C f\" using f by auto show "\C.inv (\'.counit ?a') \\<^sub>C f : ?a' \\<^sub>C f \\<^sub>C (\\<^sub>0 ?a' \\<^sub>C \'.map\<^sub>0 ?a') \\<^sub>C f\" using f \'.counit_in_hom [of ?a'] \'.iso_counit [of ?a'] by (intro C.hcomp_iso_in_hom) auto show "\\\<^sub>C[\\<^sub>0 ?a', \'.map\<^sub>0 ?a', f] : (\\<^sub>0 ?a' \\<^sub>C \'.map\<^sub>0 ?a') \\<^sub>C f \\<^sub>C \\<^sub>0 ?a' \\<^sub>C \'.map\<^sub>0 ?a' \\<^sub>C f\" using f \'.counit_in_hom [of ?a] \'.ide_map\<^sub>0_obj by (intro C.iso_in_homI) auto show "\\\<^sub>0 ?a' \\<^sub>C C.inv (\'.map\<^sub>1 f) : \\<^sub>0 ?a' \\<^sub>C \'.map\<^sub>0 ?a' \\<^sub>C f \\<^sub>C \\<^sub>0 ?a' \\<^sub>C F (G f) \\<^sub>C \'.map\<^sub>0 ?a\" using f \'.map\<^sub>1_in_hom [of f] \'.iso_map\<^sub>1_ide [of f] C.ide_iso_in_hom by (intro C.hcomp_iso_in_hom) auto show "\\\<^sub>0 ?a' \\<^sub>C F \ \\<^sub>C \'.map\<^sub>0 ?a : \\<^sub>0 ?a' \\<^sub>C F (G f) \\<^sub>C \'.map\<^sub>0 ?a \\<^sub>C \\<^sub>0 ?a' \\<^sub>C F (G f') \\<^sub>C \'.map\<^sub>0 ?a\" using assms f f' F\ F.preserves_iso C.ide_iso_in_hom by (intro C.hcomp_iso_in_hom) auto show "\\\<^sub>C\<^sup>-\<^sup>1[\\<^sub>0 ?a', F (G f'), \'.map\<^sub>0 ?a] : \\<^sub>0 ?a' \\<^sub>C F (G f') \\<^sub>C \'.map\<^sub>0 ?a \\<^sub>C (\\<^sub>0 ?a' \\<^sub>C F (G f')) \\<^sub>C \'.map\<^sub>0 ?a\" using assms f' \.map\<^sub>0_in_hom(1) [of ?a'] \.ide_map\<^sub>0_obj [of ?a'] \'.map\<^sub>0_in_hom(1) [of ?a] \'.ide_map\<^sub>0_obj [of ?a] C.assoc'_in_hom C.iso_assoc' by auto show "\C.inv (\\<^sub>1 f') \\<^sub>C \'.map\<^sub>0 ?a : (\\<^sub>0 ?a' \\<^sub>C F (G f')) \\<^sub>C \'.map\<^sub>0 ?a \\<^sub>C (f' \\<^sub>C \\<^sub>0 ?a) \\<^sub>C \'.map\<^sub>0 ?a\" using assms f' \.map\<^sub>1_in_hom \.iso_map\<^sub>1_ide \'.map\<^sub>0_in_hom(1) [of ?a] C.ide_iso_in_hom by (intro C.hcomp_iso_in_hom) auto show "\\\<^sub>C[f', \\<^sub>0 ?a, \'.map\<^sub>0 ?a] : (f' \\<^sub>C \\<^sub>0 ?a) \\<^sub>C \'.map\<^sub>0 ?a \\<^sub>C f' \\<^sub>C \\<^sub>0 ?a \\<^sub>C \'.map\<^sub>0 ?a\" using assms f' \.map\<^sub>0_in_hom(1) [of ?a] \'.map\<^sub>0_in_hom(1) [of ?a] \.ide_map\<^sub>0_obj \'.ide_map\<^sub>0_obj C.assoc_in_hom by auto show "\f' \\<^sub>C \'.counit ?a : f' \\<^sub>C \\<^sub>0 ?a \\<^sub>C \'.map\<^sub>0 ?a \\<^sub>C f' \\<^sub>C ?a\" using f f' \'.counit_in_hom by (intro C.hcomp_iso_in_hom) auto show "\\\<^sub>C[f'] : f' \\<^sub>C ?a \\<^sub>C f'\" using assms f' by auto qed qed abbreviation \ where "\ f b b' \ \\<^sub>D[G (F (\'.map\<^sub>0 b' \\<^sub>D G f \\<^sub>D \\<^sub>0 b))] \\<^sub>D (G (F (\'.map\<^sub>0 b' \\<^sub>D G f \\<^sub>D \\<^sub>0 b)) \\<^sub>D \'.\ b) \\<^sub>D \\<^sub>D[G (F (\'.map\<^sub>0 b' \\<^sub>D G f \\<^sub>D \\<^sub>0 b)), \\<^sub>0 b, \'.map\<^sub>0 b] \\<^sub>D (D.inv (\\<^sub>1 (\'.map\<^sub>0 b' \\<^sub>D G f \\<^sub>D \\<^sub>0 b)) \\<^sub>D \'.map\<^sub>0 b) \\<^sub>D (\\<^sub>D[\\<^sub>0 b', \'.map\<^sub>0 b', G f \\<^sub>D \\<^sub>0 b] \\<^sub>D \'.map\<^sub>0 b) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b' \\<^sub>D \'.map\<^sub>0 b', G f \\<^sub>D \\<^sub>0 b, \'.map\<^sub>0 b] \\<^sub>D ((\\<^sub>0 b' \\<^sub>D \'.map\<^sub>0 b') \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 b, \'.map\<^sub>0 b]) \\<^sub>D (D.inv (\'.counit b') \\<^sub>D G f \\<^sub>D D.inv (\'.counit b)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f \\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f]" lemma F_is_locally_essentially_surjective: assumes "D.obj b" and "D.obj b'" and "\f : F.map\<^sub>0 b \\<^sub>C F.map\<^sub>0 b'\" and "C.ide f" shows "\\ (\ f b b') f (F (\'.map\<^sub>0 b' \\<^sub>D G f \\<^sub>D \\<^sub>0 b)) : f \\<^sub>C F (\'.map\<^sub>0 b' \\<^sub>D G f \\<^sub>D \\<^sub>0 b)\" proof - let ?g = "\'.map\<^sub>0 b' \\<^sub>D G f \\<^sub>D \\<^sub>0 b" have g_in_hhom: "\?g : b \\<^sub>D b'\" using assms \.ide_map\<^sub>0_obj \.map\<^sub>0_in_hhom by auto have ide_g: "D.ide ?g" using assms g_in_hhom \.ide_map\<^sub>0_obj \'.ide_map\<^sub>0_obj G.preserves_ide by blast let ?\ = "\ f b b'" have "\?\ : G f \\<^sub>D G (F ?g)\" proof (intro D.comp_iso_in_hom) show "\\\<^sub>D\<^sup>-\<^sup>1[G f] : G f \\<^sub>D G f \\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)\" using assms by (intro D.iso_in_homI) auto show "\\\<^sub>D\<^sup>-\<^sup>1[G f \\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)] : G f \\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b) \\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b') \\<^sub>D G f \\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)\" proof - have "D.ide (G f \\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b))" using assms D.ide_hcomp [of "G f" "G.map\<^sub>0 (F.map\<^sub>0 b)"] by (metis D.hcomp_in_hhomE D.hseqE D.objE F.map\<^sub>0_simps(1) G.map\<^sub>0_simps(1) G.preserves_ide GF.map\<^sub>0_simp \.map\<^sub>0_simps(3) g_in_hhom) thus ?thesis using assms \.ide_map\<^sub>0_obj by (intro D.iso_in_homI) auto qed show "\D.inv (\'.counit b') \\<^sub>D G f \\<^sub>D D.inv (\'.counit b) : G.map\<^sub>0 (F.map\<^sub>0 b') \\<^sub>D G f \\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b) \\<^sub>D (\\<^sub>0 b' \\<^sub>D \'.map\<^sub>0 b') \\<^sub>D G f \\<^sub>D (\\<^sub>0 b \\<^sub>D \'.map\<^sub>0 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 "\(\\<^sub>0 b' \\<^sub>D \'.map\<^sub>0 b') \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 b, \'.map\<^sub>0 b] : (\\<^sub>0 b' \\<^sub>D \'.map\<^sub>0 b') \\<^sub>D G f \\<^sub>D (\\<^sub>0 b \\<^sub>D \'.map\<^sub>0 b) \\<^sub>D (\\<^sub>0 b' \\<^sub>D \'.map\<^sub>0 b') \\<^sub>D (G f \\<^sub>D \\<^sub>0 b) \\<^sub>D \'.map\<^sub>0 b\" proof - have "\\\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 b, \'.map\<^sub>0 b] : G f \\<^sub>D (\\<^sub>0 b \\<^sub>D \'.map\<^sub>0 b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 b) \\<^sub>D \'.map\<^sub>0 b\" using assms \.ide_map\<^sub>0_obj \'.ide_map\<^sub>0_obj \.map\<^sub>0_in_hom \'.map\<^sub>0_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 "\\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b' \\<^sub>D \'.map\<^sub>0 b', G f \\<^sub>D \\<^sub>0 b, \'.map\<^sub>0 b] : (\\<^sub>0 b' \\<^sub>D \'.map\<^sub>0 b') \\<^sub>D (G f \\<^sub>D \\<^sub>0 b) \\<^sub>D \'.map\<^sub>0 b \\<^sub>D ((\\<^sub>0 b' \\<^sub>D \'.map\<^sub>0 b') \\<^sub>D (G f \\<^sub>D \\<^sub>0 b)) \\<^sub>D \'.map\<^sub>0 b\" using assms \.ide_map\<^sub>0_obj \'.ide_map\<^sub>0_obj \.map\<^sub>0_in_hom \'.map\<^sub>0_in_hom D.assoc'_in_hom D.iso_assoc' by (intro D.iso_in_homI) auto show "\\\<^sub>D[\\<^sub>0 b', \'.map\<^sub>0 b', G f \\<^sub>D \\<^sub>0 b] \\<^sub>D \'.map\<^sub>0 b : ((\\<^sub>0 b' \\<^sub>D \'.map\<^sub>0 b') \\<^sub>D (G f \\<^sub>D \\<^sub>0 b)) \\<^sub>D \'.map\<^sub>0 b \\<^sub>D (\\<^sub>0 b' \\<^sub>D ?g) \\<^sub>D \'.map\<^sub>0 b\" using assms \.ide_map\<^sub>0_obj \'.ide_map\<^sub>0_obj \.map\<^sub>0_in_hom \'.map\<^sub>0_in_hom D.assoc_in_hom D.iso_assoc by (intro D.hcomp_iso_in_hom D.iso_in_homI) auto show "\D.inv (\\<^sub>1 ?g) \\<^sub>D \'.map\<^sub>0 b : (\\<^sub>0 b' \\<^sub>D ?g) \\<^sub>D \'.map\<^sub>0 b \\<^sub>D (G (F ?g) \\<^sub>D \\<^sub>0 b) \\<^sub>D \'.map\<^sub>0 b\" using assms \.map\<^sub>1_in_hom [of ?g] \.iso_map\<^sub>1_ide \'.map\<^sub>0_in_hom g_in_hhom ide_g by (intro D.hcomp_iso_in_hom D.iso_in_homI) auto show "\\\<^sub>D[G (F ?g), \\<^sub>0 b, \'.map\<^sub>0 b] : (G (F ?g) \\<^sub>D \\<^sub>0 b) \\<^sub>D \'.map\<^sub>0 b \\<^sub>D G (F ?g) \\<^sub>D \\<^sub>0 b \\<^sub>D \'.map\<^sub>0 b\" using assms \.ide_map\<^sub>0_obj \'.ide_map\<^sub>0_obj \.map\<^sub>0_in_hom \'.map\<^sub>0_in_hom D.assoc_in_hom D.iso_assoc by (intro D.iso_in_homI) auto show "\G (F ?g) \\<^sub>D \'.counit b : G (F ?g) \\<^sub>D \\<^sub>0 b \\<^sub>D \'.map\<^sub>0 b \\<^sub>D G (F ?g) \\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b)\" using assms \'.counit_in_hom D.ide_in_hom(2) ide_g by (intro D.hcomp_iso_in_hom) auto show "\\\<^sub>D[G (F ?g)] : G (F ?g) \\<^sub>D G.map\<^sub>0 (F.map\<^sub>0 b) \\<^sub>D G (F ?g)\" using assms ide_g by (intro D.iso_in_homI) auto qed thus "\\ ?\ f (F ?g) : f \\<^sub>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 V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \\<^sub>F proof show "\b. C.obj b \ \a. D.obj a \ C.equivalent_objects (F.map\<^sub>0 a) b" by (metis FG.biessentially_surjective_on_objects FG.map\<^sub>0_simp G.map\<^sub>0_simps(1)) show "\g g' \. \D.ide g; D.ide g'; src\<^sub>D g = src\<^sub>D g'; trg\<^sub>D g = trg\<^sub>D g'; \\ : F g \\<^sub>C F g'\\ \ \\. \\ : g \\<^sub>D g'\ \ F \ = \" proof - fix g g' \ assume g: "D.ide g" and g': "D.ide g'" assume eq_src: "src\<^sub>D g = src\<^sub>D g'" and eq_trg: "trg\<^sub>D g = trg\<^sub>D g'" assume \: "\\ : F g \\<^sub>C F g'\" obtain \ where \: "\\ : g \\<^sub>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 \\<^sub>D g'\ \ F \ = \" using \ by auto qed show "\b b' f. \D.obj b; D.obj b'; \f : F.map\<^sub>0 b \\<^sub>C F.map\<^sub>0 b'\; C.ide f\ \ \g. \g : b \\<^sub>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.map\<^sub>0 b \\<^sub>C F.map\<^sub>0 b'\" assume ide_f: "C.ide f" let ?g = "\'.map\<^sub>0 b' \\<^sub>D G f \\<^sub>D \\<^sub>0 b" have g_in_hhom: "\?g : b \\<^sub>D b'\" using b b' f \.ide_map\<^sub>0_obj \.map\<^sub>0_in_hhom by auto have ide_g: "D.ide ?g" using b b' f ide_f g_in_hhom \.ide_map\<^sub>0_obj \'.ide_map\<^sub>0_obj G.preserves_ide by blast have "f \\<^sub>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 \\<^sub>D b'\ \ D.ide g \ F g \\<^sub>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 V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \\<^sub>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 = (* 40 sec *) + locale equivalence_of_bicategories_and_pseudonatural_equivalence_left = (* 30 sec *) C: bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + D: bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + E: equivalence_of_bicategories + \: pseudonatural_equivalence V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \\<^sub>F F' \\<^sub>F' \\<^sub>0 \\<^sub>1 for F' and \\<^sub>F' and \\<^sub>0 and \\<^sub>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 V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F' \\<^sub>F' G \\<^sub>G .. interpretation F'G: composite_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C G \\<^sub>G F' \\<^sub>F' .. interpretation \': converse_pseudonatural_equivalence V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \\<^sub>F F' \\<^sub>F' \\<^sub>0 \\<^sub>1 .. interpretation \'oG: pseudonatural_equivalence_whisker_right V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F' \\<^sub>F' F \\<^sub>F G \\<^sub>G \'.map\<^sub>0 \'.map\<^sub>1 .. interpretation Go\: pseudonatural_equivalence_whisker_left V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \\<^sub>F F' \\<^sub>F' G \\<^sub>G \\<^sub>0 \\<^sub>1 .. sublocale unit: composite_pseudonatural_equivalence V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D E.I\<^sub>D.map E.I\<^sub>D.cmp E.GF.map E.GF.cmp GF'.map GF'.cmp \\<^sub>0 \\<^sub>1 Go\.map\<^sub>0 Go\.map\<^sub>1 .. sublocale counit: composite_pseudonatural_equivalence V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F'G.map F'G.cmp E.FG.map E.FG.cmp E.I\<^sub>C.map E.I\<^sub>C.cmp \'oG.map\<^sub>0 \'oG.map\<^sub>1 \\<^sub>0 \\<^sub>1 .. - sublocale equivalence_of_bicategories - V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F' \\<^sub>F' G \\<^sub>G - unit.map\<^sub>0 unit.map\<^sub>1 counit.map\<^sub>0 counit.map\<^sub>1 - .. - - abbreviation unit\<^sub>0 + + definition unit\<^sub>0 where "unit\<^sub>0 \ unit.map\<^sub>0" - abbreviation unit\<^sub>1 + definition unit\<^sub>1 where "unit\<^sub>1 \ unit.map\<^sub>1" - abbreviation counit\<^sub>0 + definition counit\<^sub>0 where "counit\<^sub>0 \ counit.map\<^sub>0" - abbreviation counit\<^sub>1 + definition counit\<^sub>1 where "counit\<^sub>1 \ counit.map\<^sub>1" + sublocale equivalence_of_bicategories + V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F' \\<^sub>F' G \\<^sub>G + unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1 + unfolding unit\<^sub>0_def unit\<^sub>1_def counit\<^sub>0_def counit\<^sub>1_def + .. + lemma induces_equivalence_of_bicategories: shows "equivalence_of_bicategories V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F' \\<^sub>F' G \\<^sub>G unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1" .. end - locale equivalence_of_bicategories_and_pseudonatural_equivalence_right = (* 40 sec *) + locale equivalence_of_bicategories_and_pseudonatural_equivalence_right = (* 30 sec *) C: bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + D: bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + E: equivalence_of_bicategories + \: pseudonatural_equivalence V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D G \\<^sub>G G' \\<^sub>G' \\<^sub>0 \\<^sub>1 for G' and \\<^sub>G' and \\<^sub>0 and \\<^sub>1 begin interpretation G'F: composite_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \\<^sub>F G' \\<^sub>G' .. interpretation FG': composite_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C G' \\<^sub>G' F \\<^sub>F .. interpretation \': converse_pseudonatural_equivalence V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D G \\<^sub>G G' \\<^sub>G' \\<^sub>0 \\<^sub>1 .. interpretation \oF: pseudonatural_equivalence_whisker_right V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D G \\<^sub>G G' \\<^sub>G' F \\<^sub>F \\<^sub>0 \\<^sub>1 .. interpretation Fo\': pseudonatural_equivalence_whisker_left V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C G' \\<^sub>G' G \\<^sub>G F \\<^sub>F \'.map\<^sub>0 \'.map\<^sub>1 .. sublocale unit: composite_pseudonatural_equivalence V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D E.I\<^sub>D.map E.I\<^sub>D.cmp E.GF.map E.GF.cmp G'F.map G'F.cmp \\<^sub>0 \\<^sub>1 \oF.map\<^sub>0 \oF.map\<^sub>1 .. sublocale counit: composite_pseudonatural_equivalence V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C FG'.map FG'.cmp E.FG.map E.FG.cmp E.I\<^sub>C.map E.I\<^sub>C.cmp Fo\'.map\<^sub>0 Fo\'.map\<^sub>1 \\<^sub>0 \\<^sub>1 .. + + definition unit\<^sub>0 + where "unit\<^sub>0 \ unit.map\<^sub>0" + + definition unit\<^sub>1 + where "unit\<^sub>1 \ unit.map\<^sub>1" + + definition counit\<^sub>0 + where "counit\<^sub>0 \ counit.map\<^sub>0" + + definition counit\<^sub>1 + where "counit\<^sub>1 \ counit.map\<^sub>1" + sublocale equivalence_of_bicategories V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \\<^sub>F G' \\<^sub>G' - unit.map\<^sub>0 unit.map\<^sub>1 counit.map\<^sub>0 counit.map\<^sub>1 + unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1 + unfolding unit\<^sub>0_def unit\<^sub>1_def counit\<^sub>0_def counit\<^sub>1_def .. - abbreviation unit\<^sub>0 - where "unit\<^sub>0 \ unit.map\<^sub>0" - - abbreviation unit\<^sub>1 - where "unit\<^sub>1 \ unit.map\<^sub>1" - - abbreviation counit\<^sub>0 - where "counit\<^sub>0 \ counit.map\<^sub>0" - - abbreviation counit\<^sub>1 - where "counit\<^sub>1 \ counit.map\<^sub>1" - lemma induces_equivalence_of_bicategories: shows "equivalence_of_bicategories V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \\<^sub>F G' \\<^sub>G' unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1" .. 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 = (* 35 sec *) + locale converse_equivalence_of_bicategories = (* 25 sec *) C: bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + D: bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + I\<^sub>C: identity_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + I\<^sub>D: identity_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + E: equivalence_of_bicategories begin sublocale counit: converse_pseudonatural_equivalence V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D I\<^sub>D.map I\<^sub>D.cmp E.GF.map E.GF.cmp \\<^sub>0 \\<^sub>1 .. sublocale unit: converse_pseudonatural_equivalence V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C E.FG.map E.FG.cmp I\<^sub>C.map I\<^sub>C.cmp \\<^sub>0 \\<^sub>1 .. + definition unit\<^sub>0 + where "unit\<^sub>0 \ unit.map\<^sub>0" + + definition unit\<^sub>1 + where "unit\<^sub>1 \ unit.map\<^sub>1" + + definition counit\<^sub>0 + where "counit\<^sub>0 \ counit.map\<^sub>0" + + definition counit\<^sub>1 + where "counit\<^sub>1 \ counit.map\<^sub>1" + sublocale equivalence_of_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C G \\<^sub>G F \\<^sub>F - unit.map\<^sub>0 unit.map\<^sub>1 counit.map\<^sub>0 counit.map\<^sub>1 + unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1 + unfolding unit\<^sub>0_def unit\<^sub>1_def counit\<^sub>0_def counit\<^sub>1_def .. - abbreviation unit\<^sub>0 - where "unit\<^sub>0 \ unit.map\<^sub>0" - - abbreviation unit\<^sub>1 - where "unit\<^sub>1 \ unit.map\<^sub>1" - - abbreviation counit\<^sub>0 - where "counit\<^sub>0 \ counit.map\<^sub>0" - - abbreviation counit\<^sub>1 - where "counit\<^sub>1 \ counit.map\<^sub>1" - lemma is_equivalence_of_bicategories: shows "equivalence_of_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C G \\<^sub>G F \\<^sub>F unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1" .. 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 = (* 80 sec *) + locale composite_equivalence_of_bicategories = (* 60 sec *) B: bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B + C: bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + D: bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + I\<^sub>B: identity_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B + I\<^sub>C: identity_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + I\<^sub>D: identity_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + F_: pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B F \\<^sub>F + G_: pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C G \\<^sub>G + H: pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C H \\<^sub>H + K: pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D K \\<^sub>K + FG: equivalence_of_bicategories V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \\<^sub>F G \\<^sub>G \\<^sub>0 \\<^sub>1 \\<^sub>0 \\<^sub>1 + HK: equivalence_of_bicategories V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D H \\<^sub>H K \\<^sub>K \\<^sub>0 \\<^sub>1 \\<^sub>0 \\<^sub>1 for V\<^sub>B :: "'b comp" (infixr "\\<^sub>B" 55) and H\<^sub>B :: "'b comp" (infixr "\\<^sub>B" 53) and \\<^sub>B :: "'b \ 'b \ 'b \ 'b" ("\\<^sub>B[_, _, _]") and \\<^sub>B :: "'b \ 'b" ("\\<^sub>B[_]") and src\<^sub>B :: "'b \ 'b" and trg\<^sub>B :: "'b \ 'b" and V\<^sub>C :: "'c comp" (infixr "\\<^sub>C" 55) and H\<^sub>C :: "'c comp" (infixr "\\<^sub>C" 53) and \\<^sub>C :: "'c \ 'c \ 'c \ 'c" ("\\<^sub>C[_, _, _]") and \\<^sub>C :: "'c \ 'c" ("\\<^sub>C[_]") and src\<^sub>C :: "'c \ 'c" and trg\<^sub>C :: "'c \ 'c" and V\<^sub>D :: "'d comp" (infixr "\\<^sub>D" 55) and H\<^sub>D :: "'d comp" (infixr "\\<^sub>D" 53) and \\<^sub>D :: "'d \ 'd \ 'd \ 'd" ("\\<^sub>D[_, _, _]") and \\<^sub>D :: "'d \ 'd" ("\\<^sub>D[_]") and src\<^sub>D :: "'d \ 'd" and trg\<^sub>D :: "'d \ 'd" and F :: "'c \ 'b" and \\<^sub>F :: "'c * 'c \ 'b" and G :: "'b \ 'c" and \\<^sub>G :: "'b * 'b \ 'c" and H :: "'d \ 'c" and \\<^sub>H :: "'d * 'd \ 'c" and K :: "'c \ 'd" and \\<^sub>K :: "'c * 'c \ 'd" and \\<^sub>0 :: "'c \ 'c" and \\<^sub>1 :: "'c \ 'c" and \\<^sub>0 :: "'b \ 'b" and \\<^sub>1 :: "'b \ 'b" and \\<^sub>0 :: "'d \ 'd" and \\<^sub>1 :: "'d \ 'd" and \\<^sub>0 :: "'c \ 'c" and \\<^sub>1 :: "'c \ 'c" begin notation B.\' ("\\<^sub>B\<^sup>-\<^sup>1[_, _, _]") text \ At this point we could make the explicit definitions: \begin{itemize} \item \\\<^sub>0 = K (\\<^sub>0 (H.map\<^sub>0 a)) \\<^sub>D \\<^sub>0 a\ \item \\\<^sub>1 = \\<^sub>D\<^sup>-\<^sup>1[K (\\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))), \\<^sub>0 (trg\<^sub>D f), f] \\<^sub>D (K (\\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))) \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[K (\\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))), K (H f), \\<^sub>0 (src\<^sub>D f)] \\<^sub>D (D.inv (\\<^sub>K (\\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f)), H f)) \\<^sub>D K (\\<^sub>1 (H f)) \\<^sub>D \\<^sub>K (G (F (H f)), \\<^sub>0 (H.map\<^sub>0 (src\<^sub>D f))) \\<^sub>D \\<^sub>0 (src\<^sub>D f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[K (G (F (H f))), K (\\<^sub>0 (H.map\<^sub>0 (src\<^sub>D f))), \\<^sub>0 (src\<^sub>D f)]\ \item \\\<^sub>0 = \\<^sub>0 a \\<^sub>B F (\\<^sub>0 (G_.map\<^sub>0 a))\ \item \\\<^sub>1 = \\<^sub>B\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>B f), F (\\<^sub>0 (G_.map\<^sub>0 (trg\<^sub>B f))), F (H (K (G f)))] \\<^sub>B (\\<^sub>0 (trg\<^sub>B f) \\<^sub>B B.inv (\\<^sub>F (\\<^sub>0 (G_.map\<^sub>0 (trg\<^sub>B f)), H (K (G f)))) \\<^sub>B F (\\<^sub>1 (G f)) \\<^sub>B \\<^sub>F (G f, \\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))) \\<^sub>B \\<^sub>B[\\<^sub>0 (trg\<^sub>B f), F (G f), F (\\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))] \\<^sub>B (\\<^sub>1 f \\<^sub>B F (\\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, \\<^sub>0 (src\<^sub>B f), F (\\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B 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 V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B H \\<^sub>H F \\<^sub>F .. sublocale KG: composite_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D G \\<^sub>G K \\<^sub>K .. interpretation IG: composite_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C G \\<^sub>G I\<^sub>C.map I\<^sub>C.cmp .. interpretation IH: composite_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C H \\<^sub>H I\<^sub>C.map I\<^sub>C.cmp .. interpretation HKG: composite_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C G \\<^sub>G HK.FG.map HK.FG.cmp .. interpretation GFH: composite_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C H \\<^sub>H FG.GF.map FG.GF.cmp .. interpretation KGFH: composite_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D GFH.map GFH.cmp K \\<^sub>K .. interpretation FHKG: composite_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B HKG.map HKG.cmp F \\<^sub>F - .. + .. interpretation \oH: pseudonatural_equivalence_whisker_right V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C I\<^sub>C.map I\<^sub>C.cmp FG.GF.map FG.GF.cmp H \\<^sub>H \\<^sub>0 \\<^sub>1 .. interpretation Ko\oH: pseudonatural_equivalence_whisker_left V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D H \\<^sub>H GFH.map GFH.cmp K \\<^sub>K \oH.map\<^sub>0 \oH.map\<^sub>1 proof - interpret Ko\oH: pseudonatural_equivalence_whisker_left V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D IH.map IH.cmp GFH.map GFH.cmp K \\<^sub>K \oH.map\<^sub>0 \oH.map\<^sub>1 .. have "IH.map = H" using H.is_extensional IH.is_extensional H.functor_axioms by force moreover have "IH.cmp = \\<^sub>H" proof fix \\ show "IH.cmp \\ = \\<^sub>H \\" using IH.cmp_def D.VV.arr_char 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 V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D H \\<^sub>H GFH.map GFH.cmp K \\<^sub>K \oH.map\<^sub>0 \oH.map\<^sub>1" using Ko\oH.pseudonatural_equivalence_whisker_left_axioms by simp qed interpretation \oG: pseudonatural_equivalence_whisker_right V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C HK.FG.map HK.FG.cmp I\<^sub>C.map I\<^sub>C.cmp G \\<^sub>G \\<^sub>0 \\<^sub>1 .. interpretation Fo\oG: pseudonatural_equivalence_whisker_left V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B HKG.map HKG.cmp G \\<^sub>G F \\<^sub>F \oG.map\<^sub>0 \oG.map\<^sub>1 proof - interpret Fo\oG: pseudonatural_equivalence_whisker_left V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B HKG.map HKG.cmp IG.map IG.cmp F \\<^sub>F \oG.map\<^sub>0 \oG.map\<^sub>1 .. have "IG.map = G" using G_.is_extensional IG.is_extensional by (meson G_.functor_axioms comp_identity_functor) moreover have "IG.cmp = \\<^sub>G" proof fix \\ show "IG.cmp \\ = \\<^sub>G \\" using IG.cmp_def B.VV.arr_char 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 V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B HKG.map HKG.cmp G \\<^sub>G F \\<^sub>F \oG.map\<^sub>0 \oG.map\<^sub>1" using Fo\oG.pseudonatural_equivalence_whisker_left_axioms by simp qed sublocale unit: composite_pseudonatural_equivalence V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D I\<^sub>D.map I\<^sub>D.cmp HK.GF.map HK.GF.cmp KGFH.map KGFH.cmp \\<^sub>0 \\<^sub>1 Ko\oH.map\<^sub>0 Ko\oH.map\<^sub>1 .. sublocale counit: composite_pseudonatural_equivalence V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B FHKG.map FHKG.cmp FG.FG.map FG.FG.cmp I\<^sub>B.map I\<^sub>B.cmp Fo\oG.map\<^sub>0 Fo\oG.map\<^sub>1 \\<^sub>0 \\<^sub>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" - abbreviation unit\<^sub>0 + definition unit\<^sub>0 where "unit\<^sub>0 \ unit.map\<^sub>0" - abbreviation unit\<^sub>1 + definition unit\<^sub>1 where "unit\<^sub>1 \ unit.map\<^sub>1" - abbreviation counit\<^sub>0 + definition counit\<^sub>0 where "counit\<^sub>0 \ counit.map\<^sub>0" - abbreviation counit\<^sub>1 + definition counit\<^sub>1 where "counit\<^sub>1 == counit.map\<^sub>1" lemma unit\<^sub>0_simp: assumes "D.obj a" shows "unit\<^sub>0 a = K (\\<^sub>0 (H.map\<^sub>0 a)) \\<^sub>D \\<^sub>0 a" - using assms unit.map\<^sub>0_def Ko\oH.map\<^sub>0_def \oH.map\<^sub>0_def by simp + using assms unit\<^sub>0_def unit.map\<^sub>0_def Ko\oH.map\<^sub>0_def \oH.map\<^sub>0_def by simp lemma unit\<^sub>1_simp: assumes "D.ide f" shows "unit\<^sub>1 f = \\<^sub>D\<^sup>-\<^sup>1[K (\\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))), \\<^sub>0 (trg\<^sub>D f), f] \\<^sub>D (K (\\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))) \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[K (\\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f))), K (H f), \\<^sub>0 (src\<^sub>D f)] \\<^sub>D (D.inv (\\<^sub>K (\\<^sub>0 (H.map\<^sub>0 (trg\<^sub>D f)), H f)) \\<^sub>D K (\\<^sub>1 (H f)) \\<^sub>D \\<^sub>K (G (F (H f)), \\<^sub>0 (H.map\<^sub>0 (src\<^sub>D f))) \\<^sub>D \\<^sub>0 (src\<^sub>D f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[K (G (F (H f))), K (\\<^sub>0 (H.map\<^sub>0 (src\<^sub>D f))), \\<^sub>0 (src\<^sub>D f)]" - using assms unit.map\<^sub>1_def Ko\oH.map\<^sub>0_def \oH.map\<^sub>0_def Ko\oH.map\<^sub>1_def \oH.map\<^sub>1_def + using assms unit\<^sub>1_def unit.map\<^sub>1_def Ko\oH.map\<^sub>0_def \oH.map\<^sub>0_def Ko\oH.map\<^sub>1_def \oH.map\<^sub>1_def by simp lemma counit\<^sub>0_simp: assumes "B.obj a" shows "counit\<^sub>0 a = \\<^sub>0 a \\<^sub>B F (\\<^sub>0 (G_.map\<^sub>0 a))" - using assms counit.map\<^sub>0_def Fo\oG.map\<^sub>0_def \oG.map\<^sub>0_def by simp + using assms counit\<^sub>0_def counit.map\<^sub>0_def Fo\oG.map\<^sub>0_def \oG.map\<^sub>0_def by simp lemma counit\<^sub>1_simp: assumes "B.ide f" shows "counit\<^sub>1 f = \\<^sub>B\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>B f), F (\\<^sub>0 (G_.map\<^sub>0 (trg\<^sub>B f))), F (H (K (G f)))] \\<^sub>B (\\<^sub>0 (trg\<^sub>B f) \\<^sub>B B.inv (\\<^sub>F (\\<^sub>0 (G_.map\<^sub>0 (trg\<^sub>B f)), H (K (G f)))) \\<^sub>B F (\\<^sub>1 (G f)) \\<^sub>B \\<^sub>F (G f, \\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))) \\<^sub>B \\<^sub>B[\\<^sub>0 (trg\<^sub>B f), F (G f), F (\\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))] \\<^sub>B (\\<^sub>1 f \\<^sub>B F (\\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, \\<^sub>0 (src\<^sub>B f), F (\\<^sub>0 (G_.map\<^sub>0 (src\<^sub>B f)))]" - using assms counit.map\<^sub>1_def Fo\oG.map\<^sub>0_def Fo\oG.map\<^sub>1_def \oG.map\<^sub>0_def \oG.map\<^sub>1_def + using assms counit\<^sub>1_def counit.map\<^sub>1_def Fo\oG.map\<^sub>0_def Fo\oG.map\<^sub>1_def + \oG.map\<^sub>0_def \oG.map\<^sub>1_def by simp sublocale equivalence_of_bicategories V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D - FH.map FH.cmp KG.map KG.cmp - unit.map\<^sub>0 unit.map\<^sub>1 counit.map\<^sub>0 counit.map\<^sub>1 + FH.map FH.cmp KG.map KG.cmp unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1 proof - interpret FH_KG: composite_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B KG.map KG.cmp FH.map FH.cmp .. interpret KG_FH: composite_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D 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 (I\<^sub>B.cmp \\)))) \\<^sub>B F (H (K (G (B.dom (fst \\) \\<^sub>B B.dom (snd \\))))) \\<^sub>B F (H (K (\\<^sub>G (B.dom (fst \\), B.dom (snd \\))))) \\<^sub>B F (H (\\<^sub>K (G (B.dom (fst \\)), G (B.dom (snd \\))))) \\<^sub>B (F (H (K (G (B.dom (fst \\))) \\<^sub>D K (G (B.dom (snd \\))))) \\<^sub>B F (\\<^sub>H (K (G (B.dom (fst \\))), K (G (B.dom (snd \\)))))) \\<^sub>B \\<^sub>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_char B.VV.dom_simp B.VV.cod_simp C.VV.arr_char C.VV.dom_simp C.VV.cod_simp D.VV.arr_char 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 (I\<^sub>B.cmp \\)))) \\<^sub>B F (H (K (G (B.dom (fst \\) \\<^sub>B B.dom (snd \\))))) \\<^sub>B F (H (K (\\<^sub>G (B.dom (fst \\), B.dom (snd \\))))) \\<^sub>B F (H (\\<^sub>K (G (B.dom (fst \\)), G (B.dom (snd \\))))) \\<^sub>B F (\\<^sub>H (K (G (B.dom (fst \\))), K (G (B.dom (snd \\))))) \\<^sub>B \\<^sub>F (H (K (G (B.dom (fst \\)))), H (K (G (B.dom (snd \\)))))" proof - have "F (H (K (G (B.dom (fst \\))) \\<^sub>D K (G (B.dom (snd \\))))) \\<^sub>B F (\\<^sub>H (K (G (B.dom (fst \\))), K (G (B.dom (snd \\))))) = F (\\<^sub>H (K (G (B.dom (fst \\))), K (G (B.dom (snd \\)))))" using True B.VV.arr_char B.VV.dom_simp B.VV.cod_simp C.VV.arr_char C.VV.dom_simp C.VV.cod_simp D.VV.arr_char 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 (I\<^sub>B.cmp \\)))) \\<^sub>B F (H (K (G (B.dom (fst \\) \\<^sub>B B.dom (snd \\))))) \\<^sub>B F (H (K (\\<^sub>G (B.dom (fst \\), B.dom (snd \\))))) \\<^sub>B (F (H (K (G (B.dom (fst \\)) \\<^sub>C G (B.dom (snd \\))))) \\<^sub>B F (H (\\<^sub>K (G (B.dom (fst \\)), G (B.dom (snd \\)))))) \\<^sub>B F (\\<^sub>H (K (G (B.dom (fst \\))), K (G (B.dom (snd \\))))) \\<^sub>B \\<^sub>F (H (K (G (B.dom (fst \\)))), H (K (G (B.dom (snd \\)))))" proof - have "F (H (K (G (B.dom (fst \\)) \\<^sub>C G (B.dom (snd \\))))) \\<^sub>B F (H (\\<^sub>K (G (B.dom (fst \\)), G (B.dom (snd \\))))) = F (H (\\<^sub>K (G (B.dom (fst \\)), G (B.dom (snd \\)))))" using True B.VV.arr_char B.VV.dom_simp B.VV.cod_simp C.VV.arr_char C.VV.dom_simp C.VV.cod_simp D.VV.arr_char 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_char B.VV.dom_simp B.VV.cod_simp C.VV.arr_char C.VV.dom_simp C.VV.cod_simp D.VV.arr_char 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 (I\<^sub>D.cmp \\)))) \\<^sub>D K (G (F (H (D.dom (fst \\) \\<^sub>D D.dom (snd \\))))) \\<^sub>D K (G (F (\\<^sub>H (D.dom (fst \\), D.dom (snd \\))))) \\<^sub>D K (G (\\<^sub>F (H (D.dom (fst \\)), H (D.dom (snd \\))))) \\<^sub>D (K (G (F (H (D.dom (fst \\))) \\<^sub>B F (H (D.dom (snd \\))))) \\<^sub>D K (\\<^sub>G (F (H (D.dom (fst \\))), F (H (D.dom (snd \\)))))) \\<^sub>D \\<^sub>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_char B.VV.dom_simp B.VV.cod_simp C.VV.arr_char C.VV.dom_simp C.VV.cod_simp D.VV.arr_char 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 (I\<^sub>D.cmp \\)))) \\<^sub>D K (G (F (H (D.dom (fst \\) \\<^sub>D D.dom (snd \\))))) \\<^sub>D K (G (F (\\<^sub>H (D.dom (fst \\), D.dom (snd \\))))) \\<^sub>D K (G (\\<^sub>F (H (D.dom (fst \\)), H (D.dom (snd \\))))) \\<^sub>D K (\\<^sub>G (F (H (D.dom (fst \\))), F (H (D.dom (snd \\))))) \\<^sub>D \\<^sub>K (G (F (H (D.dom (fst \\)))), G (F (H (D.dom (snd \\)))))" proof - have "K (G (F (H (D.dom (fst \\))) \\<^sub>B F (H (D.dom (snd \\))))) \\<^sub>D K (\\<^sub>G (F (H (D.dom (fst \\))), F (H (D.dom (snd \\))))) = K (\\<^sub>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_char B.VV.dom_simp B.VV.cod_simp C.VV.arr_char C.VV.dom_simp C.VV.cod_simp D.VV.arr_char 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 (I\<^sub>D.cmp \\)))) \\<^sub>D K (G (F (H (D.dom (fst \\) \\<^sub>D D.dom (snd \\))))) \\<^sub>D K (G (F (\\<^sub>H (D.dom (fst \\), D.dom (snd \\))))) \\<^sub>D (K (G (F (H (D.dom (fst \\)) \\<^sub>C H (D.dom (snd \\))))) \\<^sub>D K (G (\\<^sub>F (H (D.dom (fst \\)), H (D.dom (snd \\)))))) \\<^sub>D K (\\<^sub>G (F (H (D.dom (fst \\))), F (H (D.dom (snd \\))))) \\<^sub>D \\<^sub>K (G (F (H (D.dom (fst \\)))), G (F (H (D.dom (snd \\)))))" proof - have "K (G (F (H (D.dom (fst \\)) \\<^sub>C H (D.dom (snd \\))))) \\<^sub>D K (G (\\<^sub>F (H (D.dom (fst \\)), H (D.dom (snd \\))))) = K (G (\\<^sub>F (H (D.dom (fst \\)), H (D.dom (snd \\)))))" using True KGFH.cmp_def GFH.cmp_def FG.GF.cmp_def B.VV.arr_char B.VV.dom_simp B.VV.cod_simp C.VV.arr_char C.VV.dom_simp C.VV.cod_simp D.VV.arr_char 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_char B.VV.dom_simp B.VV.cod_simp C.VV.arr_char C.VV.dom_simp C.VV.cod_simp D.VV.arr_char 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 V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D - FH.map FH.cmp KG.map KG.cmp - unit.map\<^sub>0 unit.map\<^sub>1 counit.map\<^sub>0 counit.map\<^sub>1" + FH.map FH.cmp KG.map KG.cmp unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1" + unfolding unit\<^sub>0_def unit\<^sub>1_def counit\<^sub>0_def counit\<^sub>1_def using B.bicategory_axioms D.bicategory_axioms FH.pseudofunctor_axioms KG.pseudofunctor_axioms unit.pseudonatural_equivalence_axioms counit.pseudonatural_equivalence_axioms I\<^sub>B.identity_pseudofunctor_axioms I\<^sub>D.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 V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D left_map left_cmp right_map right_cmp unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1" .. 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 V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B + subbicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \\\. B.arr \ \ src\<^sub>B \ \ Obj \ trg\<^sub>B \ \ Obj\ for V\<^sub>B :: "'b comp" (infixr "\\<^sub>B" 55) and H\<^sub>B :: "'b comp" (infixr "\\<^sub>B" 53) and \\<^sub>B :: "'b \ 'b \ 'b \ 'b" ("\\<^sub>B[_, _, _]") and \\<^sub>B :: "'b \ 'b" ("\\<^sub>B[_]") and src\<^sub>B :: "'b \ 'b" and trg\<^sub>B :: "'b \ 'b" and Obj :: "'b set" + assumes dense: "\a. B.obj a \ \a'. a' \ Obj \ B.equivalent_objects a' a" begin notation B.\' ("\\<^sub>B\<^sup>-\<^sup>1[_, _, _]") notation B.lunit ("\\<^sub>B[_]") notation B.lunit' ("\\<^sub>B\<^sup>-\<^sup>1[_]") notation B.runit ("\\<^sub>B[_]") notation B.runit' ("\\<^sub>B\<^sup>-\<^sup>1[_]") notation comp (infixr "\" 55) notation hcomp (infixr "\" 53) notation in_hom ("\_ : _ \ _\") notation in_hhom ("\_ : _ \ _\") notation \ ("\[_, _, _]") notation \' ("\\<^sup>-\<^sup>1[_, _, _]") notation lunit ("\[_]") notation lunit' ("\\<^sup>-\<^sup>1[_]") notation runit ("\[_]") notation runit' ("\\<^sup>-\<^sup>1[_]") abbreviation (input) Arr where "Arr \ \\. B.arr \ \ src\<^sub>B \ \ Obj \ trg\<^sub>B \ \ Obj" sublocale emb: embedding_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B Arr .. abbreviation E where "E \ emb.map" abbreviation \\<^sub>E where "\\<^sub>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 P\<^sub>0 where "P\<^sub>0 a \ SOME a'. obj a' \ B.equivalent_objects a' a" lemma P\<^sub>0_props: assumes "B.obj a" shows "obj (P\<^sub>0 a)" and "B.equivalent_objects (P\<^sub>0 a) a" and "B.equivalent_objects a a' \ P\<^sub>0 a = P\<^sub>0 a'" and "P\<^sub>0 (P\<^sub>0 a) = P\<^sub>0 a" and "B.obj (P\<^sub>0 a)" and "P\<^sub>0 a \ Obj" proof - have "\a'. obj a' \ B.equivalent_objects a' a" using assms dense [of a] obj_char arr_char by (metis (no_types, lifting) B.equivalent_objects_def B.in_hhomE B.obj_def B.obj_simps(3)) hence 1: "obj (P\<^sub>0 a) \ B.equivalent_objects (P\<^sub>0 a) a" unfolding P\<^sub>0_def using someI_ex [of "\a'. obj a' \ B.equivalent_objects a' a"] by blast show "obj (P\<^sub>0 a)" using 1 by simp show 2: "B.equivalent_objects (P\<^sub>0 a) a" using 1 by simp show 3: "\a'. B.equivalent_objects a a' \ P\<^sub>0 a = P\<^sub>0 a'" unfolding P\<^sub>0_def using B.equivalent_objects_symmetric B.equivalent_objects_transitive by meson show "P\<^sub>0 (P\<^sub>0 a) = P\<^sub>0 a" using 2 3 [of "P\<^sub>0 a"] B.equivalent_objects_symmetric by auto show "B.obj (P\<^sub>0 a)" using 1 B.equivalent_objects_def by auto thus "P\<^sub>0 a \ Obj" using 1 3 obj_char arr_char by auto qed text \ For each object \a\ of \B\, we choose an adjoint equivalence from \a\ to \P\<^sub>0 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 \\<^sub>B P\<^sub>0 a\ \ (\d \ \. adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B e d \ \))" definition d where "d a = (SOME d. \\ \. adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B (e a) d \ \)" definition \ where "\ a = (SOME \. \\. adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) \ \)" definition \ where "\ a = (SOME \. adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) (\ a) \)" lemma chosen_adjoint_equivalence: assumes "B.obj a" shows "adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) (\ a) (\ a)" and "\e a : a \\<^sub>B P\<^sub>0 a\" and "B.ide (d a)" and "B.ide (e a)" and "B.iso (\ a)" and "B.iso (\ a)" proof - have "\e. \e : a \\<^sub>B P\<^sub>0 a\ \ (\d \ \. adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B e d \ \)" proof - obtain e where e: "\e : a \\<^sub>B P\<^sub>0 a\ \ B.equivalence_map e" using assms P\<^sub>0_props(2) B.equivalent_objects_symmetric B.equivalent_objects_def by meson obtain d \ \ where d: "adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B 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 \\<^sub>B P\<^sub>0 a\ \ adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) (\ a) (\ a)" using d_def e_def \_def \_def arr_char someI_ex [of "\e. \e : a \\<^sub>B P\<^sub>0 a\ \ (\d \ \. adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B e d \ \)"] someI_ex [of "\d. (\\ \. adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B (e a) d \ \)"] someI_ex [of "\\. \\. adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) \ \"] someI_ex [of "\\. adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) (\ a) \"] by simp interpret adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e a\ \d a\ \\ a\ \\ a\ using 1 by simp show "adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B (e a) (d a) (\ a) (\ a)" .. show "\e a : a \\<^sub>B P\<^sub>0 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_hom\<^sub>B [intro]: assumes "B.obj a" shows "\e a : a \\<^sub>B P\<^sub>0 a\" and "\d a : P\<^sub>0 a \\<^sub>B a\" and "\e a : e a \\<^sub>B e a\" and "\d a : d a \\<^sub>B d a\" and "\\ a : a \\<^sub>B a\" and "\\ a : P\<^sub>0 a \\<^sub>B P\<^sub>0 a\" and "\\ a : a \\<^sub>B d a \\<^sub>B e a\" and "\\ a : e a \\<^sub>B d a \\<^sub>B P\<^sub>0 a\" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e a\ \d a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence by simp show e: "\e a : a \\<^sub>B P\<^sub>0 a\" using assms chosen_adjoint_equivalence by simp show "\d a : P\<^sub>0 a \\<^sub>B a\" using e antipar by auto show "\e a : e a \\<^sub>B e a\" by auto show "\d a : d a \\<^sub>B d a\" by auto show "\\ a : a \\<^sub>B a\" using unit_in_hom e by auto show "\\ a : P\<^sub>0 a \\<^sub>B P\<^sub>0 a\" using counit_in_hom e by auto show "\\ a : a \\<^sub>B d a \\<^sub>B e a\" using unit_in_hom e by auto show "\\ a : e a \\<^sub>B d a \\<^sub>B P\<^sub>0 a\" using counit_in_hom e by auto qed lemma equivalence_data_simps\<^sub>B [simp]: assumes "B.obj a" shows "B.ide (d a)" and "B.ide (e a)" and "B.iso (\ a)" and "B.iso (\ a)" and "src\<^sub>B (e a) = a" and "trg\<^sub>B (e a) = P\<^sub>0 a" and "src\<^sub>B (d a) = P\<^sub>0 a" and "trg\<^sub>B (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 "src\<^sub>B (\ a) = a" and "trg\<^sub>B (\ a) = a" and "src\<^sub>B (\ a) = P\<^sub>0 a" and "trg\<^sub>B (\ a) = P\<^sub>0 a" and "B.dom (\ a) = a" and "B.cod (\ a) = d a \\<^sub>B e a" and "B.dom (\ a) = e a \\<^sub>B d a" and "B.cod (\ a) = P\<^sub>0 a" using assms chosen_adjoint_equivalence equivalence_data_in_hom\<^sub>B B.in_hhom_def apply auto by (meson B.in_homE)+ lemma equivalence_data_in_hom [intro]: assumes "obj a" shows "\e a : a \ P\<^sub>0 a\" and "\d a : P\<^sub>0 a \ a\" and "\e a : e a \ e a\" and "\d a : d a \ d a\" and "\\ a : a \ a\" and "\\ a : P\<^sub>0 a \ P\<^sub>0 a\" and "\\ a : a \ d a \ e a\" and "\\ a : e a \ d a \ P\<^sub>0 a\" proof - have a: "B.obj a \ a \ Obj" using assms obj_char arr_char by auto have P\<^sub>0a: "obj (P\<^sub>0 a) \ P\<^sub>0 a \ Obj" using assms a P\<^sub>0_props [of a] arr_char obj_char by auto show ea: "\e a : a \ P\<^sub>0 a\" using a P\<^sub>0a arr_char 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_char in_hom_char by auto show da: "\d a : P\<^sub>0 a \ a\" using a P\<^sub>0a arr_char src_def trg_def P\<^sub>0_props(1) equivalence_data_in_hom\<^sub>B(2) by auto show "\d a : d a \ d a\" using a da chosen_adjoint_equivalence(3) ide_char in_hom_char by auto show \a: "\\ a : a \ d a \ e a\" proof show 1: "arr (\ a)" using assms a P\<^sub>0a da ea arr_char equivalence_data_in_hom\<^sub>B(7) hcomp_closed by (metis (no_types, lifting) equivalence_data_in_hom\<^sub>B(5) B.in_hhom_def) show "dom (\ a) = a" using a 1 dom_char equivalence_data_in_hom\<^sub>B(7) by auto show "cod (\ a) = d a \ e a" using a 1 da ea cod_char in_hom_char equivalence_data_in_hom\<^sub>B(7) hcomp_def by (metis (no_types, lifting) hcomp_char in_hhomE B.in_homE) qed show \a: "\\ a : e a \ d a \ P\<^sub>0 a\" proof show 1: "arr (\ a)" using assms a P\<^sub>0a da ea arr_char equivalence_data_in_hom\<^sub>B(8) hcomp_closed by (metis (no_types, lifting) equivalence_data_in_hom\<^sub>B(6) B.in_hhom_def) show "dom (\ a) = e a \ d a" using a 1 da ea dom_char in_hom_char equivalence_data_in_hom\<^sub>B(8) hcomp_def by (metis (no_types, lifting) hcomp_char in_hhomE B.in_homE) show "cod (\ a) = P\<^sub>0 a" using a 1 cod_char equivalence_data_in_hom\<^sub>B(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_char obj_simps(2-3) vconn_implies_hpar(1-2)) show "\\ a : P\<^sub>0 a \ P\<^sub>0 a\" using P\<^sub>0a \a src_def trg_def src_cod trg_cod by (metis (no_types, lifting) in_hhom_def in_hom_char 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) = P\<^sub>0 a" and "src (d a) = P\<^sub>0 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) = P\<^sub>0 a" and "trg (\ a) = P\<^sub>0 a" and "dom (\ a) = a" and "cod (\ a) = d a \\<^sub>B e a" and "dom (\ a) = e a \\<^sub>B d a" and "cod (\ a) = P\<^sub>0 a" (* These can be done in a one-liner, but it takes 20 sec. *) using assms equivalence_data_in_hom(2) ide_char obj_char apply auto[1] using assms equivalence_data_in_hom(1) ide_char obj_char apply auto[1] using assms obj_char arr_char iso_char B.iso_is_arr apply auto[1] using assms obj_char arr_char iso_char P\<^sub>0_props B.iso_is_arr apply auto[1] using assms obj_char arr_char P\<^sub>0_props src_def apply auto[1] using assms obj_char arr_char P\<^sub>0_props trg_def apply auto[1] using assms obj_char arr_char P\<^sub>0_props src_def apply auto[1] using assms obj_char arr_char P\<^sub>0_props trg_def apply auto[1] using assms obj_char arr_char dom_char P\<^sub>0_props apply auto[1] using assms obj_char arr_char cod_char P\<^sub>0_props apply auto[1] using assms obj_char arr_char dom_char P\<^sub>0_props apply auto[1] using assms obj_char arr_char cod_char P\<^sub>0_props apply auto[1] using assms obj_char arr_char B.iso_is_arr src_def apply auto[1] using assms obj_char arr_char B.iso_is_arr trg_def apply auto[1] using assms obj_char arr_char dom_char P\<^sub>0_props B.iso_is_arr src_def apply auto[1] using assms obj_char arr_char cod_char P\<^sub>0_props B.iso_is_arr trg_def apply auto[1] using assms obj_char arr_char dom_char B.iso_is_arr apply auto[1] using assms obj_char arr_char cod_char B.iso_is_arr apply auto[1] using assms obj_char arr_char dom_char P\<^sub>0_props B.iso_is_arr apply auto[1] using assms obj_char arr_char cod_char P\<^sub>0_props B.iso_is_arr by auto[1] definition P where "P \ = e (trg\<^sub>B \) \\<^sub>B \ \\<^sub>B d (src\<^sub>B \)" lemma P_in_hom\<^sub>B [intro]: assumes "B.arr \" shows "\P \ : P\<^sub>0 (src\<^sub>B \) \\<^sub>B P\<^sub>0 (trg\<^sub>B \)\" and "\P \ : P (B.dom \) \\<^sub>B P (B.cod \)\" unfolding P_def using assms by auto lemma P_simps\<^sub>B [simp]: assumes "B.arr \" shows "B.arr (P \)" and "src\<^sub>B (P \) = P\<^sub>0 (src\<^sub>B \)" and "trg\<^sub>B (P \) = P\<^sub>0 (trg\<^sub>B \)" and "B.dom (P \) = P (B.dom \)" and "B.cod (P \) = P (B.cod \)" using assms P_in_hom\<^sub>B by blast+ lemma P_in_hom [intro]: assumes "B.arr \" shows "\P \ : P\<^sub>0 (src\<^sub>B \) \ P\<^sub>0 (trg\<^sub>B \)\" and "\P \ : P (B.dom \) \ P (B.cod \)\" proof - show 1: "\P \ : P\<^sub>0 (src\<^sub>B \) \ P\<^sub>0 (trg\<^sub>B \)\" using assms P_in_hom\<^sub>B(1) arr_char src_def trg_def P\<^sub>0_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_char cod_char by (intro in_homI) auto qed lemma P_simps [simp]: assumes "B.arr \" shows "arr (P \)" and "src (P \) = P\<^sub>0 (src\<^sub>B \)" and "trg (P \) = P\<^sub>0 (trg\<^sub>B \)" and "dom (P \) = P (B.dom \)" and "cod (P \) = P (B.cod \)" using assms P_in_hom by blast+ interpretation P: "functor" V\<^sub>B 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_simps\<^sub>B(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 (\ \\<^sub>B \) = P \ \ P \" proof - fix \ \ assume seq: "B.seq \ \" show "P (\ \\<^sub>B \) = P \ \ P \" proof - have 4: "P (\ \\<^sub>B \) = e (trg\<^sub>B \) \\<^sub>B \ \\<^sub>B \ \\<^sub>B d (src\<^sub>B \)" unfolding P_def using seq B.src_vcomp B.trg_vcomp by simp also have "... = P \ \\<^sub>B P \" unfolding P_def using seq 0 4 B.whisker_left B.whisker_right by (metis equivalence_data_simps\<^sub>B(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_char comp_char by auto finally show ?thesis by blast qed qed qed interpretation faithful_functor V\<^sub>B comp P proof fix \ \' assume par: "B.par \ \'" have 1: "src\<^sub>B \ = src\<^sub>B \' \ trg\<^sub>B \ = trg\<^sub>B \'" using par by (metis B.src_dom B.trg_dom) assume eq: "P \ = P \'" interpret src: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e (src\<^sub>B \)\ \d (src\<^sub>B \)\ \\ (src\<^sub>B \)\ \\ (src\<^sub>B \)\ using par chosen_adjoint_equivalence by simp interpret trg: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e (trg\<^sub>B \)\ \d (trg\<^sub>B \)\ \\ (trg\<^sub>B \)\ \\ (trg\<^sub>B \)\ using par chosen_adjoint_equivalence by simp show "\ = \'" using eq par 1 unfolding P_def using B.equivalence_cancel_left [of "e (trg\<^sub>B \)" "\ \\<^sub>B d (src\<^sub>B \)" "\' \\<^sub>B d (src\<^sub>B \')"] B.equivalence_cancel_right [of "d (src\<^sub>B \)" \ \'] B.equivalence_map_def src.dual_equivalence trg.equivalence_in_bicategory_axioms by auto qed interpretation P: weak_arrow_of_homs V\<^sub>B src\<^sub>B trg\<^sub>B comp src trg P proof fix \ assume \: "B.arr \" interpret src: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e (src\<^sub>B \)\ \d (src\<^sub>B \)\ \\ (src\<^sub>B \)\ \\ (src\<^sub>B \)\ using \ chosen_adjoint_equivalence by simp show "isomorphic (P (src\<^sub>B \)) (src (P \))" proof (unfold isomorphic_def) show "\f. \f : P (src\<^sub>B \) \ src (P \)\ \ iso f" proof - let ?\ = "\ (src\<^sub>B \) \\<^sub>B (e (src\<^sub>B \) \\<^sub>B \\<^sub>B[d (src\<^sub>B \)])" have "\?\ : P (src\<^sub>B \) \ src (P \)\ \ iso ?\" proof - have 1: "\?\ : P (src\<^sub>B \) \\<^sub>B P\<^sub>0 (src\<^sub>B \)\ \ B.iso ?\" unfolding P_def using \ by auto moreover have "arr ?\ \ arr (B.inv ?\)" using \ 1 arr_char P\<^sub>0_props(1) by (metis (no_types, lifting) P\<^sub>0_props(6) B.arrI B.arr_inv equivalence_data_simps\<^sub>B(16) B.obj_src src.counit_simps(4-5) B.src_inv B.src_vcomp B.trg_inv B.trg_vcomp) moreover have "P\<^sub>0 (src\<^sub>B \) = src (P \)" using \ by simp ultimately show ?thesis using \ in_hom_char arr_char iso_char by (metis (no_types, lifting) src_closed equivalence_data_simps\<^sub>B(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 V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e (trg\<^sub>B \)\ \d (trg\<^sub>B \)\ \\ (trg\<^sub>B \)\ \\ (trg\<^sub>B \)\ using \ chosen_adjoint_equivalence by simp show "isomorphic (P (trg\<^sub>B \)) (trg (P \))" proof (unfold isomorphic_def) show "\f. \f : P (trg\<^sub>B \) \ trg (P \)\ \ iso f" proof - let ?\ = "\ (trg\<^sub>B \) \\<^sub>B (e (trg\<^sub>B \) \\<^sub>B \\<^sub>B[d (trg\<^sub>B \)])" have "\?\ : P (trg\<^sub>B \) \ trg (P \)\ \ iso ?\" proof - have 1: "\?\ : P (trg\<^sub>B \) \\<^sub>B P\<^sub>0 (trg\<^sub>B \)\ \ B.iso ?\" unfolding P_def using \ by auto moreover have "arr ?\ \ arr (B.inv ?\)" using \ 1 arr_char P\<^sub>0_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 (trg\<^sub>B \)) \ arr (P\<^sub>0 (trg\<^sub>B \))" using \ arr_char P\<^sub>0_props(1) P_simps(1) obj_char by auto moreover have "P\<^sub>0 (trg\<^sub>B \) = trg (P \)" using \ by simp ultimately show ?thesis using in_hom_char iso_char 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.map\<^sub>0 a\ and to \P\<^sub>0 a\. \ lemma P_map\<^sub>0_simp [simp]: assumes "B.obj a" shows "P.map\<^sub>0 a = P\<^sub>0 a" using assms P.map\<^sub>0_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 V\<^sub>B comp \(\\\. fst \\ \\<^sub>B snd \\)\ P .. no_notation B.in_hom ("\_ : _ \\<^sub>B _\") definition CMP where "CMP f g \ (e (trg\<^sub>B f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g, d (src\<^sub>B g)]) \\<^sub>B (e (trg\<^sub>B f) \\<^sub>B f \\<^sub>B \\<^sub>B[g \\<^sub>B d (src\<^sub>B g)] \\<^sub>B (B.inv (\ (trg\<^sub>B g)) \\<^sub>B g \\<^sub>B d (src\<^sub>B g))) \\<^sub>B (e (trg\<^sub>B f) \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d (src\<^sub>B f), e (trg\<^sub>B g), g \\<^sub>B d (src\<^sub>B g)]) \\<^sub>B \\<^sub>B[e (trg\<^sub>B f), f, d (src\<^sub>B f) \\<^sub>B P g] \\<^sub>B \\<^sub>B[e (trg\<^sub>B f) \\<^sub>B f, d (src\<^sub>B f), P g] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e (trg\<^sub>B f), f, d (src\<^sub>B f)] \\<^sub>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 "src\<^sub>B f = trg\<^sub>B g" shows "\CMP f g : P\<^sub>0 (src\<^sub>B g) \ P\<^sub>0 (trg\<^sub>B f)\" and "\CMP f g : P f \ P g \ P (f \\<^sub>B g)\" and "\CMP f g : P\<^sub>0 (src\<^sub>B g) \\<^sub>B P\<^sub>0 (trg\<^sub>B f)\" and "\CMP f g : P f \\<^sub>B P g \\<^sub>B P (f \\<^sub>B g)\" proof - show 1: "\CMP f g : P f \\<^sub>B P g \\<^sub>B P (f \\<^sub>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 : P\<^sub>0 (src\<^sub>B g) \\<^sub>B P\<^sub>0 (trg\<^sub>B 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 \\<^sub>B g)\" using assms 1 B.arrI B.vconn_implies_hpar(1-2) P_simps(1) arr_char hcomp_eqI hseqI' in_hom_char by force show "\CMP f g : P\<^sub>0 (src\<^sub>B g) \ P\<^sub>0 (trg\<^sub>B f)\" using 2 3 arr_char src_def trg_def by fastforce qed lemma CMP_simps [simp]: assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g" shows "arr (CMP f g)" and "src (CMP f g) = P\<^sub>0 (src\<^sub>B g)" and "trg (CMP f g) = P\<^sub>0 (trg\<^sub>B f)" and "dom (CMP f g) = P f \ P g" and "cod (CMP f g) = P (f \\<^sub>B g)" and "B.arr (CMP f g)" and "src\<^sub>B (CMP f g) = P\<^sub>0 (src\<^sub>B g)" and "trg\<^sub>B (CMP f g) = P\<^sub>0 (trg\<^sub>B f)" and "B.dom (CMP f g) = P f \\<^sub>B P g" and "B.cod (CMP f g) = P (f \\<^sub>B g)" using assms CMP_in_hom [of f g] by auto lemma iso_CMP: assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B 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_char P.as_nat_iso.components_are_iso by (intro B.isos_compose B.iso_hcomp) auto thus "iso (CMP f g)" using assms iso_char arr_char CMP_simps(1) by auto qed abbreviation (input) SRC where "SRC \ \ d (src\<^sub>B \) \\<^sub>B e (src\<^sub>B \)" abbreviation (input) TRG where "TRG \ \ d (trg\<^sub>B \) \\<^sub>B e (trg\<^sub>B \)" definition LUNIT where "LUNIT f \ \\<^sub>B[f] \\<^sub>B (B.inv (\ (trg\<^sub>B f)) \\<^sub>B f)" definition RUNIT where "RUNIT f \ \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ (src\<^sub>B 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 : src\<^sub>B f \\<^sub>B trg\<^sub>B f\" and "\LUNIT f : TRG f \\<^sub>B f \\<^sub>B f\" proof - interpret e_trg_f: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e (trg\<^sub>B f)\ \d (trg\<^sub>B f)\ \\ (trg\<^sub>B f)\ \\ (trg\<^sub>B f)\ using assms chosen_adjoint_equivalence by simp show "\LUNIT f : src\<^sub>B f \\<^sub>B trg\<^sub>B f\" unfolding LUNIT_def using assms e_trg_f.unit_is_iso by auto show "\LUNIT f : TRG f \\<^sub>B f \\<^sub>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 "src\<^sub>B (LUNIT f) = src\<^sub>B f" and "trg\<^sub>B (LUNIT f) = trg\<^sub>B f" and "B.dom (LUNIT f) = TRG f \\<^sub>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 : src\<^sub>B f \\<^sub>B trg\<^sub>B f\" and "\RUNIT f : f \\<^sub>B SRC f \\<^sub>B f\" proof - interpret e_src_f: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e (src\<^sub>B f)\ \d (src\<^sub>B f)\ \\ (src\<^sub>B f)\ \\ (src\<^sub>B f)\ using assms chosen_adjoint_equivalence by simp show "\RUNIT f : src\<^sub>B f \\<^sub>B trg\<^sub>B f\" unfolding RUNIT_def using assms e_src_f.unit_is_iso by auto show "\RUNIT f : f \\<^sub>B SRC f \\<^sub>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 "src\<^sub>B (RUNIT f) = src\<^sub>B f" and "trg\<^sub>B (RUNIT f) = trg\<^sub>B f" and "B.dom (RUNIT f) = f \\<^sub>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 "\ \\<^sub>B LUNIT (B.dom \) = LUNIT (B.cod \) \\<^sub>B (TRG \ \\<^sub>B \)" proof - interpret e_trg_\: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e (trg\<^sub>B \)\ \d (trg\<^sub>B \)\ \\ (trg\<^sub>B \)\ \\ (trg\<^sub>B \)\ using assms chosen_adjoint_equivalence by simp show ?thesis proof - have "\ \\<^sub>B LUNIT (B.dom \) = (\ \\<^sub>B \\<^sub>B[B.dom \]) \\<^sub>B (B.inv (\ (trg\<^sub>B \)) \\<^sub>B B.dom \)" unfolding LUNIT_def using assms B.comp_assoc by simp also have "... = \\<^sub>B[B.cod \] \\<^sub>B (trg\<^sub>B \ \\<^sub>B \) \\<^sub>B (B.inv (\ (trg\<^sub>B \)) \\<^sub>B B.dom \)" using assms B.lunit_naturality B.comp_assoc by simp also have "... = \\<^sub>B[B.cod \] \\<^sub>B (B.inv (\ (trg\<^sub>B \)) \\<^sub>B \)" using assms B.interchange [of "trg\<^sub>B \" "B.inv (\ (trg\<^sub>B \))" \ "B.dom \"] e_trg_\.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp also have "... = \\<^sub>B[B.cod \] \\<^sub>B (B.inv (\ (trg\<^sub>B \)) \\<^sub>B B.cod \) \\<^sub>B (TRG \ \\<^sub>B \)" using assms B.interchange [of "B.inv (\ (trg\<^sub>B \))" "TRG \" "B.cod \" \] e_trg_\.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp also have "... = LUNIT (B.cod \) \\<^sub>B (TRG \ \\<^sub>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 "\ \\<^sub>B RUNIT (B.dom \) = RUNIT (B.cod \) \\<^sub>B (\ \\<^sub>B SRC \)" proof - interpret e_src_\: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e (src\<^sub>B \)\ \d (src\<^sub>B \)\ \\ (src\<^sub>B \)\ \\ (src\<^sub>B \)\ using assms chosen_adjoint_equivalence by simp show ?thesis proof - have "\ \\<^sub>B RUNIT (B.dom \) = (\ \\<^sub>B \\<^sub>B[B.dom \]) \\<^sub>B (B.dom \ \\<^sub>B B.inv (\ (src\<^sub>B \)))" unfolding RUNIT_def using assms B.comp_assoc by simp also have "... = \\<^sub>B[B.cod \] \\<^sub>B (\ \\<^sub>B src\<^sub>B \) \\<^sub>B (B.dom \ \\<^sub>B B.inv (\ (src\<^sub>B \)))" using assms B.runit_naturality B.comp_assoc by simp also have "... = \\<^sub>B[B.cod \] \\<^sub>B (\ \\<^sub>B B.inv (\ (src\<^sub>B \)))" using assms B.interchange [of \ "B.dom \" "src\<^sub>B \" "B.inv (\ (src\<^sub>B \))"] e_src_\.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp also have "... = \\<^sub>B[B.cod \] \\<^sub>B (B.cod \ \\<^sub>B B.inv (\ (src\<^sub>B \))) \\<^sub>B (\ \\<^sub>B SRC \)" using assms B.interchange [of "B.cod \" \ "B.inv (\ (src\<^sub>B \))" "SRC \"] e_src_\.unit_is_iso B.comp_arr_dom B.comp_cod_arr by simp also have "... = RUNIT (B.cod \) \\<^sub>B (\ \\<^sub>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 "src\<^sub>B f = trg\<^sub>B g" shows "LUNIT (f \\<^sub>B g) \\<^sub>B \\<^sub>B[d (trg\<^sub>B f) \\<^sub>B e (trg\<^sub>B f), f, g] = LUNIT f \\<^sub>B g" proof - interpret e_trg_f: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e (trg\<^sub>B f)\ \d (trg\<^sub>B f)\ \\ (trg\<^sub>B f)\ \\ (trg\<^sub>B f)\ using assms chosen_adjoint_equivalence by simp have "LUNIT (f \\<^sub>B g) \\<^sub>B \\<^sub>B[TRG f, f, g] = \\<^sub>B[f \\<^sub>B g] \\<^sub>B (B.inv (\ (trg\<^sub>B f)) \\<^sub>B f \\<^sub>B g) \\<^sub>B \\<^sub>B[TRG f, f, g]" unfolding LUNIT_def using assms B.comp_assoc by simp also have "... = (\\<^sub>B[f \\<^sub>B g] \\<^sub>B \\<^sub>B[trg\<^sub>B f, f, g]) \\<^sub>B ((B.inv (\ (trg\<^sub>B f)) \\<^sub>B f) \\<^sub>B g)" using assms B.assoc_naturality [of "B.inv (\ (trg\<^sub>B f))" f g] e_trg_f.unit_is_iso B.comp_assoc by simp also have "... = (\\<^sub>B[f] \\<^sub>B g) \\<^sub>B ((B.inv (\ (trg\<^sub>B f)) \\<^sub>B f) \\<^sub>B g)" using assms B.lunit_hcomp by simp also have "... = LUNIT f \\<^sub>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 "src\<^sub>B f = trg\<^sub>B g" shows "RUNIT (f \\<^sub>B g) = (f \\<^sub>B RUNIT g) \\<^sub>B \\<^sub>B[f, g, SRC g]" proof - interpret e_src_g: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e (src\<^sub>B g)\ \d (src\<^sub>B g)\ \\ (src\<^sub>B g)\ \\ (src\<^sub>B g)\ using assms chosen_adjoint_equivalence by simp have "(f \\<^sub>B RUNIT g) \\<^sub>B \\<^sub>B[f, g, SRC g] = (f \\<^sub>B \\<^sub>B[g]) \\<^sub>B (f \\<^sub>B g \\<^sub>B B.inv (\ (src\<^sub>B g))) \\<^sub>B \\<^sub>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 \\<^sub>B \\<^sub>B[g]) \\<^sub>B \\<^sub>B[f, g, src\<^sub>B g]) \\<^sub>B ((f \\<^sub>B g) \\<^sub>B B.inv (\ (src\<^sub>B g)))" using assms B.assoc_naturality [of f g "B.inv (\ (src\<^sub>B g))"] e_src_g.unit_is_iso B.comp_assoc by simp also have "... = \\<^sub>B[f \\<^sub>B g] \\<^sub>B ((f \\<^sub>B g) \\<^sub>B B.inv (\ (src\<^sub>B g)))" using assms B.runit_hcomp by simp also have "... = RUNIT (f \\<^sub>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 "src\<^sub>B f = trg\<^sub>B g" shows "(f \\<^sub>B LUNIT g) \\<^sub>B \\<^sub>B[f, SRC f, g] = RUNIT f \\<^sub>B g" proof - interpret e_trg_g: adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e (trg\<^sub>B g)\ \d (trg\<^sub>B g)\ \\ (trg\<^sub>B g)\ \\ (trg\<^sub>B g)\ using assms chosen_adjoint_equivalence by simp show ?thesis proof - have "(f \\<^sub>B LUNIT g) \\<^sub>B \\<^sub>B[f, SRC f, g] = (f \\<^sub>B \\<^sub>B[g]) \\<^sub>B (f \\<^sub>B B.inv (\ (trg\<^sub>B g)) \\<^sub>B g) \\<^sub>B \\<^sub>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 \\<^sub>B \\<^sub>B[g]) \\<^sub>B \\<^sub>B[f, src\<^sub>B f, g]) \\<^sub>B ((f \\<^sub>B B.inv (\ (trg\<^sub>B g))) \\<^sub>B g)" using assms B.assoc_naturality [of f "B.inv (\ (trg\<^sub>B g))" g] e_trg_g.unit_is_iso B.comp_assoc by auto also have "... = (\\<^sub>B[f] \\<^sub>B g) \\<^sub>B ((f \\<^sub>B B.inv (\ (trg\<^sub>B g))) \\<^sub>B g)" using assms B.triangle by simp also have "... = RUNIT f \\<^sub>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 "src\<^sub>B \ = trg\<^sub>B \" shows "CMP (B.cod \) (B.cod \) \\<^sub>B (P \ \\<^sub>B P \) = P (\ \\<^sub>B \) \\<^sub>B CMP (B.dom \) (B.dom \)" proof - let ?a = "src\<^sub>B \" let ?b = "src\<^sub>B \" let ?c = "trg\<^sub>B \" let ?f = "B.dom \" let ?g = "B.cod \" let ?h = "B.dom \" let ?k = "B.cod \" have "CMP ?g ?k \\<^sub>B (P \ \\<^sub>B P \) = (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \\<^sub>B (e ?c \\<^sub>B ?g \\<^sub>B LUNIT (?k \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B ?g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, ?g, d ?b \\<^sub>B P ?k] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B ?g, d ?b, P ?k] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, ?g, d ?b] \\<^sub>B P ?k) \\<^sub>B (P \ \\<^sub>B P \)" unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp also have "... = (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \\<^sub>B (e ?c \\<^sub>B ?g \\<^sub>B LUNIT (?k \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B ?g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, ?g, d ?b \\<^sub>B P ?k] \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B ?g, d ?b, P ?k] \\<^sub>B (((e ?c \\<^sub>B \) \\<^sub>B d ?b) \\<^sub>B P \)) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \\<^sub>B P ?h)" proof - have "(\\<^sub>B\<^sup>-\<^sup>1[e ?c, ?g, d ?b] \\<^sub>B P ?k) \\<^sub>B (P \ \\<^sub>B P \) = \\<^sub>B\<^sup>-\<^sup>1[e ?c, ?g, d ?b] \\<^sub>B P \ \\<^sub>B P ?k \\<^sub>B P \" using assms P_def B.interchange by fastforce also have "... = ((e ?c \\<^sub>B \) \\<^sub>B d ?b) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \\<^sub>B P ?k \\<^sub>B P \" using assms P_def B.assoc'_naturality [of "e ?c" \ "d ?b"] by simp also have "... = ((e ?c \\<^sub>B \) \\<^sub>B d ?b) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \\<^sub>B P \ \\<^sub>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 \\<^sub>B \) \\<^sub>B d ?b) \\<^sub>B P \) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \\<^sub>B P ?h)" using assms B.interchange by auto finally have "(\\<^sub>B\<^sup>-\<^sup>1[e ?c, ?g, d ?b] \\<^sub>B P ?k) \\<^sub>B (P \ \\<^sub>B P \) = (((e ?c \\<^sub>B \) \\<^sub>B d ?b) \\<^sub>B P \) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \\<^sub>B P ?h)" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \\<^sub>B (e ?c \\<^sub>B ?g \\<^sub>B LUNIT (?k \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B ?g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \\<^sub>B d ?a]) \\<^sub>B (\\<^sub>B[e ?c, ?g, d ?b \\<^sub>B P ?k] \\<^sub>B ((e ?c \\<^sub>B \) \\<^sub>B d ?b \\<^sub>B P \)) \\<^sub>B \\<^sub>B[e ?c \\<^sub>B ?f, d ?b, P ?h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \\<^sub>B P ?h)" using assms B.assoc_naturality [of "e ?c \\<^sub>B \" "d ?b" "P \"] B.comp_assoc by simp also have "... = (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \\<^sub>B (e ?c \\<^sub>B ?g \\<^sub>B LUNIT (?k \\<^sub>B d ?a)) \\<^sub>B ((e ?c \\<^sub>B ?g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \\<^sub>B d ?a]) \\<^sub>B (e ?c \\<^sub>B \ \\<^sub>B d ?b \\<^sub>B P \)) \\<^sub>B \\<^sub>B[e ?c, ?f, d ?b \\<^sub>B P ?h] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B ?f, d ?b, P ?h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \\<^sub>B P ?h)" using assms B.assoc_naturality [of "e ?c" \ "d ?b \\<^sub>B P \"] B.comp_assoc by simp also have "... = (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \\<^sub>B ((e ?c \\<^sub>B ?g \\<^sub>B LUNIT (?k \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B \ \\<^sub>B SRC \ \\<^sub>B \ \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B ?f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, ?f, d ?b \\<^sub>B P ?h] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B ?f, d ?b, P ?h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \\<^sub>B P ?h)" proof - have "(e ?c \\<^sub>B ?g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \\<^sub>B d ?a]) \\<^sub>B (e ?c \\<^sub>B \ \\<^sub>B d ?b \\<^sub>B P \) = (e ?c \\<^sub>B \ \\<^sub>B TRG \ \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B (e ?c \\<^sub>B ?f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \\<^sub>B d ?a])" proof - have "(e ?c \\<^sub>B ?g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \\<^sub>B d ?a]) \\<^sub>B (e ?c \\<^sub>B \ \\<^sub>B d ?b \\<^sub>B P \) = e ?c \\<^sub>B (?g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \\<^sub>B d ?a]) \\<^sub>B (\ \\<^sub>B d ?b \\<^sub>B P \)" using assms P_def B.whisker_left by simp also have "... = e ?c \\<^sub>B \ \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \\<^sub>B d ?a] \\<^sub>B (d ?b \\<^sub>B P \)" using assms P_def B.comp_cod_arr B.interchange [of "?g" \ "\\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?k \\<^sub>B d ?a]" "d ?b \\<^sub>B P \"] by simp also have "... = e ?c \\<^sub>B \ \\<^sub>B (TRG \ \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \\<^sub>B d ?a]" using assms P_def B.assoc'_naturality [of "d ?b" "e ?b" "\ \\<^sub>B d ?a"] by simp also have "... = e ?c \\<^sub>B (\ \\<^sub>B TRG \ \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B (?f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \\<^sub>B d ?a])" using assms B.comp_arr_dom B.interchange [of \ "?f" "TRG \ \\<^sub>B \ \\<^sub>B d ?a""\\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \\<^sub>B d ?a]"] by fastforce also have "... = (e ?c \\<^sub>B \ \\<^sub>B TRG \ \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B (e ?c \\<^sub>B ?f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \\<^sub>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 \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \\<^sub>B (e ?c \\<^sub>B \ \\<^sub>B \ \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B ?f \\<^sub>B LUNIT (?h \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B ?f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, ?f, d ?b \\<^sub>B P ?h] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B ?f, d ?b, P ?h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \\<^sub>B P ?h)" proof - have "((e ?c \\<^sub>B ?g \\<^sub>B LUNIT (?k \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B \ \\<^sub>B TRG \ \\<^sub>B \ \\<^sub>B d ?a)) = e ?c \\<^sub>B \ \\<^sub>B LUNIT (B.cod (\ \\<^sub>B d ?a)) \\<^sub>B ((d ?b \\<^sub>B e ?b) \\<^sub>B \ \\<^sub>B d ?a)" using assms comp_arr_dom B.comp_cod_arr B.whisker_left B.interchange [of "?g" \ "LUNIT (?k \\<^sub>B d ?a)" "(d ?b \\<^sub>B e ?b) \\<^sub>B \ \\<^sub>B d ?a"] by fastforce also have "... = e ?c \\<^sub>B \ \\<^sub>B (\ \\<^sub>B d ?a) \\<^sub>B LUNIT (?h \\<^sub>B d ?a)" using assms LUNIT_naturality [of "\ \\<^sub>B d ?a"] by simp also have "... = (e ?c \\<^sub>B \ \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B (e ?c \\<^sub>B ?f \\<^sub>B LUNIT (?h \\<^sub>B d ?a))" using assms B.comp_arr_dom B.comp_cod_arr B.whisker_left B.interchange [of \ "?f" "\ \\<^sub>B d ?a" "LUNIT (?h \\<^sub>B d ?a)"] by simp finally have "((e ?c \\<^sub>B ?g \\<^sub>B LUNIT (?k \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B \ \\<^sub>B TRG \ \\<^sub>B \ \\<^sub>B d ?a)) = (e ?c \\<^sub>B \ \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B (e ?c \\<^sub>B ?f \\<^sub>B LUNIT (?h \\<^sub>B d ?a))" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = P (\ \\<^sub>B \) \\<^sub>B (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?f, ?h, d ?a]) \\<^sub>B (e ?c \\<^sub>B ?f \\<^sub>B LUNIT (?h \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B ?f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ?h \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, ?f, d ?b \\<^sub>B P ?h] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B ?f, d ?b, P ?h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, ?f, d ?b] \\<^sub>B P ?h)" proof - have "(e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \\<^sub>B (e ?c \\<^sub>B \ \\<^sub>B \ \\<^sub>B d ?a) = P (\ \\<^sub>B \) \\<^sub>B (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?f, ?h, d ?a])" proof - have "(e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a]) \\<^sub>B (e ?c \\<^sub>B \ \\<^sub>B \ \\<^sub>B d ?a) = e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g, ?k, d ?a] \\<^sub>B (\ \\<^sub>B \ \\<^sub>B d ?a)" using assms B.whisker_left by simp also have "... = e ?c \\<^sub>B ((\ \\<^sub>B \) \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?f, ?h, d ?a]" using assms B.assoc'_naturality [of \ \ "d ?a"] by simp also have "... = P (\ \\<^sub>B \) \\<^sub>B (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>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 (\ \\<^sub>B \) \\<^sub>B CMP ?f ?h" unfolding CMP_def LUNIT_def using assms by simp finally show ?thesis by simp qed interpretation EV: self_evaluation_map V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B .. notation EV.eval ("\_\") abbreviation (input) SRCt ("\<^bold>S\<^bold>R\<^bold>C") where "\<^bold>S\<^bold>R\<^bold>C \ \ \<^bold>\d (src\<^sub>B \)\<^bold>\ \<^bold>\ \<^bold>\e (src\<^sub>B \)\<^bold>\" abbreviation (input) TRGt ("\<^bold>T\<^bold>R\<^bold>G") where "\<^bold>T\<^bold>R\<^bold>G \ \ \<^bold>\d (trg\<^sub>B \)\<^bold>\ \<^bold>\ \<^bold>\e (trg\<^sub>B \)\<^bold>\" abbreviation (input) PRJt ("\<^bold>P\<^bold>R\<^bold>J") where "\<^bold>P\<^bold>R\<^bold>J \ \ \<^bold>\e (trg\<^sub>B \)\<^bold>\ \<^bold>\ \<^bold>\\\<^bold>\ \<^bold>\ \<^bold>\d (src\<^sub>B \)\<^bold>\" 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 "src\<^sub>B f = trg\<^sub>B g" and "src\<^sub>B g = trg\<^sub>B h" shows "P \\<^sub>B[f, g, h] \\<^sub>B CMP (f \\<^sub>B g) h \\<^sub>B (CMP f g \\<^sub>B P h) = CMP f (g \\<^sub>B h) \\<^sub>B (P f \\<^sub>B CMP g h) \\<^sub>B \\<^sub>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 = "trg\<^sub>B f" let ?b = "trg\<^sub>B g" let ?c = "trg\<^sub>B h" let ?d = "src\<^sub>B h" have "P \\<^sub>B[f, g, h] \\<^sub>B CMP (f \\<^sub>B g) h \\<^sub>B (CMP f g \\<^sub>B P h) = P \\<^sub>B[f, g, h] \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B LUNIT (h \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B d ?c)) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp also have "... = P \\<^sub>B[f, g, h] \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B LUNIT (h \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B d ?c)) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" using assms B.whisker_right P_def by simp (* 6 sec *) also have "... = P \\<^sub>B[f, g, h] \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B LUNIT h \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B d ?c)) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B LUNIT (h \\<^sub>B d ?d) = e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B (LUNIT h \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>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 \\<^sub>B (f \\<^sub>B g) \\<^sub>B LUNIT h \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])" using assms B.whisker_left by simp finally have "e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B LUNIT (h \\<^sub>B d ?d) = (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B LUNIT h \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (P \\<^sub>B[f, g, h] \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B d ?c)) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "(e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B LUNIT h \\<^sub>B d ?d) = e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, h, d ?d] \\<^sub>B ((f \\<^sub>B g) \\<^sub>B LUNIT h \\<^sub>B d ?d)" using assms B.whisker_left by simp also have "... = e ?a \\<^sub>B (((f \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]" using assms B.assoc'_naturality [of "f \\<^sub>B g" "LUNIT h" "d ?d"] by simp also have "... = (e ?a \\<^sub>B ((f \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d])" using assms B.whisker_left by simp finally have "(e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B LUNIT h \\<^sub>B d ?d) = (e ?a \\<^sub>B ((f \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d])" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B d ?c)) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "P \\<^sub>B[f, g, h] \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B d ?d) = e ?a \\<^sub>B \\<^sub>B[f, g, h] \\<^sub>B ((f \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B d ?d" using assms B.whisker_left B.whisker_right P_def by simp also have "... = e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d" using assms B.assoc_naturality [of f g "LUNIT h"] by simp also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d)" using assms B.whisker_left B.whisker_right by simp finally have "P \\<^sub>B[f, g, h] \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B d ?d) = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d)" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B (((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B RUNIT f \\<^sub>B g \\<^sub>B d ?c) \\<^sub>B P h)) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "((e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B d ?c)) \\<^sub>B P h) = (e ?a \\<^sub>B (RUNIT f \\<^sub>B g \\<^sub>B d ?c) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \\<^sub>B d ?c]) \\<^sub>B P h" using assms TRIANGLE [of f "g \\<^sub>B d ?c"] B.invert_side_of_triangle by simp also have "... = ((e ?a \\<^sub>B RUNIT f \\<^sub>B g \\<^sub>B d ?c) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \\<^sub>B d ?c]) \\<^sub>B P h)" using assms B.whisker_left B.whisker_right P_def by simp finally have "((e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B d ?c)) \\<^sub>B P h) = ((e ?a \\<^sub>B RUNIT f \\<^sub>B g \\<^sub>B d ?c) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \\<^sub>B d ?c]) \\<^sub>B P h)" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f \\<^sub>B g, d ?c, P h] \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B d ?c) \\<^sub>B P h)) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B TRG g, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B RUNIT f \\<^sub>B g \\<^sub>B d ?c) \\<^sub>B P h) = (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c] \\<^sub>B (RUNIT f \\<^sub>B g \\<^sub>B d ?c)) \\<^sub>B P h" using assms B.whisker_left B.whisker_right P_def by simp also have "... = (e ?a \\<^sub>B ((RUNIT f \\<^sub>B g) \\<^sub>B d ?c) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B TRG g, g, d ?c]) \\<^sub>B P h" using assms B.assoc'_naturality [of "RUNIT f" g "d ?c"] by simp also have "... = ((e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B d ?c) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B TRG g, g, d ?c]) \\<^sub>B P h)" using assms B.whisker_left B.whisker_right P_def by simp finally have "((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B RUNIT f \\<^sub>B g \\<^sub>B d ?c) \\<^sub>B P h) = ((e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B d ?c) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B TRG g, g, d ?c]) \\<^sub>B P h)" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f \\<^sub>B g, d ?c, P h] \\<^sub>B (((e ?a \\<^sub>B (RUNIT f \\<^sub>B g)) \\<^sub>B d ?c) \\<^sub>B P h)) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "(\\<^sub>B\<^sup>-\<^sup>1[e ?a, f \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B d ?c) \\<^sub>B P h) = \\<^sub>B\<^sup>-\<^sup>1[e ?a, f \\<^sub>B g, d ?c] \\<^sub>B (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B d ?c) \\<^sub>B P h" using assms B.whisker_left B.whisker_right P_def by simp also have "... = ((e ?a \\<^sub>B (RUNIT f \\<^sub>B g)) \\<^sub>B d ?c) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h" using assms B.assoc'_naturality [of "e ?a" "RUNIT f \\<^sub>B g" "d ?c"] by simp also have "... = (((e ?a \\<^sub>B (RUNIT f \\<^sub>B g)) \\<^sub>B d ?c) \\<^sub>B P h) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h)" using assms B.whisker_left B.whisker_right P_def by simp finally have "(\\<^sub>B\<^sup>-\<^sup>1[e ?a, f \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B d ?c) \\<^sub>B P h) = (((e ?a \\<^sub>B (RUNIT f \\<^sub>B g)) \\<^sub>B d ?c) \\<^sub>B P h) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h)" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (\\<^sub>B[e ?a, f \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B ((e ?a \\<^sub>B (RUNIT f \\<^sub>B g)) \\<^sub>B d ?c \\<^sub>B P h)) \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?b \\<^sub>B e ?b, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, d ?b \\<^sub>B e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" using assms B.assoc_naturality [of "e ?a \\<^sub>B (RUNIT f \\<^sub>B g)" "d ?c" "P h"] B.comp_assoc by simp also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B ((e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B d ?c \\<^sub>B P h)) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" using assms B.assoc_naturality [of "e ?a" "RUNIT f \\<^sub>B g" "d ?c \\<^sub>B P h"] B.comp_assoc by simp also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B ((e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d])) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "((e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B d ?c \\<^sub>B P h)) = e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d] \\<^sub>B (d ?c \\<^sub>B P h)" using assms B.comp_cod_arr B.whisker_left P_def B.interchange [of "f \\<^sub>B g" "RUNIT f \\<^sub>B g"] by simp also have "... = e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]" using assms B.comp_arr_dom P_def by simp finally have "((e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B d ?c \\<^sub>B P h)) = e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d])) \\<^sub>B (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "(e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) = e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]" using assms B.comp_arr_dom B.comp_cod_arr B.whisker_left B.interchange [of "f \\<^sub>B g" "RUNIT f \\<^sub>B g" "\\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]" "\\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]"] by simp (* 3 sec *) also have "... = (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d])" using assms B.comp_arr_dom B.whisker_left B.interchange [of "RUNIT f \\<^sub>B g" "(f \\<^sub>B SRC f) \\<^sub>B g" "\\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]" "\\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]"] by simp finally have "(e ?a \\<^sub>B (f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) = (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d])" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B[f, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "(e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) = e ?a \\<^sub>B (\\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]" using assms B.whisker_left by simp also have "... = e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B (d ?c \\<^sub>B e ?c) \\<^sub>B h, d ?d] \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, (d ?c \\<^sub>B e ?c) \\<^sub>B h, d ?d]) \\<^sub>B \\<^sub>B[f, g, ((d ?c \\<^sub>B e ?c) \\<^sub>B h) \\<^sub>B d ?d]" proof - have "(\\<^sub>B\<^sup>-\<^sup>1[f, g, (d ?c \\<^sub>B e ?c) \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B (d ?c \\<^sub>B e ?c) \\<^sub>B h, d ?d] \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, (d ?c \\<^sub>B e ?c) \\<^sub>B h, d ?d]) = \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, (d ?c \\<^sub>B e ?c) \\<^sub>B h, d ?d] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g, ((d ?c \\<^sub>B e ?c) \\<^sub>B h) \\<^sub>B d ?d]" using assms B.pentagon' B.comp_assoc by simp moreover have "B.inv (\\<^sub>B\<^sup>-\<^sup>1[f, g, (d ?c \\<^sub>B e ?c) \\<^sub>B h] \\<^sub>B d ?d) = \\<^sub>B[f, g, (d ?c \\<^sub>B e ?c) \\<^sub>B h] \\<^sub>B d ?d" using assms by simp ultimately show ?thesis using assms B.comp_assoc B.invert_opposite_sides_of_square [of "\\<^sub>B\<^sup>-\<^sup>1[f, g, (d ?c \\<^sub>B e ?c) \\<^sub>B h] \\<^sub>B d ?d" "\\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B (d ?c \\<^sub>B e ?c) \\<^sub>B h, d ?d] \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, (d ?c \\<^sub>B e ?c) \\<^sub>B h, d ?d])" "\\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, (d ?c \\<^sub>B e ?c) \\<^sub>B h, d ?d]" "\\<^sub>B\<^sup>-\<^sup>1[f, g, ((d ?c \\<^sub>B e ?c) \\<^sub>B h) \\<^sub>B d ?d]"] by simp (* 3 sec *) qed also have "... = (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d])" using assms B.whisker_left by simp finally have "(e ?a \\<^sub>B \\<^sub>B[f, g, TRG h \\<^sub>B h] \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) = (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d])" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B RUNIT f \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f \\<^sub>B SRC f, g, (d ?c \\<^sub>B e ?c) \\<^sub>B h \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "(e ?a \\<^sub>B \\<^sub>B[f, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) = e ?a \\<^sub>B \\<^sub>B[f, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d] \\<^sub>B ((RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])" using assms B.whisker_left by simp also have "... = e ?a \\<^sub>B (RUNIT f \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B \\<^sub>B[f \\<^sub>B SRC f, g, TRG h \\<^sub>B h \\<^sub>B d ?d]" using assms B.hseqI' B.assoc_naturality [of "RUNIT f" g "\\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]"] by simp also have "... = (e ?a \\<^sub>B RUNIT f \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f \\<^sub>B SRC f, g, TRG h \\<^sub>B h \\<^sub>B d ?d])" using assms B.whisker_left by simp finally have "(e ?a \\<^sub>B \\<^sub>B[f, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (RUNIT f \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) = (e ?a \\<^sub>B RUNIT f \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f \\<^sub>B SRC f, g, TRG h \\<^sub>B h \\<^sub>B d ?d])" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d))) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, d ?b \\<^sub>B e ?b, g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c \\<^sub>B e ?c, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f \\<^sub>B SRC f, g, TRG h \\<^sub>B h \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "e ?a \\<^sub>B RUNIT f \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c \\<^sub>B e ?c, h, d ?d] = (e ?a \\<^sub>B RUNIT f \\<^sub>B g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>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 \\<^sub>B SRC f" "g \\<^sub>B ((TRG h \\<^sub>B h) \\<^sub>B d ?d)" "g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]"] by simp (* 5 sec *) also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B \\<^sub>B[f, d ?b \\<^sub>B e ?b, g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])" using assms TRIANGLE [of f "g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d"] by simp also have "... = (e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, d ?b \\<^sub>B e ?b, g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])" using assms B.whisker_left B.comp_assoc by simp finally have "e ?a \\<^sub>B RUNIT f \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d] = (e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, d ?b \\<^sub>B e ?b, g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d])" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B LUNIT g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, d ?b \\<^sub>B e ?b, g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f \\<^sub>B SRC f, g, TRG h \\<^sub>B h \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d) = e ?a \\<^sub>B f \\<^sub>B (LUNIT g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b \\<^sub>B e ?b, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d]" using assms LUNIT_hcomp [of g "(TRG h \\<^sub>B h) \\<^sub>B d ?d"] B.invert_side_of_triangle by simp also have "... = (e ?a \\<^sub>B f \\<^sub>B LUNIT g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d])" using assms B.whisker_left by simp finally have "e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d) = (e ?a \\<^sub>B f \\<^sub>B LUNIT g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d])" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B (LUNIT g \\<^sub>B TRG h \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, d ?b \\<^sub>B e ?b, g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f \\<^sub>B SRC f, g, TRG h \\<^sub>B h \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "(e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B LUNIT g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d) = e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h \\<^sub>B h, d ?d] \\<^sub>B (LUNIT g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d)" using assms B.whisker_left by simp also have "... = e ?a \\<^sub>B f \\<^sub>B ((LUNIT g \\<^sub>B TRG h \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g \\<^sub>B g, TRG h \\<^sub>B h, d ?d]" using assms B.assoc'_naturality [of "LUNIT g" "TRG h \\<^sub>B h" "d ?d"] by simp also have "... = (e ?a \\<^sub>B f \\<^sub>B (LUNIT g \\<^sub>B TRG h \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g \\<^sub>B g, TRG h \\<^sub>B h, d ?d])" using assms B.whisker_left by simp finally have "(e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B LUNIT g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d) = (e ?a \\<^sub>B f \\<^sub>B (LUNIT g \\<^sub>B TRG h \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g \\<^sub>B g, TRG h \\<^sub>B h, d ?d])" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = ((e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B SRC g \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \\<^sub>B g) \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, d ?b \\<^sub>B e ?b, g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f \\<^sub>B SRC f, g, TRG h \\<^sub>B h \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "(e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B (LUNIT g \\<^sub>B TRG h \\<^sub>B h) \\<^sub>B d ?d) = e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B TRG h \\<^sub>B h, d ?d] \\<^sub>B (f \\<^sub>B (LUNIT g \\<^sub>B TRG h \\<^sub>B h) \\<^sub>B d ?d)" using assms B.whisker_left by simp also have "... = e ?a \\<^sub>B ((f \\<^sub>B LUNIT g \\<^sub>B TRG h \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \\<^sub>B g) \\<^sub>B TRG h \\<^sub>B h, d ?d]" using assms B.assoc'_naturality [of f "LUNIT g \\<^sub>B TRG h \\<^sub>B h" "d ?d"] by simp also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B TRG h \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \\<^sub>B g) \\<^sub>B TRG h \\<^sub>B h, d ?d])" using assms B.whisker_left by simp finally have "(e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B (LUNIT g \\<^sub>B TRG h \\<^sub>B h) \\<^sub>B d ?d) = (e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B TRG h \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \\<^sub>B g) \\<^sub>B TRG h \\<^sub>B h, d ?d])" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \\<^sub>B g) \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, d ?b \\<^sub>B e ?b, g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f \\<^sub>B SRC f, g, TRG h \\<^sub>B h \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" proof - have "(e ?a \\<^sub>B (f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B TRG h \\<^sub>B h) \\<^sub>B d ?d) = e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B LUNIT h) \\<^sub>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 \\<^sub>B e ?c) \\<^sub>B h"] by simp (* 5 sec *) thus ?thesis using assms by simp qed finally have L: "P \\<^sub>B[f, g, h] \\<^sub>B CMP (f \\<^sub>B g) h \\<^sub>B (CMP f g \\<^sub>B P h) = (e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \\<^sub>B g) \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, d ?b \\<^sub>B e ?b, g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f \\<^sub>B SRC f, g, TRG h \\<^sub>B h \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" by simp have "CMP f (g \\<^sub>B h) \\<^sub>B (P f \\<^sub>B CMP g h) \\<^sub>B \\<^sub>B[P f, P g, P h] = (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B LUNIT ((g \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P (g \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P (g \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P (g \\<^sub>B h)) \\<^sub>B (P f \\<^sub>B (e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \\<^sub>B (e ?b \\<^sub>B g \\<^sub>B LUNIT (h \\<^sub>B d ?d)) \\<^sub>B (e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h)) \\<^sub>B \\<^sub>B[P f, P g, P h]" unfolding CMP_def LUNIT_def using assms B.comp_assoc by simp also have "... = (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B LUNIT ((g \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P (g \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P (g \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P (g \\<^sub>B h)) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B LUNIT (h \\<^sub>B d ?d)) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>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 \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b \\<^sub>B e ?b, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P (g \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P (g \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P (g \\<^sub>B h)) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B LUNIT (h \\<^sub>B d ?d)) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "e ?a \\<^sub>B f \\<^sub>B LUNIT ((g \\<^sub>B h) \\<^sub>B d ?d) = e ?a \\<^sub>B f \\<^sub>B (LUNIT (g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g \\<^sub>B h, d ?d]" using assms LUNIT_hcomp [of "g \\<^sub>B h" "d ?d"] B.invert_side_of_triangle by simp also have "... = (e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g \\<^sub>B h, d ?d])" using assms B.whisker_left by simp finally have "e ?a \\<^sub>B f \\<^sub>B LUNIT ((g \\<^sub>B h) \\<^sub>B d ?d) = (e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g \\<^sub>B h, d ?d])" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P (g \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P (g \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P (g \\<^sub>B h)) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B LUNIT (h \\<^sub>B d ?d)) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "(e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B h) \\<^sub>B d ?d) = e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B h, d ?d] \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h) \\<^sub>B d ?d)" using assms B.whisker_left by simp also have "... = e ?a \\<^sub>B ((f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d]" using assms B.assoc'_naturality [of f "LUNIT (g \\<^sub>B h)" "d ?d"] LUNIT_in_hom [of "g \\<^sub>B h"] by auto also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d])" using assms B.whisker_left by simp finally have "(e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B LUNIT (g \\<^sub>B h) \\<^sub>B d ?d) = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d])" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P (g \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P (g \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P (g \\<^sub>B h)) \\<^sub>B ((P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B RUNIT g \\<^sub>B h \\<^sub>B d ?d)) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B LUNIT (h \\<^sub>B d ?d) = P f \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h, h \\<^sub>B d ?d]" using assms TRIANGLE [of g "h \\<^sub>B d ?d"] B.invert_side_of_triangle by simp also have "... = (P f \\<^sub>B e ?b \\<^sub>B RUNIT g \\<^sub>B h \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h, h \\<^sub>B d ?d])" using assms B.whisker_left P_def by simp finally have "P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B LUNIT (h \\<^sub>B d ?d) = (P f \\<^sub>B e ?b \\<^sub>B RUNIT g \\<^sub>B h \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, TRG h, h \\<^sub>B d ?d])" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P (g \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P (g \\<^sub>B h)] \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P (g \\<^sub>B h)) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "(P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B RUNIT g \\<^sub>B h \\<^sub>B d ?d) = P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d] \\<^sub>B (RUNIT g \\<^sub>B h \\<^sub>B d ?d)" using assms B.whisker_left P_def by simp also have "... = P f \\<^sub>B e ?b \\<^sub>B ((RUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]" using assms B.assoc'_naturality [of "RUNIT g" h "d ?d"] by auto also have "... = (P f \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d])" using assms B.whisker_left P_def by simp finally have "(P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B RUNIT g \\<^sub>B h \\<^sub>B d ?d) = (P f \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d])" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P (g \\<^sub>B h)] \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P (g \\<^sub>B h)] \\<^sub>B (((e ?a \\<^sub>B f) \\<^sub>B d ?b) \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "(\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P (g \\<^sub>B h)) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) = \\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d" using assms B.comp_arr_dom B.comp_cod_arr P_def B.interchange [of "\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b]" "P f" "P (g \\<^sub>B h)" "e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d"] by simp also have "... = (((e ?a \\<^sub>B f) \\<^sub>B d ?b) \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d)" using assms B.comp_arr_dom B.comp_cod_arr B.interchange [of "(e ?a \\<^sub>B f) \\<^sub>B d ?b" "\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b]" "e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d" "e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d"] by simp (* 4 sec *) finally have "(\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P (g \\<^sub>B h)) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) = (((e ?a \\<^sub>B f) \\<^sub>B d ?b) \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d)" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b \\<^sub>B e ?b, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P (g \\<^sub>B h)] \\<^sub>B ((e ?a \\<^sub>B f) \\<^sub>B d ?b \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" using assms P_def B.comp_assoc B.assoc_naturality [of "e ?a \\<^sub>B f" "d ?b" "e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d"] by simp also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, (d ?b \\<^sub>B e ?b) \\<^sub>B g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b \\<^sub>B e ?b, g \\<^sub>B h, d ?d]) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B d ?b \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P ((g \\<^sub>B d ?c \\<^sub>B e ?c) \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P ((g \\<^sub>B d ?c \\<^sub>B e ?c) \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B d ?c \\<^sub>B e ?c) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?c \\<^sub>B e ?c, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?c \\<^sub>B e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" using assms P_def B.comp_assoc B.assoc_naturality [of "e ?a" f "d ?b \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d"] by simp also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d]) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B TRG g \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P ((g \\<^sub>B d ?c \\<^sub>B e ?c) \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P ((g \\<^sub>B d ?c \\<^sub>B e ?c) \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "(e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B d ?b \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) = e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d] \\<^sub>B (d ?b \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d)" using assms B.whisker_left by simp also have "... = e ?a \\<^sub>B f \\<^sub>B (TRG g \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d]" using assms B.assoc'_naturality [of "d ?b" "e ?b" "(RUNIT g \\<^sub>B h) \\<^sub>B d ?d"] by auto also have "... = (e ?a \\<^sub>B f \\<^sub>B SRC f \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d])" using assms B.whisker_left by simp finally have "(e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, (g \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B d ?b \\<^sub>B e ?b \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) = (e ?a \\<^sub>B f \\<^sub>B SRC f \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d])" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B (TRG g \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "(e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B TRG g \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) = e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g \\<^sub>B h, d ?d] \\<^sub>B (TRG g \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d)" using assms B.whisker_left by simp also have "... = e ?a \\<^sub>B f \\<^sub>B ((TRG g \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]" using assms B.assoc'_naturality [of "TRG g" "RUNIT g \\<^sub>B h" "d ?d"] by simp also have "... = (e ?a \\<^sub>B f \\<^sub>B (TRG g \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \\<^sub>B SRC g) \\<^sub>B h, d ?d])" using assms B.whisker_left by simp finally have "(e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B TRG g \\<^sub>B (RUNIT g \\<^sub>B h) \\<^sub>B d ?d) = (e ?a \\<^sub>B f \\<^sub>B (TRG g \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \\<^sub>B SRC g) \\<^sub>B h, d ?d])" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B TRG g \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "(e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B (TRG g \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d) = e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d] \\<^sub>B (f \\<^sub>B (TRG g \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d)" using assms B.whisker_left by simp also have "... = e ?a \\<^sub>B ((f \\<^sub>B SRC f \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]" using assms B.assoc'_naturality [of f "(d ?b \\<^sub>B e ?b) \\<^sub>B (RUNIT g \\<^sub>B h)" "d ?d"] by simp also have "... = (e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B (g \\<^sub>B SRC g) \\<^sub>B h, d ?d])" using assms B.whisker_left by simp finally have "(e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B g \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B (TRG g \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d) = (e ?a \\<^sub>B (f \\<^sub>B TRG g \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B (g \\<^sub>B SRC g) \\<^sub>B h, d ?d])" using assms by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B \\<^sub>B[g, d ?c \\<^sub>B e ?c, h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d = e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B(g \\<^sub>B LUNIT h) \\<^sub>B \\<^sub>B[g, SRC g, h]) \\<^sub>B d ?d" using assms TRIANGLE [of g h] by simp also have "... = (e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B \\<^sub>B[g, TRG h, h]) \\<^sub>B d ?d)" using assms B.whisker_left B.whisker_right by simp finally have "e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B (RUNIT g \\<^sub>B h)) \\<^sub>B d ?d = (e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B \\<^sub>B[g, TRG h, h]) \\<^sub>B d ?d)" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B ((e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B TRG g \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B TRG g \\<^sub>B \\<^sub>B[g, SRC g, h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d = e ?a \\<^sub>B (f \\<^sub>B (LUNIT g \\<^sub>B h) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \\<^sub>B d ?d" using assms LUNIT_hcomp [of g h] B.invert_side_of_triangle by simp also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \\<^sub>B d ?d)" using assms B.whisker_left B.whisker_right by simp finally have "e ?a \\<^sub>B (f \\<^sub>B LUNIT (g \\<^sub>B h)) \\<^sub>B d ?d = (e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \\<^sub>B d ?d)" by simp thus ?thesis using assms B.comp_assoc by simp qed also have "... = ((e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B (TRG g \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B d ?d)) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \\<^sub>B h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B TRG g \\<^sub>B \\<^sub>B[g, SRC g, h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "(e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B TRG g \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) = e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h] \\<^sub>B (TRG g \\<^sub>B g \\<^sub>B LUNIT h)) \\<^sub>B d ?d" using assms B.whisker_left B.whisker_right by auto also have "... = e ?a \\<^sub>B (f \\<^sub>B ((TRG g \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, SRC g \\<^sub>B h]) \\<^sub>B d ?d" using assms B.assoc'_naturality [of "TRG g" g "LUNIT h"] by auto also have "... = (e ?a \\<^sub>B (f \\<^sub>B (TRG g \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \\<^sub>B h]) \\<^sub>B d ?d)" using assms B.whisker_left B.whisker_right by auto finally have "(e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B TRG g \\<^sub>B g \\<^sub>B LUNIT h) \\<^sub>B d ?d) = (e ?a \\<^sub>B (f \\<^sub>B (TRG g \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \\<^sub>B h]) \\<^sub>B d ?d)" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \\<^sub>B h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B TRG g \\<^sub>B \\<^sub>B[g, d ?c \\<^sub>B e ?c, h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B (g \\<^sub>B d ?c \\<^sub>B e ?c) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B d ?c \\<^sub>B e ?c) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" proof - have "(e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B (TRG g \\<^sub>B g) \\<^sub>B LUNIT h) \\<^sub>B d ?d) = e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B LUNIT h) \\<^sub>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 \\<^sub>B g" h "LUNIT h"] by simp (* 4 sec *) thus ?thesis using assms by simp qed finally have R: "CMP f (g \\<^sub>B h) \\<^sub>B (P f \\<^sub>B CMP g h) \\<^sub>B \\<^sub>B[P f, P g, P h] = (e ?a \\<^sub>B (f \\<^sub>B LUNIT g \\<^sub>B LUNIT h) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \\<^sub>B h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B \\<^sub>B[g, TRG h, h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>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 \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, (TRG g \\<^sub>B g) \\<^sub>B TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g \\<^sub>B g, TRG h \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f, d ?b \\<^sub>B e ?b, g \\<^sub>B (TRG h \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG h, h, d ?d]) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B[f \\<^sub>B SRC f, g, TRG h \\<^sub>B h \\<^sub>B d ?d]) \\<^sub>B (e ?a \\<^sub>B ((f \\<^sub>B SRC f) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c \\<^sub>B P h] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B (f \\<^sub>B SRC f) \\<^sub>B g, d ?c, P h] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, (f \\<^sub>B SRC f) \\<^sub>B g, d ?c] \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B SRC f, g, d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, SRC f, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B ((e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, g \\<^sub>B d ?c]) \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a, f, d ?b \\<^sub>B P g] \\<^sub>B P h) \\<^sub>B (\\<^sub>B[e ?a \\<^sub>B f, d ?b, P g] \\<^sub>B P h) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B P g) \\<^sub>B P h)" let ?LHSt = "(\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, (\<^bold>T\<^bold>R\<^bold>G g \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\ \<^bold>T\<^bold>R\<^bold>G h \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>T\<^bold>R\<^bold>G g \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>T\<^bold>R\<^bold>G h \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>T\<^bold>R\<^bold>G g, \<^bold>\g\<^bold>\, (\<^bold>T\<^bold>R\<^bold>G h \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>\e ?b\<^bold>\, \<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>T\<^bold>R\<^bold>G h \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\e ?a\<^bold>\ \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C f) \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>T\<^bold>R\<^bold>G h, \<^bold>\h\<^bold>\, \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C f, \<^bold>\g\<^bold>\, \<^bold>T\<^bold>R\<^bold>G h \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\e ?a\<^bold>\ \<^bold>\ ((\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C f) \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\d ?c\<^bold>\, \<^bold>\e ?c\<^bold>\, \<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?a\<^bold>\, (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C f) \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\d ?c\<^bold>\ \<^bold>\ \<^bold>P\<^bold>R\<^bold>J h\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?a\<^bold>\ \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C f) \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\d ?c\<^bold>\, \<^bold>P\<^bold>R\<^bold>J h\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e ?a\<^bold>\, (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C f) \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\d ?c\<^bold>\\<^bold>] \<^bold>\ \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\ ((\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C f, \<^bold>\g\<^bold>\, \<^bold>\d ?c\<^bold>\\<^bold>]) \<^bold>\ \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\ ((\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>S\<^bold>R\<^bold>C f, \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\d ?c\<^bold>\\<^bold>]) \<^bold>\ \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\ ((\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\d ?b\<^bold>\, \<^bold>\e ?b\<^bold>\, \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\d ?c\<^bold>\\<^bold>]) \<^bold>\ \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\e ?a\<^bold>\, \<^bold>\f\<^bold>\, \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>P\<^bold>R\<^bold>J g\<^bold>] \<^bold>\ \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\f\<^bold>\, \<^bold>\d ?b\<^bold>\, \<^bold>P\<^bold>R\<^bold>J g\<^bold>] \<^bold>\ \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\ ((\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e ?a\<^bold>\, \<^bold>\f\<^bold>\, \<^bold>\d ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>P\<^bold>R\<^bold>J g) \<^bold>\ \<^bold>P\<^bold>R\<^bold>J h)" let ?RHS = "(e ?a \\<^sub>B (f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, g, TRG h \\<^sub>B h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B (f \\<^sub>B SRC f \\<^sub>B \\<^sub>B[g, TRG h, h]) \\<^sub>B d ?d) \\<^sub>B (e ?a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f, TRG g \\<^sub>B (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[TRG g, (g \\<^sub>B SRC g) \\<^sub>B h, d ?d]) \\<^sub>B (e ?a \\<^sub>B f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d]) \\<^sub>B \\<^sub>B[e ?a, f, d ?b \\<^sub>B P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B \\<^sub>B[e ?a \\<^sub>B f, d ?b, P ((g \\<^sub>B SRC g) \\<^sub>B h)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?a, f, d ?b] \\<^sub>B e ?b \\<^sub>B ((g \\<^sub>B SRC g) \\<^sub>B h) \\<^sub>B d ?d) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B SRC g, h, d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, SRC g, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B e ?b \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?c, e ?c, h \\<^sub>B d ?d]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b, g, d ?c \\<^sub>B P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B[e ?b \\<^sub>B g, d ?c, P h]) \\<^sub>B (P f \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, g, d ?c] \\<^sub>B P h) \\<^sub>B \\<^sub>B[P f, P g, P h]" let ?RHSt = "(\<^bold>\e ?a\<^bold>\ \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>T\<^bold>R\<^bold>G g, \<^bold>\g\<^bold>\, \<^bold>T\<^bold>R\<^bold>G h \<^bold>\ \<^bold>\h\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\d ?d\<^bold>\) \<^bold>\ (\<^bold>\e ?a\<^bold>\ \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C f \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\, \<^bold>T\<^bold>R\<^bold>G h, \<^bold>\h\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\d ?d\<^bold>\) \<^bold>\ (\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>T\<^bold>R\<^bold>G g \<^bold>\ (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>T\<^bold>R\<^bold>G g, (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\d ?b\<^bold>\, \<^bold>\e ?b\<^bold>\, ((\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?a\<^bold>\, \<^bold>\f\<^bold>\, \<^bold>\d ?b\<^bold>\ \<^bold>\ (\<^bold>\e ?b\<^bold>\ \<^bold>\ ((\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ \<^bold>\d ?d\<^bold>\)\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?a\<^bold>\ \<^bold>\ \<^bold>\f\<^bold>\, \<^bold>\d ?b\<^bold>\, (\<^bold>\e ?b\<^bold>\ \<^bold>\ ((\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ \<^bold>\d ?d\<^bold>\)\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e ?a\<^bold>\, \<^bold>\f\<^bold>\, \<^bold>\d ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\e ?b\<^bold>\ \<^bold>\ ((\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C g) \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ \<^bold>\d ?d\<^bold>\) \<^bold>\ (\<^bold>P\<^bold>R\<^bold>J f \<^bold>\ \<^bold>\e ?b\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>S\<^bold>R\<^bold>C g, \<^bold>\h\<^bold>\, \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>P\<^bold>R\<^bold>J f \<^bold>\ \<^bold>\e ?b\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>S\<^bold>R\<^bold>C g, \<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>P\<^bold>R\<^bold>J f \<^bold>\ \<^bold>\e ?b\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\d ?c\<^bold>\, \<^bold>\e ?c\<^bold>\, \<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\d ?d\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>P\<^bold>R\<^bold>J f \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?b\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\d ?c\<^bold>\ \<^bold>\ \<^bold>P\<^bold>R\<^bold>J h\<^bold>]) \<^bold>\ (\<^bold>P\<^bold>R\<^bold>J f \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?b\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\d ?c\<^bold>\, \<^bold>P\<^bold>R\<^bold>J h\<^bold>]) \<^bold>\ (\<^bold>P\<^bold>R\<^bold>J f \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e ?b\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\d ?c\<^bold>\\<^bold>] \<^bold>\ \<^bold>P\<^bold>R\<^bold>J h) \<^bold>\ \<^bold>\\<^bold>[\<^bold>P\<^bold>R\<^bold>J f, \<^bold>P\<^bold>R\<^bold>J g, \<^bold>P\<^bold>R\<^bold>J h\<^bold>]" have E: "?LHS = ?RHS" proof - have "?LHS = \?LHSt\" using assms B.\_def B.\'.map_ide_simp B.VVV.ide_char B.VVV.arr_char B.VV.ide_char B.VV.arr_char 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_char B.VVV.arr_char B.VV.ide_char B.VV.arr_char 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_char B.VV.arr_char P.FF_def hcomp_def arr_char P\<^sub>0_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)) \\<^sub>B HoPP.map fg" using fg 1 B.VV.ide_char B.VV.arr_char B.VV.cod_char HoPP.preserves_arr P.FF_def hcomp_char comp_char by auto also have "... = PoH.map fg \\<^sub>B CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg))" using fg CMP_naturality [of "fst fg" "snd fg"] B.VV.ide_char B.VV.arr_char B.VV.dom_char B.VV.cod_char P.FF_def hcomp_def arr_char P\<^sub>0_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 "P\<^sub>0 (src\<^sub>B (snd fg)) \ Obj \ P\<^sub>0 (trg\<^sub>B (fst fg)) \ Obj" using fg P\<^sub>0_props(6) B.VV.arrE B.obj_src B.obj_trg by meson moreover have "B.seq (P (fst fg \\<^sub>B snd fg)) (CMP (fst (B.VV.dom fg)) (snd (B.VV.dom fg)))" using fg 1 B.VV.ide_char B.VV.arr_char B.VV.dom_char by simp ultimately show ?thesis using fg 1 comp_char arr_char in_hom_char 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_char B.VV.arr_char CMP.map_simp_ide by unfold_locales simp definition \\<^sub>P where "\\<^sub>P = CMP.map" interpretation \\<^sub>P: natural_isomorphism B.VV.comp comp HoPP.map PoH.map \\<^sub>P unfolding \\<^sub>P_def using CMP.natural_isomorphism_axioms by simp no_notation B.in_hom ("\_ : _ \\<^sub>B _\") lemma \\<^sub>P_in_hom [intro]: assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g" shows "\\\<^sub>P (f, g) : src (P g) \ trg (P f)\" and "\\\<^sub>P (f, g) : P f \ P g \ P (f \\<^sub>B g)\" proof - show 1: "\\\<^sub>P (f, g) : P f \ P g \ P (f \\<^sub>B g)\" using assms B.VV.ide_char B.VV.arr_char P.FF_def by auto show "\\\<^sub>P (f, g) : src (P g) \ trg (P f)\" using 1 by (metis (no_types, lifting) hcomp_simps(2) in_hhom_def in_hom_char src_hcomp vconn_implies_hpar(1-2)) qed lemma \\<^sub>P_simps [simp]: assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g" shows "arr (\\<^sub>P (f, g))" and "src (\\<^sub>P (f, g)) = src (P g)" and "trg (\\<^sub>P (f, g)) = trg (P f)" and "dom (\\<^sub>P (f, g)) = P f \ P g" and "cod (\\<^sub>P (f, g)) = P (f \\<^sub>B g)" using assms \\<^sub>P_in_hom by blast+ sublocale prj: pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \ \\<^sub>B src trg P \\<^sub>P proof fix f g h assume f: "B.ide f" and g: "B.ide g" and h: "B.ide h" and fg: "src\<^sub>B f = trg\<^sub>B g" and gh: "src\<^sub>B g = trg\<^sub>B h" have 1: "ide (P f) \ ide (P g) \ ide (P h)" using f g h P.preserves_ide by simp have "P \\<^sub>B[f, g, h] \ \\<^sub>P (f \\<^sub>B g, h) \ (\\<^sub>P (f, g) \ P h) = P \\<^sub>B[f, g, h] \\<^sub>B \\<^sub>P (f \\<^sub>B g, h) \\<^sub>B (\\<^sub>P (f, g) \\<^sub>B P h)" using f g h fg gh B.VV.arr_char B.VV.dom_char B.VV.cod_char P.FF_def by (intro comp_eqI hcomp_eqI seqI hseqI') auto also have "... = CMP f (g \\<^sub>B h) \\<^sub>B (P f \\<^sub>B CMP g h) \\<^sub>B \\<^sub>B[P f, P g, P h]" using f g h fg gh CMP_coherence CMP.map_simp_ide B.VV.ide_char B.VV.arr_char \\<^sub>P_def by simp also have "... = \\<^sub>P (f, g \\<^sub>B h) \\<^sub>B (P f \\<^sub>B \\<^sub>P (g, h)) \\<^sub>B \\<^sub>B[P f, P g, P h]" using f g h fg gh B.VV.ide_char B.VV.arr_char \\<^sub>P_def by simp also have "... = \\<^sub>P (f, g \\<^sub>B h) \ (P f \\<^sub>B \\<^sub>P (g, h)) \ \\<^sub>B[P f, P g, P h]" proof - have 2: "arr ((P f \ \\<^sub>P (g, h)) \ \[P f, P g, P h])" using f g h fg gh B.VV.arr_char B.VV.dom_char P.FF_def by auto moreover have "(P f \ \\<^sub>P (g, h)) \ \[P f, P g, P h] = (P f \\<^sub>B \\<^sub>P (g, h)) \\<^sub>B \\<^sub>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 \\<^sub>B \\<^sub>P (g, h))" using f g h fg gh arr_hcompI by (intro arr_hcompI hseqI') auto moreover have "B.dom (P f \\<^sub>B \\<^sub>P (g, h)) = P f \\<^sub>B P g \\<^sub>B P h" proof - have "B.dom (P f \\<^sub>B \\<^sub>P (g, h)) = P f \\<^sub>B B.dom (\\<^sub>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_char) also have "... = P f \\<^sub>B P g \\<^sub>B P h" using g h gh dom_char hcomp_char \\<^sub>P_simps(1-4) by auto finally show ?thesis by blast qed moreover have "B.dom (\\<^sub>P (f, g \\<^sub>B h)) = B.cod ((P f \\<^sub>B \\<^sub>P (g, h)) \\<^sub>B \\<^sub>B[P f, P g, P h])" proof - have "B.dom (\\<^sub>P (f, g \\<^sub>B h)) = dom (\\<^sub>P (f, g \\<^sub>B h))" using f g h fg gh B.VV.arr_char dom_char by simp also have "... = cod ((P f \ \\<^sub>P (g, h)) \ \[P f, P g, P h])" using f g h fg gh 2 VV.arr_char \\<^sub>P_simps by simp also have "... = B.cod ((P f \ \\<^sub>P (g, h)) \ \[P f, P g, P h])" using 2 cod_char by presburger also have "... = B.cod ((P f \\<^sub>B \\<^sub>P (g, h)) \\<^sub>B \\<^sub>B[P f, P g, P h])" proof - have "(P f \ \\<^sub>P (g, h)) \ \[P f, P g, P h] = (P f \\<^sub>B \\<^sub>P (g, h)) \\<^sub>B \\<^sub>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 \\<^sub>B[P f, P g, P h] = P f \\<^sub>B P g \\<^sub>B P h" using f g h fg gh 1 ide_char by auto ultimately show ?thesis using f g h fg gh B.VV.arr_char assoc_simp assoc_simps(1) dom_char cod_char by (intro comp_eqI' seqI arr_compI arr_hcompI) auto qed also have "... = \\<^sub>P (f, g \\<^sub>B h) \ (P f \ \\<^sub>P (g, h)) \ \[P f, P g, P h]" using f g h fg gh assoc_simp hcomp_eqI \\<^sub>P_simps(1) by auto finally show "P \\<^sub>B[f, g, h] \ \\<^sub>P (f \\<^sub>B g, h) \ (\\<^sub>P (f, g) \ P h) = \\<^sub>P (f, g \\<^sub>B h) \ (P f \ \\<^sub>P (g, h)) \ \[P f, P g, P h]" by blast qed lemma pseudofunctor_prj: shows "pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \ \\<^sub>B src trg P \\<^sub>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 \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>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 V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e a\ \d a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence(1) by simp let ?x = "(e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)" have x: "\?x : P.map\<^sub>0 a \\<^sub>B P a\" using assms P_def P.map\<^sub>0_def P_map\<^sub>0_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 \\<^sub>B d a : P\<^sub>0 a \\<^sub>B a\ \ B.ide (a \\<^sub>B d a)" using assms B.obj_simps by auto show x\<^sub>S: "\?x : P.map\<^sub>0 a \ P a\" using assms x P_map\<^sub>0_simp P_simps\<^sub>B(1) arr_char B.arrI equivalence_data_simps\<^sub>B(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_char) show "iso ?x" using assms x\<^sub>S B.isos_compose iso_char arr_char by auto have *: "\?x : P\<^sub>0 a \ P\<^sub>0 a\" using assms x\<^sub>S P\<^sub>0_props vconn_implies_hpar(1-2) by (intro in_hhomI) auto show "?x \ \\<^sub>B[P.map\<^sub>0 a] = (P \\<^sub>B[a] \ \\<^sub>P (a, a)) \ (?x \ ?x)" proof - have 0: "\d a \\<^sub>B e a \\<^sub>B a \\<^sub>B d a : P\<^sub>0 a \\<^sub>B a\" using assms by force have 1: "B.seq (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]" using assms by (elim B.objE, intro B.seqI) auto have "(P \\<^sub>B[a] \ \\<^sub>P (a, a)) \ (?x \ ?x) = P \\<^sub>B[a] \ \\<^sub>P (a, a) \ (?x \ ?x)" using comp_assoc by simp also have "... = P \\<^sub>B[a] \ (CMP a a \ (P a \ P a)) \ (?x \ ?x)" unfolding \\<^sub>P_def CMP.map_def using assms B.VV.arr_char B.VV.cod_char P.FF_def by auto also have "... = P \\<^sub>B[a] \ (CMP a a \ (P a \\<^sub>B P a)) \ (?x \\<^sub>B ?x)" using assms x\<^sub>S hcomp_char src_def trg_def by auto also have "... = P \\<^sub>B[a] \\<^sub>B (CMP a a \\<^sub>B (P a \\<^sub>B P a)) \\<^sub>B (?x \\<^sub>B ?x)" proof - have "\P \\<^sub>B[a] \ (CMP a a \ (P a \\<^sub>B P a)) \ (?x \\<^sub>B ?x) : P\<^sub>0 a \ P\<^sub>0 a \ P a\" proof (intro comp_in_homI) show "\?x \\<^sub>B ?x : P\<^sub>0 a \ P\<^sub>0 a \ P a \ P a\" proof - have "\?x \ ?x : P\<^sub>0 a \ P\<^sub>0 a \ P a \ P a\" using assms x\<^sub>S * P\<^sub>0_props P.map\<^sub>0_simps(3) by (intro hcomp_in_vhom) auto moreover have "?x \ ?x = ?x \\<^sub>B ?x" using x\<^sub>S * by (intro hcomp_eqI) auto ultimately show ?thesis by simp qed show "\P a \\<^sub>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 \\<^sub>B a)\" using assms CMP_in_hom(2) by auto show "\P \\<^sub>B[a] : P (a \\<^sub>B a) \ P a\" using assms by auto qed thus ?thesis by (intro comp_eqI hcomp_eqI) auto qed also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a] \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a)) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, d a \\<^sub>B P a] \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a) \\<^sub>B (P a \\<^sub>B P a) \\<^sub>B (?x \\<^sub>B ?x)" using assms B.comp_assoc CMP_def by auto also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, d a \\<^sub>B P a] \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a) \\<^sub>B (P a \\<^sub>B P a) \\<^sub>B (?x \\<^sub>B ?x)" using assms 1 B.whisker_left B.comp_assoc by fastforce also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, d a \\<^sub>B P a] \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a) \\<^sub>B ((P a \\<^sub>B P a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]))) \\<^sub>B (B.inv (\ a) \\<^sub>B B.inv (\ a))" using assms B.interchange B.comp_assoc by simp also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B ((e a \\<^sub>B a \\<^sub>B B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a])) \\<^sub>B \\<^sub>B[e a, a, d a \\<^sub>B P a] \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B (B.inv (\ a) \\<^sub>B B.inv (\ a))" proof - have "(P a \\<^sub>B P a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) = ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]))" unfolding P_def using assms B.comp_cod_arr [of "(e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])" "(e a \\<^sub>B a \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B a \\<^sub>B d a)"] by auto thus ?thesis using B.comp_assoc by simp qed also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B ((e a \\<^sub>B a \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, d a \\<^sub>B P a]) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B (B.inv (\ a) \\<^sub>B B.inv (\ a))" proof - have "(e a \\<^sub>B a \\<^sub>B B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) = e a \\<^sub>B (a \\<^sub>B B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B (a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a])" proof - have "B.seq (a \\<^sub>B B.inv (\ a) \\<^sub>B a \\<^sub>B d a) (a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>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 \\<^sub>B a \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]" using assms 1 B.whisker_left by fastforce finally have "(e a \\<^sub>B a \\<^sub>B B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) = e a \\<^sub>B a \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]" by blast thus ?thesis using B.comp_assoc by simp qed also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B (\\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a)) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B (B.inv (\ a) \\<^sub>B B.inv (\ a))" proof - have "(e a \\<^sub>B a \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, d a \\<^sub>B P a] = \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a])" using assms 1 a_da P_def B.assoc_naturality [of "e a" a "(B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]"] by fastforce thus ?thesis using B.comp_assoc by simp qed also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B (((e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a \\<^sub>B P a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B[a, d a, P a]) \\<^sub>B \\<^sub>B[e a, a \\<^sub>B d a, P a] \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B (B.inv (\ a) \\<^sub>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_char P.preserves_ide by auto have 2: "src\<^sub>B (e a) = trg\<^sub>B a \ src\<^sub>B a = trg\<^sub>B (d a) \ src\<^sub>B (d a) = trg\<^sub>B (P a)" using assms by auto have "((e a \\<^sub>B \\<^sub>B[a, d a, P a]) \\<^sub>B \\<^sub>B[e a, a \\<^sub>B d a, P a]) \\<^sub>B (\\<^sub>B[e a, a, d a] \\<^sub>B P a) = \\<^sub>B[e a, a, d a \\<^sub>B P a] \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a]" using assms 1 2 B.pentagon B.comp_assoc by force hence "(e a \\<^sub>B \\<^sub>B[a, d a, P a]) \\<^sub>B \\<^sub>B[e a, a \\<^sub>B d a, P a] = (\\<^sub>B[e a, a, d a \\<^sub>B P a] \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a]) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a)" using assms 1 2 B.invert_side_of_triangle(2) [of "\\<^sub>B[e a, a, d a \\<^sub>B P a] \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a]" "(e a \\<^sub>B \\<^sub>B[a, d a, P a]) \\<^sub>B \\<^sub>B[e a, a \\<^sub>B d a, P a]" "\\<^sub>B[e a, a, d a] \\<^sub>B P a"] by fastforce hence "\\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a) = \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a \\<^sub>B P a] \\<^sub>B (e a \\<^sub>B \\<^sub>B[a, d a, P a]) \\<^sub>B \\<^sub>B[e a, a \\<^sub>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 \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B (e a \\<^sub>B a \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B ((e a \\<^sub>B \\<^sub>B[a, d a, P a]) \\<^sub>B \\<^sub>B[e a, a \\<^sub>B d a, P a]) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B (B.inv (\ a) \\<^sub>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) \\<^sub>B a \\<^sub>B d a) \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]" using assms by (elim B.objE, intro B.seqI) auto moreover have "src\<^sub>B a = trg\<^sub>B ((B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a])" using assms a_da trg_vcomp by fastforce moreover have "src\<^sub>B a = a \ trg\<^sub>B a = a" using assms by auto moreover have "B.dom ((B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) = d a \\<^sub>B e a \\<^sub>B a \\<^sub>B d a" using assms a_da by fastforce moreover have "B.cod ((B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) = a \\<^sub>B a \\<^sub>B d a" using assms a_da by fastforce ultimately have "((e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a \\<^sub>B P a] = \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B (e a \\<^sub>B a \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a])" unfolding P_def using assms B.assoc'_naturality [of "e a" a "(B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]"] by auto thus ?thesis using B.comp_assoc by simp qed also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, d a \\<^sub>B P a]) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B (B.inv (\ a) \\<^sub>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_char P.preserves_ide by auto have 2: "src\<^sub>B (e a) = trg\<^sub>B a \ src\<^sub>B a = trg\<^sub>B (d a) \ src\<^sub>B (d a) = trg\<^sub>B (P a)" using assms by auto have "((e a \\<^sub>B \\<^sub>B[a, d a, P a]) \\<^sub>B \\<^sub>B[e a, a \\<^sub>B d a, P a]) \\<^sub>B (\\<^sub>B[e a, a, d a] \\<^sub>B P a) = \\<^sub>B[e a, a, d a \\<^sub>B P a] \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a]" using assms 1 2 B.pentagon B.comp_assoc by fastforce hence "(e a \\<^sub>B \\<^sub>B[a, d a, P a]) \\<^sub>B \\<^sub>B[e a, a \\<^sub>B d a, P a] = \\<^sub>B[e a, a, d a \\<^sub>B P a] \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a)" using assms 1 2 P.preserves_ide B.comp_assoc B.invert_side_of_triangle(2) [of "\\<^sub>B[e a, a, d a \\<^sub>B P a] \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a]" "(e a \\<^sub>B \\<^sub>B[a, d a, P a]) \\<^sub>B \\<^sub>B[e a, a \\<^sub>B d a, P a]" "\\<^sub>B[e a, a, d a] \\<^sub>B P a"] by force thus ?thesis using B.comp_assoc by simp qed also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B (B.inv (\ a) \\<^sub>B B.inv (\ a))" proof - have "(e a \\<^sub>B a \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, d a \\<^sub>B P a] = \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a])" proof - have 1: "B.seq (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]" using assms by (elim B.objE, intro B.seqI) auto moreover have "src\<^sub>B a = trg\<^sub>B ((B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a])" using a_da by force moreover have "B.arr a \ B.dom a = a \ B.cod a = a \ src\<^sub>B a = a \ trg\<^sub>B a = a" using assms by auto moreover have "B.dom ((B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) = d a \\<^sub>B e a \\<^sub>B a \\<^sub>B d a" using a_da by auto moreover have "B.cod ((B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) = a \\<^sub>B a \\<^sub>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) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]"] by auto qed thus ?thesis using B.comp_assoc by simp qed also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a))) \\<^sub>B (B.inv (\ a) \\<^sub>B P\<^sub>0 a)" proof - have "((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B (B.inv (\ a) \\<^sub>B B.inv (\ a)) = ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) \\<^sub>B (B.inv (\ a) \\<^sub>B P\<^sub>0 a)" proof - have "((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B (B.inv (\ a) \\<^sub>B B.inv (\ a)) = (((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B ((e a \\<^sub>B d a) \\<^sub>B B.inv (\ a))) \\<^sub>B (B.inv (\ a) \\<^sub>B P\<^sub>0 a)" using assms B.comp_arr_dom B.comp_cod_arr B.comp_assoc B.interchange [of "e a \\<^sub>B d a" "B.inv (\ a)" "B.inv (\ a)" "P\<^sub>0 a"] by simp also have "... = ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) \\<^sub>B (B.inv (\ a) \\<^sub>B P\<^sub>0 a)" using assms B.interchange by simp also have "... = ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) \\<^sub>B (B.inv (\ a) \\<^sub>B P\<^sub>0 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 \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B (\\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (((e a \\<^sub>B a) \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a))) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B P\<^sub>0 a) \\<^sub>B (B.inv (\ a) \\<^sub>B P\<^sub>0 a)" proof - have "(\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) = (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B P a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a))" proof - have "B.seq \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])" using assms a_da by (elim B.objE, intro B.seqI) auto moreover have "B.seq (P a) ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a))" using assms a_da P_def by auto ultimately show ?thesis using assms B.interchange by simp qed also have "... = (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a))" proof - have "P a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) = (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)" using assms x P_def B.comp_cod_arr by blast moreover have "\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) = \\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a" proof - have "\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) = B.inv ((e a \\<^sub>B \\<^sub>B[d a]) \\<^sub>B \\<^sub>B[e a, a, d a])" proof - have "B.inv ((e a \\<^sub>B \\<^sub>B[d a]) \\<^sub>B \\<^sub>B[e a, a, d a]) = B.inv \\<^sub>B[e a, a, d a] \\<^sub>B B.inv (e a \\<^sub>B \\<^sub>B[d a])" proof - have "B.seq (e a \\<^sub>B \\<^sub>B[d a]) \\<^sub>B[e a, a, d a]" proof - have "B.iso ((e a \\<^sub>B \\<^sub>B[d a]) \\<^sub>B \\<^sub>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 \\<^sub>B[e a, a, d a]" using assms by force ultimately show ?thesis using assms B.inv_comp by auto qed also have "... = \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])" using assms by (elim B.objE) (simp add: assms) finally show ?thesis by simp qed also have "... = B.inv (\\<^sub>B[e a] \\<^sub>B d a)" using assms B.triangle [of "d a" "e a"] by simp also have "... = \\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a" using assms by simp finally show ?thesis by blast qed ultimately show ?thesis by simp qed also have "... = (((e a \\<^sub>B a) \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) \\<^sub>B P\<^sub>0 a" using assms 0 B.comp_cod_arr B.comp_arr_dom by (elim B.objE) auto also have "... = ((e a \\<^sub>B a) \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) \\<^sub>B P\<^sub>0 a" using B.comp_assoc by simp also have "... = (((e a \\<^sub>B a) \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B P\<^sub>0 a)" proof - have "B.seq ((e a \\<^sub>B a) \\<^sub>B d a) (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]))" using assms 0 a_da by (elim B.objE, intro B.seqI) auto moreover have "B.seq ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) (P\<^sub>0 a)" using assms 0 apply (intro B.seqI, auto simp add: B.obj_simps(5) P\<^sub>0_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 "(\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B P a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) = (((e a \\<^sub>B a) \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B P\<^sub>0 a)" by blast thus ?thesis using B.comp_assoc by simp qed also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B (((e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B d a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a))) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P\<^sub>0 a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B P\<^sub>0 a) \\<^sub>B (B.inv (\ a) \\<^sub>B P\<^sub>0 a)" proof - have "\\<^sub>B[e a \\<^sub>B a, d a, P a] \\<^sub>B (((e a \\<^sub>B a) \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) = ((e a \\<^sub>B a) \\<^sub>B d a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P\<^sub>0 a]" proof - have "B.hseq (e a) a" using assms by force moreover have "src\<^sub>B (d a) = trg\<^sub>B ?x" using assms B.trg_vcomp [of "e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]" "B.inv (\ a)"] by simp moreover have "B.cod ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) = P a" using assms x by blast ultimately show ?thesis using assms B.assoc_naturality [of "e a \\<^sub>B a" "d a" ?x] by auto qed thus ?thesis using B.comp_assoc by simp qed also have "... = P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a] \\<^sub>B (d a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a))) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P\<^sub>0 a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B P\<^sub>0 a) \\<^sub>B (B.inv (\ a) \\<^sub>B P\<^sub>0 a)" proof - have "((e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B d a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) = (e a \\<^sub>B a) \\<^sub>B (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a] \\<^sub>B (d a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a))" proof - have "B.ide (e a \\<^sub>B a)" using assms by force moreover have "B.seq ((B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a]) (d a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>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 \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (B.inv (\ a) \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \\<^sub>B (d a \\<^sub>B B.inv (\ a))) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P\<^sub>0 a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B P\<^sub>0 a) \\<^sub>B (B.inv (\ a) \\<^sub>B P\<^sub>0 a)" proof - have "(B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a] \\<^sub>B (d a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) = (a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (B.inv (\ a) \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \\<^sub>B (d a \\<^sub>B B.inv (\ a))" proof - have "(B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a] \\<^sub>B (d a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) = (B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[d a, e a, a \\<^sub>B d a] \\<^sub>B (d a \\<^sub>B e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B (d a \\<^sub>B B.inv (\ a))" using assms B.whisker_left B.comp_assoc by simp also have "... = ((B.inv (\ a) \\<^sub>B a \\<^sub>B d a) \\<^sub>B ((d a \\<^sub>B e a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \\<^sub>B (d a \\<^sub>B B.inv (\ a))" using assms B.assoc'_naturality [of "d a" "e a" "\\<^sub>B\<^sup>-\<^sup>1[d a]"] B.comp_assoc by simp also have "... = (B.inv (\ a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \\<^sub>B (d a \\<^sub>B B.inv (\ a))" using assms B.interchange [of "B.inv (\ a)" "d a \\<^sub>B e a" "a \\<^sub>B d a" "\\<^sub>B\<^sup>-\<^sup>1[d a]"] B.comp_arr_dom B.comp_cod_arr by simp also have "... = (a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (B.inv (\ a) \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \\<^sub>B (d a \\<^sub>B B.inv (\ a))" using assms B.interchange [of a "B.inv (\ a)" "\\<^sub>B\<^sup>-\<^sup>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 \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a]) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P\<^sub>0 a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B P\<^sub>0 a) \\<^sub>B (B.inv (\ a) \\<^sub>B P\<^sub>0 a)" proof - have "(B.inv (\ a) \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \\<^sub>B (d a \\<^sub>B B.inv (\ a)) = \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a]" proof - have "(B.inv (\ a) \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a, e a, d a] \\<^sub>B (d a \\<^sub>B B.inv (\ a)) = B.inv (\ a \\<^sub>B d a) \\<^sub>B B.inv \\<^sub>B[d a, e a, d a] \\<^sub>B B.inv (d a \\<^sub>B \ a)" using assms by simp also have "... = B.inv ((d a \\<^sub>B \ a) \\<^sub>B \\<^sub>B[d a, e a, d a] \\<^sub>B (\ a \\<^sub>B d a))" proof - have "B.iso ((d a \\<^sub>B \ a) \\<^sub>B \\<^sub>B[d a, e a, d a] \\<^sub>B (\ a \\<^sub>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 (\\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a])" using assms triangle_right by simp also have "... = \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>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 \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a]) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P\<^sub>0 a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B P\<^sub>0 a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a \\<^sub>B d a]) \\<^sub>B B.inv (\ a) \\<^sub>B \\<^sub>B[P\<^sub>0 a]" proof - have "\\<^sub>B[e a \\<^sub>B d a] \\<^sub>B (B.inv (\ a) \\<^sub>B P\<^sub>0 a) = B.inv (\ a) \\<^sub>B \\<^sub>B[P\<^sub>0 a]" using assms B.runit_naturality [of "B.inv (\ a)"] B.unitor_coincidence P\<^sub>0_props by simp hence "B.inv (\ a) \\<^sub>B P\<^sub>0 a = \\<^sub>B\<^sup>-\<^sup>1[e a \\<^sub>B d a] \\<^sub>B B.inv (\ a) \\<^sub>B \\<^sub>B[P\<^sub>0 a]" using assms P\<^sub>0_props(5) B.invert_side_of_triangle(1) by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) \\<^sub>B \\<^sub>B[P\<^sub>0 a]" proof - have P\<^sub>0_a: "B.obj (P\<^sub>0 a) \ B.arr (P\<^sub>0 a)" using assms a_da by fastforce have i_a: "B.\ a = \\<^sub>B[a]" using assms B.unitor_coincidence B.\_ide_simp B.obj_simps by (metis (no_types, lifting) B.objE) have "P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a]) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P\<^sub>0 a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B P\<^sub>0 a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a \\<^sub>B d a] = e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]" proof - have "P \\<^sub>B[a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a, a, d a]) \\<^sub>B (e a \\<^sub>B a \\<^sub>B \\<^sub>B[a \\<^sub>B d a]) \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B \\<^sub>B[e a, a, a \\<^sub>B a \\<^sub>B d a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B (a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a]) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, P\<^sub>0 a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B P\<^sub>0 a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a \\<^sub>B d a] = \(\<^bold>\e a\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\a\<^bold>\\<^sub>0\<^bold>] \<^bold>\ \<^bold>\d a\<^bold>\) \<^bold>\ (\<^bold>\e a\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\a\<^bold>\\<^sub>0, \<^bold>\a\<^bold>\\<^sub>0, \<^bold>\d a\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\e a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\\<^bold>[\<^bold>\a\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\d a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e a\<^bold>\, \<^bold>\a\<^bold>\\<^sub>0, \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\d a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e a\<^bold>\, \<^bold>\a\<^bold>\\<^sub>0, \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\d a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e a\<^bold>\, \<^bold>\a\<^bold>\\<^sub>0, \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\d a\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\e a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\\<^sub>0) \<^bold>\ (\<^bold>\a\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\d a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\d a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\d a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\\<^sub>0, \<^bold>\d a\<^bold>\, \<^bold>\P\<^sub>0 a\<^bold>\\<^sub>0\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e a\<^bold>\, \<^bold>\a\<^bold>\\<^sub>0, \<^bold>\d a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\e a\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\d a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\P\<^sub>0 a\<^bold>\\<^sub>0) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e a\<^bold>\ \<^bold>\ \<^bold>\d a\<^bold>\\<^bold>]\" unfolding P_def using assms B.obj_def [of a] P\<^sub>0_a i_a B.\_def B.\'.map_ide_simp B.VVV.ide_char B.VVV.arr_char B.VV.ide_char B.VV.arr_char B.\_ide_simp B.\_ide_simp by (elim B.objE) simp also have "... = \\<^bold>\e a\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\d a\<^bold>\\<^bold>]\" using assms P\<^sub>0_a by (intro EV.eval_eqI) auto also have "... = e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>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 \ \\<^sub>B[P.map\<^sub>0 a]" proof - have "B.arr \\<^sub>B[P.map\<^sub>0 a] \ src\<^sub>B \\<^sub>B[P.map\<^sub>0 a] \ Obj \ trg\<^sub>B \\<^sub>B[P.map\<^sub>0 a] \ Obj \ P\<^sub>0 a \\<^sub>B P\<^sub>0 a \ Obj \ B.cod \\<^sub>B[P.map\<^sub>0 a] = P\<^sub>0 a" using assms P\<^sub>0_props arr_char 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 P\<^sub>0_props [of a] by simp qed finally show ?thesis by argo qed qed thus ?thesis by simp qed sublocale PoE: composite_pseudofunctor comp hcomp \ \\<^sub>B src trg V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \ \\<^sub>B src trg E \\<^sub>E P \\<^sub>P .. sublocale EoP: composite_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \ \\<^sub>B src trg V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B P \\<^sub>P E \\<^sub>E .. sublocale I\<^sub>C: identity_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B .. sublocale I\<^sub>S: identity_pseudofunctor comp hcomp \ \\<^sub>B src trg .. no_notation B.in_hom ("\_ : _ \\<^sub>B _\") abbreviation (input) unit\<^sub>0 where "unit\<^sub>0 \ e" definition unit\<^sub>1 where "unit\<^sub>1 f = (e (trg\<^sub>B f) \\<^sub>B \\<^sub>B[f]) \\<^sub>B \\<^sub>B[e (trg\<^sub>B f), f, src\<^sub>B f] \\<^sub>B ((e (trg\<^sub>B f) \\<^sub>B f) \\<^sub>B B.inv (\ (src\<^sub>B f))) \\<^sub>B \\<^sub>B[e (trg\<^sub>B f) \\<^sub>B f, d (src\<^sub>B f), e (src\<^sub>B f)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e (trg\<^sub>B f), f, d (src\<^sub>B f)] \\<^sub>B e (src\<^sub>B f))" + abbreviation (input) \\<^sub>0 + where "\\<^sub>0 \ unit\<^sub>0" + abbreviation (input) \\<^sub>1 where "\\<^sub>1 \ unit\<^sub>1" lemma unit\<^sub>1_in_hom [intro]: assumes "B.ide f" shows "\\\<^sub>1 f : src\<^sub>B f \\<^sub>B P\<^sub>0 (trg\<^sub>B f)\" and "\\\<^sub>1 f : (e (trg\<^sub>B f) \\<^sub>B f \\<^sub>B d (src\<^sub>B f)) \\<^sub>B e (src\<^sub>B f) \\<^sub>B e (trg\<^sub>B f) \\<^sub>B f\" proof - show "\\\<^sub>1 f : (e (trg\<^sub>B f) \\<^sub>B f \\<^sub>B d (src\<^sub>B f)) \\<^sub>B e (src\<^sub>B f) \\<^sub>B e (trg\<^sub>B f) \\<^sub>B f\" using assms unit\<^sub>1_def by force thus "\\\<^sub>1 f : src\<^sub>B f \\<^sub>B P\<^sub>0 (trg\<^sub>B f)\" using assms B.vconn_implies_hpar(1-2) by auto qed lemma unit\<^sub>1_simps [simp]: assumes "B.ide f" shows "B.arr (\\<^sub>1 f)" and "src\<^sub>B (\\<^sub>1 f) = src\<^sub>B f" and "trg\<^sub>B (\\<^sub>1 f) = P\<^sub>0 (trg\<^sub>B f)" and "B.dom (\\<^sub>1 f) = (e (trg\<^sub>B f) \\<^sub>B f \\<^sub>B d (src\<^sub>B f)) \\<^sub>B e (src\<^sub>B f)" and "B.cod (\\<^sub>1 f) = e (trg\<^sub>B f) \\<^sub>B f" and "B.iso (\\<^sub>1 f)" using assms unit\<^sub>1_in_hom apply auto unfolding unit\<^sub>1_def - by (intro B.isos_compose) auto (* 15 sec *) + by (intro B.isos_compose) auto (* 10 sec *) lemma unit\<^sub>1_in_hom\<^sub>S [intro]: assumes "ide f" shows "\\\<^sub>1 f : src f \ P\<^sub>0 (trg f)\" and "\\\<^sub>1 f : PoE.map f \ e (src f) \ e (trg f) \ I\<^sub>S.map f\" using assms ide_char arr_char P\<^sub>0_props(1,6) P.preserves_ide hcomp_def src_def trg_def P_def emb.map_def equivalence_data_simps(2) in_hom_char by auto lemma unit\<^sub>1_simps\<^sub>S [simp]: assumes "ide f" shows "arr (\\<^sub>1 f)" and "src (\\<^sub>1 f) = src f" and "trg (\\<^sub>1 f) = P\<^sub>0 (trg f)" and "dom (\\<^sub>1 f) = PoE.map f \ e (src f)" and "cod (\\<^sub>1 f) = e (trg f) \ I\<^sub>S.map f" and "iso (\\<^sub>1 f)" using assms unit\<^sub>1_in_hom\<^sub>S iso_char ide_char arr_char P.as_nat_iso.components_are_iso by auto sublocale unit: pseudonatural_equivalence comp hcomp \ \\<^sub>B src trg comp hcomp \ \\<^sub>B src trg I\<^sub>S.map I\<^sub>S.cmp \P \ E\ PoE.cmp unit\<^sub>0 unit\<^sub>1 proof show "\a. obj a \ ide (e a)" using obj_char ide_char arr_char P\<^sub>0_props(1) by force show "\a. obj a \ \e a : src (I\<^sub>S.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 \ \\<^sub>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 \ \\\<^sub>1 f : PoE.map f \ e (src f) \ e (trg f) \ I\<^sub>S.map f\" using unit\<^sub>1_in_hom\<^sub>S(2) by simp show "\f. ide f \ iso (\\<^sub>1 f)" by simp show *: "\\. arr \ \ \\<^sub>1 (cod \) \ (PoE.map \ \ e (src \)) = (e (trg \) \ I\<^sub>S.map \) \ \\<^sub>1 (dom \)" proof - fix \ assume \: "arr \" have 1: "B.arr \" using \ arr_char by simp let ?f = "B.dom \" let ?g = "B.cod \" let ?a = "src\<^sub>B \" let ?b = "trg\<^sub>B \" have "\\<^sub>1 (cod \) \ (PoE.map \ \ e (src \)) = ((e ?b \\<^sub>B \\<^sub>B[?g]) \\<^sub>B \\<^sub>B[e ?b, ?g, ?a] \\<^sub>B ((e ?b \\<^sub>B ?g) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?b \\<^sub>B ?g, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?b, ?g, d ?a] \\<^sub>B e ?a)) \\<^sub>B ((e ?b \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B e ?a)" unfolding unit\<^sub>1_def P_def emb.map_def hcomp_def src_def using \ comp_char cod_simp arr_char P\<^sub>0_props [of ?a] P\<^sub>0_props [of ?b] - by auto (* 7 sec *) + by auto (* 10 sec *) also have "... = (e ?b \\<^sub>B \\<^sub>B[?g]) \\<^sub>B \\<^sub>B[e ?b, ?g, ?a] \\<^sub>B ((e ?b \\<^sub>B ?g) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?b \\<^sub>B ?g, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?b, ?g, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?b \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B e ?a)" using 1 B.comp_assoc by simp also have "... = (e ?b \\<^sub>B \\<^sub>B[?g]) \\<^sub>B \\<^sub>B[e ?b, ?g, ?a] \\<^sub>B ((e ?b \\<^sub>B ?g) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?b \\<^sub>B ?g, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?b, ?g, d ?a] \\<^sub>B (e ?b \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B e ?a \\<^sub>B e ?a)" using 1 B.interchange [of "\\<^sub>B\<^sup>-\<^sup>1[e ?b, ?g, d ?a]" "e ?b \\<^sub>B \ \\<^sub>B d ?a" "e ?a" "e ?a"] by simp also have "... = (e ?b \\<^sub>B \\<^sub>B[?g]) \\<^sub>B \\<^sub>B[e ?b, ?g, ?a] \\<^sub>B ((e ?b \\<^sub>B ?g) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?b \\<^sub>B ?g, d ?a, e ?a] \\<^sub>B (((e ?b \\<^sub>B \) \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \\<^sub>B e ?a \\<^sub>B e ?a)" using 1 B.assoc'_naturality [of "e ?b" \ "d ?a"] by simp also have "... = (e ?b \\<^sub>B \\<^sub>B[?g]) \\<^sub>B \\<^sub>B[e ?b, ?g, ?a] \\<^sub>B ((e ?b \\<^sub>B ?g) \\<^sub>B B.inv (\ ?a)) \\<^sub>B (\\<^sub>B[e ?b \\<^sub>B ?g, d ?a, e ?a] \\<^sub>B (((e ?b \\<^sub>B \) \\<^sub>B d ?a) \\<^sub>B e ?a)) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \\<^sub>B e ?a)" using 1 B.comp_assoc B.interchange [of "(e ?b \\<^sub>B \) \\<^sub>B d ?a" "\\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a]" "e ?a" "e ?a" ] by simp also have "... = (e ?b \\<^sub>B \\<^sub>B[?g]) \\<^sub>B \\<^sub>B[e ?b, ?g, ?a] \\<^sub>B (((e ?b \\<^sub>B ?g) \\<^sub>B B.inv (\ ?a)) \\<^sub>B ((e ?b \\<^sub>B \) \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B \\<^sub>B[e ?b \\<^sub>B ?f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \\<^sub>B e ?a)" using 1 B.assoc_naturality [of "e ?b \\<^sub>B \" "d ?a" "e ?a"] B.comp_assoc by simp also have "... = (e ?b \\<^sub>B \\<^sub>B[?g]) \\<^sub>B \\<^sub>B[e ?b, ?g, ?a] \\<^sub>B ((e ?b \\<^sub>B \) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?b \\<^sub>B ?f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \\<^sub>B e ?a)" proof - have "((e ?b \\<^sub>B ?g) \\<^sub>B B.inv (\ ?a)) \\<^sub>B ((e ?b \\<^sub>B \) \\<^sub>B d ?a \\<^sub>B e ?a) = (e ?b \\<^sub>B ?g) \\<^sub>B (e ?b \\<^sub>B \) \\<^sub>B B.inv (\ ?a) \\<^sub>B (d ?a \\<^sub>B e ?a)" using 1 B.interchange by simp also have "... = (e ?b \\<^sub>B \) \\<^sub>B B.inv (\ ?a)" using 1 B.comp_arr_dom B.comp_cod_arr by simp finally have "((e ?b \\<^sub>B ?g) \\<^sub>B B.inv (\ ?a)) \\<^sub>B ((e ?b \\<^sub>B \) \\<^sub>B d ?a \\<^sub>B e ?a) = (e ?b \\<^sub>B \) \\<^sub>B B.inv (\ ?a)" by simp thus ?thesis by simp qed also have "... = (e ?b \\<^sub>B \\<^sub>B[?g]) \\<^sub>B (\\<^sub>B[e ?b, ?g, ?a] \\<^sub>B ((e ?b \\<^sub>B \) \\<^sub>B ?a)) \\<^sub>B ((e ?b \\<^sub>B ?f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?b \\<^sub>B ?f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \\<^sub>B e ?a)" proof - have "(e ?b \\<^sub>B \) \\<^sub>B B.inv (\ ?a) = (e ?b \\<^sub>B \) \\<^sub>B (e ?b \\<^sub>B ?f) \\<^sub>B ?a \\<^sub>B B.inv (\ ?a)" using 1 B.comp_arr_dom B.comp_cod_arr by simp also have "... = ((e ?b \\<^sub>B \) \\<^sub>B ?a) \\<^sub>B ((e ?b \\<^sub>B ?f) \\<^sub>B B.inv (\ ?a))" using 1 B.interchange by simp finally have "(e ?b \\<^sub>B \) \\<^sub>B B.inv (\ ?a) = ((e ?b \\<^sub>B \) \\<^sub>B ?a) \\<^sub>B ((e ?b \\<^sub>B ?f) \\<^sub>B B.inv (\ ?a))" by blast thus ?thesis using B.comp_assoc by simp qed also have "... = ((e ?b \\<^sub>B \\<^sub>B[?g]) \\<^sub>B (e ?b \\<^sub>B \ \\<^sub>B ?a)) \\<^sub>B \\<^sub>B[e ?b, ?f, ?a] \\<^sub>B ((e ?b \\<^sub>B ?f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?b \\<^sub>B ?f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \\<^sub>B e ?a)" using 1 B.assoc_naturality [of "e ?b" \ "?a"] B.comp_assoc by simp also have "... = (e ?b \\<^sub>B \) \\<^sub>B (e ?b \\<^sub>B \\<^sub>B[?f]) \\<^sub>B \\<^sub>B[e ?b, ?f, ?a] \\<^sub>B ((e ?b \\<^sub>B ?f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?b \\<^sub>B ?f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?b, ?f, d ?a] \\<^sub>B e ?a)" proof - have "(e ?b \\<^sub>B \\<^sub>B[?g]) \\<^sub>B (e ?b \\<^sub>B \ \\<^sub>B ?a) = e ?b \\<^sub>B e ?b \\<^sub>B \\<^sub>B[?g] \\<^sub>B (\ \\<^sub>B ?a)" using 1 B.interchange [of "e ?b" "e ?b" "\\<^sub>B[?g]" "\ \\<^sub>B ?a"] by simp also have "... = e ?b \\<^sub>B e ?b \\<^sub>B \ \\<^sub>B \\<^sub>B[?f]" using 1 B.runit_naturality by simp also have "... = (e ?b \\<^sub>B \) \\<^sub>B (e ?b \\<^sub>B \\<^sub>B[?f])" using 1 B.interchange [of "e ?b" "e ?b" \ "\\<^sub>B[?f]"] by simp finally have "(e ?b \\<^sub>B \\<^sub>B[?g]) \\<^sub>B (e ?b \\<^sub>B \ \\<^sub>B ?a) = (e ?b \\<^sub>B \) \\<^sub>B (e ?b \\<^sub>B \\<^sub>B[?f])" by blast thus ?thesis using B.comp_assoc by simp qed also have "... = (e (trg \) \ I\<^sub>S.map \) \ \\<^sub>1 (dom \)" proof - have "arr ((e (trg \) \ I\<^sub>S.map \) \ \\<^sub>1 (dom \))" using \ by simp hence 2: "arr ((e ?b \\<^sub>B \) \ ((e ?b \\<^sub>B \\<^sub>B[?f]) \\<^sub>B \\<^sub>B[e ?b, ?f, ?a] \\<^sub>B ((e ?b \\<^sub>B ?f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?b \\<^sub>B ?f, d ?a, e ?a] \\<^sub>B (B.inv \\<^sub>B[e ?b, ?f, d ?a] \\<^sub>B e ?a)))" unfolding unit\<^sub>1_def using \ 1 arr_char P\<^sub>0_props trg_def hcomp_def dom_char by simp show ?thesis unfolding unit\<^sub>1_def using \ 1 arr_char P\<^sub>0_props trg_def hcomp_def dom_char apply simp using 2 comp_eqI by blast qed finally show "\\<^sub>1 (cod \) \ (PoE.map \ \ e (src \)) = (e (trg \) \ I\<^sub>S.map \) \ \\<^sub>1 (dom \)" by simp qed show "\a. obj a \ (e a \ I\<^sub>S.unit a) \ \\<^sup>-\<^sup>1[e a] \ \[e a] = unit\<^sub>1 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 V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e a\ \d a\ \\ a\ \\ a\ using 0 chosen_adjoint_equivalence(1) [of a] by simp have src: "src\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) = P\<^sub>0 a" using 0 B.src_vcomp [of "e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]" "B.inv (\ a)"] emb.map_def by simp have trg: "trg\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) = P\<^sub>0 a" using 0 B.trg_vcomp [of "e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]" "B.inv (\ a)"] emb.map_def by simp have 1: "arr a \ src\<^sub>B a = a" using a src_def by auto have 2: "seq (P a) ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a))" proof - have "B.arr (P a \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>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 \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a))" apply (unfold arr_char) by (simp add: 0 P\<^sub>0_props(6) src trg) ultimately show ?thesis using seq_char by simp qed have 3: "obj (P\<^sub>0 a)" using 0 P\<^sub>0_props arr_char obj_char by blast have 4: "B.seq (P a) ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a))" using 2 seq_char by simp have "\\<^sub>1 a \ (PoE.unit a \ e a) = \\<^sub>1 a \ (P a \ (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) \\<^sub>B e a)" proof - have "\\<^sub>1 a \ (PoE.unit a \ e a) = \\<^sub>1 a \ (PoE.unit a \\<^sub>B e a)" using 0 a hcomp_eqI by force moreover have "PoE.unit a = P a \ (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>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 "... = \\<^sub>1 a \ (P a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) \\<^sub>B e a)" using 2 comp_simp by force also have "... = \\<^sub>1 a \\<^sub>B (P a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) \\<^sub>B e a)" proof - have "arr (\\<^sub>1 a)" using a by auto moreover have "arr (P a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) \\<^sub>B e a)" by (metis (no_types, lifting) a 0 2 B.src_vcomp B.vseq_implies_hpar(1) arr_char arr_compI equivalence_data_in_hom(3) equivalence_data_simps\<^sub>B(6) hcomp_closed in_hom_char src) moreover have "B.seq (\\<^sub>1 a) (P a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) \\<^sub>B e a)" using 0 3 P_def obj_char - by (elim B.objE, intro B.seqI) auto (* 6 sec *) + by (elim B.objE, intro B.seqI) auto (* 10 sec *) ultimately show ?thesis using 2 comp_char by meson qed also have "... = (e a \\<^sub>B \\<^sub>B[a]) \\<^sub>B \\<^sub>B[e a, a, a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B B.inv (\ a)) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, e a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B e a) \\<^sub>B (((e a \\<^sub>B a \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])) \\<^sub>B B.inv (\ a) \\<^sub>B e a)" using 0 1 unit\<^sub>1_def P_def emb.map\<^sub>0_def emb.map_def B.comp_assoc B.obj_def' by simp also have "... = (e a \\<^sub>B \\<^sub>B[a]) \\<^sub>B \\<^sub>B[e a, a, a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B B.inv (\ a)) \\<^sub>B \\<^sub>B[e a \\<^sub>B a, d a, e a] \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B e a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) \\<^sub>B e a))" proof - have "(e a \\<^sub>B a \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) = e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]" using 0 B.comp_cod_arr by simp thus ?thesis by simp qed also have "... = (e a \\<^sub>B \\<^sub>B[a]) \\<^sub>B \\<^sub>B[e a, a, a] \\<^sub>B ((e a \\<^sub>B a) \\<^sub>B B.inv (\ a)) \\<^sub>B (\\<^sub>B[e a \\<^sub>B a, d a, e a] \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a) \\<^sub>B e a)) \\<^sub>B (B.inv (\ a) \\<^sub>B e a)" proof - have "(\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B e a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) \\<^sub>B e a) = (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B e a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B e a) \\<^sub>B (B.inv (\ a) \\<^sub>B e a)" using 0 B.whisker_right by simp also have "... = ((\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B e a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B e a)) \\<^sub>B (B.inv (\ a) \\<^sub>B e a)" using B.comp_assoc by simp also have "... = (\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B e a) \\<^sub>B (B.inv (\ a) \\<^sub>B e a)" proof - have "B.seq \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>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 "... = ((\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a) \\<^sub>B e a) \\<^sub>B (B.inv (\ a) \\<^sub>B e a)" proof - have "\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) = \\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a" proof - have "e a \\<^sub>B \\<^sub>B[d a] = (\\<^sub>B[e a] \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a]" using 0 1 B.triangle' [of "e a" "d a"] by simp moreover have "B.hseq (e a) \\<^sub>B[d a]" using 0 by auto moreover have "src\<^sub>B (e a) = a" using 0 by simp ultimately have "(\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a) \\<^sub>B (e a \\<^sub>B \\<^sub>B[d a]) = \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a]" using 0 B.invert_side_of_triangle(1) [of "e a \\<^sub>B \\<^sub>B[d a]" "\\<^sub>B[e a] \\<^sub>B d a" "\\<^sub>B\<^sup>-\<^sup>1[e a, src\<^sub>B (e a), d a]"] by simp moreover have "B.arr \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a]" using 0 B.assoc'_in_hom [of "e a" a "d a"] by fastforce ultimately have "\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a = \\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])" using 0 B.invert_side_of_triangle(2) [of "\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a]" "\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a" "e a \\<^sub>B \\<^sub>B[d a]"] by simp thus ?thesis by simp qed thus ?thesis by simp qed finally have "(\\<^sub>B\<^sup>-\<^sup>1[e a, a, d a] \\<^sub>B e a) \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) \\<^sub>B e a) = ((\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a) \\<^sub>B e a) \\<^sub>B (B.inv (\ a) \\<^sub>B e a)" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e a \\<^sub>B \\<^sub>B[a]) \\<^sub>B \\<^sub>B[e a, a, a] \\<^sub>B (((e a \\<^sub>B a) \\<^sub>B B.inv (\ a)) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a \\<^sub>B e a)) \\<^sub>B \\<^sub>B[e a, d a, e a] \\<^sub>B (B.inv (\ a) \\<^sub>B e a)" using 0 B.assoc_naturality [of "\\<^sub>B\<^sup>-\<^sup>1[e a]" "d a" "e a"] B.comp_assoc by simp also have "... = (e a \\<^sub>B \\<^sub>B[a]) \\<^sub>B \\<^sub>B[e a, a, a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B a) \\<^sub>B ((e a \\<^sub>B B.inv (\ a)) \\<^sub>B \\<^sub>B[e a, d a, e a] \\<^sub>B (B.inv (\ a) \\<^sub>B e a))" proof - have "((e a \\<^sub>B a) \\<^sub>B B.inv (\ a)) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a \\<^sub>B e a) = (e a \\<^sub>B a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B B.inv (\ a) \\<^sub>B (d a \\<^sub>B e a)" using 0 B.interchange [of "e a \\<^sub>B a" "\\<^sub>B\<^sup>-\<^sup>1[e a]" "B.inv (\ a)" "d a \\<^sub>B e a"] by force also have "... = \\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B e a \\<^sub>B a \\<^sub>B B.inv (\ a)" using 0 B.comp_arr_dom B.comp_cod_arr by simp also have "... = (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B a) \\<^sub>B (e a \\<^sub>B B.inv (\ a))" using 0 B.interchange [of "\\<^sub>B\<^sup>-\<^sup>1[e a]" "e a" a "B.inv (\ a)"] by fastforce ultimately have "((e a \\<^sub>B a) \\<^sub>B B.inv (\ a)) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B d a \\<^sub>B e a) = (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B a) \\<^sub>B (e a \\<^sub>B B.inv (\ a))" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = ((e a \\<^sub>B \\<^sub>B[a]) \\<^sub>B \\<^sub>B[e a, a, a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B \\<^sub>B[e a]" by (metis B.comp_assoc adjoint_equivalence_in_bicategory_def adjunction_in_bicategory.triangle_right dual_adjoint_equivalence) also have "... = \\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B \\<^sub>B[e a]" proof - have "(e a \\<^sub>B \\<^sub>B[a]) \\<^sub>B \\<^sub>B[e a, a, a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B a) = e a \\<^sub>B a" proof - have "(e a \\<^sub>B \\<^sub>B[a]) \\<^sub>B \\<^sub>B[e a, a, a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B a) = ((e a \\<^sub>B \\<^sub>B[a]) \\<^sub>B \\<^sub>B[e a, a, a]) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B a)" using B.comp_assoc by simp also have "... = (\\<^sub>B[e a] \\<^sub>B a) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B a)" using 0 B.runit_char(2) [of "e a"] B.unitor_coincidence by simp also have "... = e a \\<^sub>B a" using 0 B.whisker_right [of a "\\<^sub>B[e a]" "\\<^sub>B\<^sup>-\<^sup>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 \ I\<^sub>S.unit a) \ \\<^sup>-\<^sup>1[e a] \ \[e a]" proof - have "(e a \ I\<^sub>S.unit a) \ \\<^sup>-\<^sup>1[e a] \ \[e a] = (e a \\<^sub>B a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B \\<^sub>B[e a]" proof - have 1: "arr \\<^sub>B[e a] \ arr \\<^sub>B\<^sup>-\<^sup>1[e a]" using a lunit_simp runit'_simp lunit_simps runit'_simps by simp moreover have "arr (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B \\<^sub>B[e a])" using 1 arr_char dom_char cod_char comp_char by auto moreover have 2: "hseq (e a) (I\<^sub>S.unit a)" using a I\<^sub>S.unit_char by (intro hseqI') auto moreover have "B.seq (e a \ I\<^sub>S.unit a) (\\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B \\<^sub>B[e a])" using 0 2 hcomp_char arr_char I\<^sub>S.unit_char' apply (intro B.seqI) apply auto[4] using a by force ultimately show ?thesis using a comp_char I\<^sub>S.unit_char' lunit_simp runit'_simp hcomp_char by auto (* 4 sec *) qed also have "... = \\<^sub>B\<^sup>-\<^sup>1[e a] \\<^sub>B \\<^sub>B[e a]" using 0 B.comp_cod_arr by simp finally show ?thesis by simp qed finally show "(e a \ I\<^sub>S.unit a) \ \\<^sup>-\<^sup>1[e a] \ \[e a] = \\<^sub>1 a \ (PoE.unit a \ e a)" by argo qed show "\f g. \ide f; ide g; src g = trg f\ \ (e (trg g) \ I\<^sub>S.cmp (g, f)) \ \[e (trg g), I\<^sub>S.map g, I\<^sub>S.map f] \ (\\<^sub>1 g \ I\<^sub>S.map f) \ inv \[PoE.map g, e (src g), I\<^sub>S.map f] \ (PoE.map g \ \\<^sub>1 f) \ \[PoE.map g, PoE.map f, e (src f)] = \\<^sub>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 = "src\<^sub>B f" let ?b = "src\<^sub>B g" let ?c = "trg\<^sub>B g" have hseq: "ide f \ ide g \ arr f \ arr g \ src g = trg f" using f g fg ide_char by auto have hseq\<^sub>B: "B.ide f \ B.ide g \ B.arr f \ B.arr g \ src\<^sub>B g = trg\<^sub>B f" using f g fg ide_char 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) \ I\<^sub>S.cmp (g, f)) \ \[e (trg g), I\<^sub>S.map g, I\<^sub>S.map f] \ (\\<^sub>1 g \ I\<^sub>S.map f) \ inv \[PoE.map g, e (src g), I\<^sub>S.map f] \ (PoE.map g \ \\<^sub>1 f) \ \[PoE.map g, PoE.map f, e (src f)] = (e ?c \ I\<^sub>S.cmp (g, f)) \ \[e ?c, I\<^sub>S.map g, I\<^sub>S.map f] \ (\\<^sub>1 g \ I\<^sub>S.map f) \ inv \[PoE.map g, e ?b, I\<^sub>S.map f] \ (PoE.map g \ \\<^sub>1 f) \ \[PoE.map g, PoE.map f, e ?a]" using 1 by simp also have "... = (e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" proof - have "(e ?c \ I\<^sub>S.cmp (g, f)) \ \[e ?c, I\<^sub>S.map g, I\<^sub>S.map f] \ (\\<^sub>1 g \ I\<^sub>S.map f) \ inv \[PoE.map g, e ?b, I\<^sub>S.map f] \ (PoE.map g \ unit\<^sub>1 f) \ \[PoE.map g, PoE.map f, e ?a] = (e ?c \ g \ f) \ \[e ?c, g, f] \ (\\<^sub>1 g \ f) \ inv \[P (E g), e ?b, f] \ (P (E g) \ \\<^sub>1 f) \ \[P (E g), P (E f), e ?a]" using f g fg by simp also have "... = (e ?c \ g \ f) \\<^sub>B \[e ?c, g, f] \\<^sub>B (\\<^sub>1 g \ f) \\<^sub>B inv \[P (E g), e ?b, f] \\<^sub>B (P (E g) \ \\<^sub>1 f) \\<^sub>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_char VV.arr_char src_def obj_src assoc_in_hom by force show "\P (E g) \ \\<^sub>1 f : P (E g) \ P (E f) \ e ?a \ P (E g) \ e ?b \ f\" using f g fg 1 unit\<^sub>1_in_hom\<^sub>S [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_char VV.arr_char assoc_in_hom by simp moreover have "iso \[P (E g), e ?b, f]" using f g fg 1 VVV.arr_char VV.arr_char iso_assoc by simp ultimately show ?thesis using inv_in_hom by simp qed show "\unit\<^sub>1 g \ f : (P (E g) \ e ?b) \ f \ (e ?c \ g) \ f\" using f g fg 1 unit\<^sub>1_in_hom\<^sub>S [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_char 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_char VV.arr_char 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) \\<^sub>B \[e ?c, g, f] \\<^sub>B (\\<^sub>1 g \ f) \\<^sub>B B.inv \[P (E g), e ?b, f] \\<^sub>B (P (E g) \ \\<^sub>1 f) \\<^sub>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 hseq\<^sub>B src_def trg_def emb.map\<^sub>0_def emb.map_def arr_char obj_char P\<^sub>0_props(6) by (intro iso_assoc) auto thus ?thesis using f g fg inv_char [of "\[P (E g), e ?b, f]"] by simp qed also have "... = (e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B \[e ?c, g, f] \\<^sub>B (\\<^sub>1 g \\<^sub>B f) \\<^sub>B B.inv \[P (E g), e ?b, f] \\<^sub>B (P (E g) \\<^sub>B \\<^sub>1 f) \\<^sub>B \[P (E g), P (E f), e ?a]" proof - have "arr (unit\<^sub>0 ?c)" using hseq obj_trg trg_def by auto moreover have "arr (g \\<^sub>B f)" using hseq hseq\<^sub>B arr_char by auto ultimately show ?thesis unfolding hcomp_def using f g fg hseq\<^sub>B src_def trg_def P\<^sub>0_props(1) emb.map\<^sub>0_def emb.map_def by simp qed also have "... = (e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g, f] \\<^sub>B (\\<^sub>1 g \\<^sub>B f) \\<^sub>B B.inv \\<^sub>B[P (E g), e ?b, f] \\<^sub>B (P (E g) \\<^sub>B \\<^sub>1 f) \\<^sub>B \\<^sub>B[P (E g), P (E f), e ?a]" proof - have "arr (e ?a) \ arr (e ?b) \ arr (e ?c)" using hseq hseq\<^sub>B src_def trg_def obj_char src.preserves_arr trg.preserves_arr by auto thus ?thesis using f g fg VVV.arr_char VV.arr_char src_def trg_def hseq\<^sub>B emb.map\<^sub>0_def emb.map_def by simp qed also have "... = (e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g, f] \\<^sub>B (\\<^sub>1 g \\<^sub>B f) \\<^sub>B (B.inv \\<^sub>B[P g, e ?b, f] \\<^sub>B (P g \\<^sub>B \\<^sub>1 f) \\<^sub>B \\<^sub>B[P g, P f, e ?a])" using f g fg 1 emb.map_def by simp also have "... = (e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g, f] \\<^sub>B (\\<^sub>1 g \\<^sub>B f) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \\<^sub>B (P g \\<^sub>B \\<^sub>1 f) \\<^sub>B \\<^sub>B[P g, P f, e ?a])" using hseq\<^sub>B P.preserves_ide ide_char B.assoc'_eq_inv_assoc [of "P g" "e ?b" f] by simp also have "... = (e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g, f] \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B[g]) \\<^sub>B \\<^sub>B[e ?c, g, ?b] \\<^sub>B ((e ?c \\<^sub>B g) \\<^sub>B B.inv (\ ?b)) \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \\<^sub>B (P g \\<^sub>B (e ?b \\<^sub>B \\<^sub>B[f]) \\<^sub>B \\<^sub>B[e ?b, f, ?a] \\<^sub>B ((e ?b \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a)) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" unfolding unit\<^sub>1_def using hseq\<^sub>B B.comp_assoc by auto also have "... = ((e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g, f]) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B[g]) \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c, g, ?b] \\<^sub>B f) \\<^sub>B (((e ?c \\<^sub>B g) \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \\<^sub>B (P g \\<^sub>B e ?b \\<^sub>B \\<^sub>B[f]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, ?a]) \\<^sub>B (P g \\<^sub>B (e ?b \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" using hseq\<^sub>B ide_char P.preserves_ide B.whisker_left B.whisker_right B.comp_assoc by auto also have "... = \\<^sub>B[e ?c, g, f] \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B[g]) \\<^sub>B f) \\<^sub>B ((\\<^sub>B[e ?c, g, ?b] \\<^sub>B f) \\<^sub>B (((e ?c \\<^sub>B g) \\<^sub>B B.inv (\ ?b)) \\<^sub>B f)) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \\<^sub>B (P g \\<^sub>B e ?b \\<^sub>B \\<^sub>B[f]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, ?a]) \\<^sub>B (P g \\<^sub>B (e ?b \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" proof - have "(e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g, f] = \\<^sub>B[e ?c, g, f]" using hseq\<^sub>B B.comp_cod_arr by auto thus ?thesis using B.comp_assoc by simp qed also have "... = \\<^sub>B[e ?c, g, f] \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B[g]) \\<^sub>B f) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f)) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \\<^sub>B (P g \\<^sub>B e ?b \\<^sub>B \\<^sub>B[f]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, ?a]) \\<^sub>B (P g \\<^sub>B (e ?b \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" proof - have "(\\<^sub>B[e ?c, g, ?b] \\<^sub>B f) \\<^sub>B (((e ?c \\<^sub>B g) \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) = ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f)" proof - have "(\\<^sub>B[e ?c, g, ?b] \\<^sub>B f) \\<^sub>B (((e ?c \\<^sub>B g) \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) = \\<^sub>B[e ?c, g, ?b] \\<^sub>B ((e ?c \\<^sub>B g) \\<^sub>B B.inv (\ ?b)) \\<^sub>B f" using hseq\<^sub>B B.whisker_right by auto also have "... = (e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b)) \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f" using hseq\<^sub>B B.assoc_naturality [of "e ?c" g "B.inv (\ ?b)"] by simp also have "... = ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f)" using hseq\<^sub>B B.whisker_right by auto finally show ?thesis by simp qed thus ?thesis using B.comp_assoc by simp qed also have "... = (\\<^sub>B[e ?c, g, f] \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f)) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \\<^sub>B (P g \\<^sub>B e ?b \\<^sub>B \\<^sub>B[f]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, ?a]) \\<^sub>B (P g \\<^sub>B (e ?b \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" using hseq\<^sub>B B.whisker_left B.whisker_right B.comp_assoc by simp also have "... = (e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \\<^sub>B (P g \\<^sub>B e ?b \\<^sub>B \\<^sub>B[f]) \\<^sub>B ((P g \\<^sub>B \\<^sub>B[e ?b, f, ?a]) \\<^sub>B (P g \\<^sub>B (e ?b \\<^sub>B f) \\<^sub>B B.inv (\ ?a))) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" using hseq\<^sub>B B.comp_assoc B.assoc_naturality [of "e ?c" "\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))" f] by fastforce also have "... = (e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \\<^sub>B ((P g \\<^sub>B e ?b \\<^sub>B \\<^sub>B[f]) \\<^sub>B (P g \\<^sub>B e ?b \\<^sub>B f \\<^sub>B B.inv (\ ?a))) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" proof - have "(P g \\<^sub>B \\<^sub>B[e ?b, f, ?a]) \\<^sub>B (P g \\<^sub>B (e ?b \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) = (P g \\<^sub>B e ?b \\<^sub>B f \\<^sub>B B.inv (\ ?a)) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a])" proof - have "(P g \\<^sub>B \\<^sub>B[e ?b, f, ?a]) \\<^sub>B (P g \\<^sub>B (e ?b \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) = P g \\<^sub>B \\<^sub>B[e ?b, f, ?a] \\<^sub>B ((e ?b \\<^sub>B f) \\<^sub>B B.inv (\ ?a))" using hseq\<^sub>B ide_char P.preserves_ide B.whisker_left by auto also have "... = P g \\<^sub>B (e ?b \\<^sub>B f \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]" using hseq\<^sub>B B.assoc_naturality [of "e ?b" f "B.inv (\ ?a)"] by auto also have "... = (P g \\<^sub>B e ?b \\<^sub>B f \\<^sub>B B.inv (\ ?a)) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a])" using hseq\<^sub>B ide_char 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 \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f] \\<^sub>B (P g \\<^sub>B e ?b \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" using hseq\<^sub>B ide_char P.preserves_ide B.whisker_left B.comp_assoc by auto also have "... = (e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f) \\<^sub>B (((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f) \\<^sub>B ((P g \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" using hseq\<^sub>B B.comp_assoc B.assoc'_naturality [of "P g" "e ?b" "\\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))"] by simp also have "... = (e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f) \\<^sub>B ((\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" proof - have "((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f) \\<^sub>B ((P g \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) = (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))" proof - have "((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f) \\<^sub>B ((P g \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) = (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B (P g \\<^sub>B e ?b) \\<^sub>B (f \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))" proof - have "B.seq (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) (P g \\<^sub>B e ?b)" using hseq\<^sub>B P.preserves_ide P_def src_def trg_def by auto moreover have "B.seq f (\\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))" using hseq\<^sub>B by simp ultimately show ?thesis using B.interchange [of "\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b" "P g \\<^sub>B e ?b" f "\\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))"] by presburger qed also have "... = (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B d ?b) \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))" using hseq\<^sub>B B.whisker_right P_def by auto also have "... = (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))" using hseq\<^sub>B 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 \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B ((\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" proof - have "(\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) = (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a))" proof - have "(\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) = \\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))" using hseq\<^sub>B B.interchange by simp also have "... = \\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B (\\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a)" using hseq\<^sub>B B.comp_cod_arr B.comp_arr_dom by simp also have "... = (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a))" using hseq\<^sub>B B.interchange by auto finally show ?thesis by simp qed thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" proof - have "(\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) = \\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))" using hseq\<^sub>B B.interchange by simp also have "... = \\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B (\\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a)" using hseq\<^sub>B B.comp_cod_arr B.comp_arr_dom by simp also have "... = (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)" using hseq\<^sub>B B.interchange by auto finally have "(\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) = (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = ((e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B ((e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" proof - have "\\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) = \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B (\\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a))" using hseq\<^sub>B B.comp_arr_dom B.comp_cod_arr by simp also have "... = \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)" using hseq\<^sub>B B.interchange by simp also have "... = (\\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)" using B.comp_assoc by simp also have "... = (((e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a]) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)" using hseq\<^sub>B B.assoc_naturality [of "e ?c" "g \\<^sub>B d ?b \\<^sub>B e ?b" "\\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))"] by auto also have "... = ((e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)" using B.comp_assoc by simp finally have "\\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) = ((e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" proof - have "(e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B ((e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) = e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))" using hseq\<^sub>B B.whisker_left B.interchange by fastforce also have "... = e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))" using hseq\<^sub>B B.comp_arr_dom B.comp_cod_arr by auto finally have "(e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B ((e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) = e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))" by simp thus ?thesis by simp qed finally show ?thesis by blast qed also have "... = (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f]) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, ?a] \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" proof - (* Working from the above expression to the common form, as in the previous part. *) have "(e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f]) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, ?a] \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a) = (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f]) \\<^sub>B (\\<^sub>B[e ?c, g \\<^sub>B f, ?a] \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B B.inv (\ ?a))) \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" using B.comp_assoc by simp also have "... = ((e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f]) \\<^sub>B (e ?c \\<^sub>B (g \\<^sub>B f) \\<^sub>B B.inv (\ ?a))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" using hseq\<^sub>B B.assoc_naturality [of "e ?c" "g \\<^sub>B f" "B.inv (\ ?a)"] B.comp_assoc by simp also have "... = (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B B.inv (\ ?a))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B e ?a)) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" using hseq\<^sub>B B.whisker_left B.comp_assoc by auto also have "... = (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B B.inv (\ ?a))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a)) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" proof - have "((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B e ?a) = ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a)" using hseq\<^sub>B 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 \\<^sub>B \\<^sub>B[g \\<^sub>B f] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B B.inv (\ ?a))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B (((g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f \\<^sub>B d ?a))) \\<^sub>B e ?a)) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" proof - have "((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) = (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, ?b, f \\<^sub>B d ?a] \\<^sub>B (g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B e ?a" using hseq\<^sub>B B.whisker_right B.whisker_left by auto also have "... = (e ?c \\<^sub>B (((g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a" using hseq\<^sub>B B.assoc'_naturality [of g "B.inv (\ ?b)" "f \\<^sub>B d ?a"] by auto also have "... = ((e ?c \\<^sub>B (((g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f \\<^sub>B d ?a))) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a)" using hseq\<^sub>B B.whisker_right B.whisker_left by auto finally have "((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) = ((e ?c \\<^sub>B (((g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f \\<^sub>B d ?a))) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a)" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B B.inv (\ ?a))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a)) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" using hseq\<^sub>B B.whisker_right B.whisker_left B.comp_assoc by auto also have "... = (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B B.inv (\ ?a))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B (((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f) \\<^sub>B d ?a)) \\<^sub>B e ?a)) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" proof - have "((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) = (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a] \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B e ?a" using hseq\<^sub>B B.whisker_left B.whisker_right by auto also have "... = (e ?c \\<^sub>B (((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f) \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a]) \\<^sub>B e ?a" using hseq\<^sub>B B.assoc'_naturality [of "\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))" f "d ?a"] by auto also have "... = ((e ?c \\<^sub>B (((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f) \\<^sub>B d ?a)) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a]) \\<^sub>B e ?a)" using hseq\<^sub>B B.whisker_left B.whisker_right by auto finally have "((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) = ((e ?c \\<^sub>B (((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f) \\<^sub>B d ?a)) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a]) \\<^sub>B e ?a)" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B B.inv (\ ?a))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (((e ?c \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f)) \\<^sub>B d ?a) \\<^sub>B e ?a)) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \\<^sub>B (d ?b \\<^sub>B e ?b)) \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" proof - have "(\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B (((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f) \\<^sub>B d ?a)) \\<^sub>B e ?a) = \\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B (e ?c \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B d ?a)) \\<^sub>B e ?a" using hseq\<^sub>B B.whisker_right by auto also have "... = ((e ?c \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f)) \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \\<^sub>B (d ?b \\<^sub>B e ?b)) \\<^sub>B f, d ?a] \\<^sub>B e ?a" using hseq\<^sub>B B.assoc'_naturality [of "e ?c" "(\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f" "d ?a"] by fastforce also have "... = (((e ?c \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f)) \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \\<^sub>B (d ?b \\<^sub>B e ?b)) \\<^sub>B f, d ?a] \\<^sub>B e ?a)" using hseq\<^sub>B B.whisker_right by auto finally have "(\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B (((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f) \\<^sub>B d ?a)) \\<^sub>B e ?a) = (((e ?c \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f)) \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \\<^sub>B (d ?b \\<^sub>B e ?b)) \\<^sub>B f, d ?a] \\<^sub>B e ?a)" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B B.inv (\ ?a))) \\<^sub>B (\\<^sub>B[e ?c, g \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B ((e ?c \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f)) \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B \\<^sub>B[e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \\<^sub>B (d ?b \\<^sub>B e ?b)) \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" proof - have "\\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (((e ?c \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f)) \\<^sub>B d ?a) \\<^sub>B e ?a) = ((e ?c \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f)) \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a, e ?a]" using hseq\<^sub>B B.assoc_naturality [of "(e ?c \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f))" "d ?a" "e ?a"] by fastforce thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B B.inv (\ ?a))) \\<^sub>B (e ?c \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f) \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[e ?c, (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \\<^sub>B (d ?b \\<^sub>B e ?b)) \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" using hseq\<^sub>B B.comp_assoc B.assoc_naturality [of "e ?c" "(\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f" "d ?a \\<^sub>B e ?a"] by force also have "... = (e ?c \\<^sub>B (g \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B[g, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (e ?c \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))) \\<^sub>B f) \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B \\<^sub>B[e ?c, (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \\<^sub>B (d ?b \\<^sub>B e ?b)) \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" proof - have "e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) = (e ?c \\<^sub>B (g \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B (e ?c \\<^sub>B \\<^sub>B[g, f, d ?a \\<^sub>B e ?a])" using hseq\<^sub>B 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 \\<^sub>B (g \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B (e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a))) \\<^sub>B (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B \\<^sub>B[e ?c, (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \\<^sub>B (d ?b \\<^sub>B e ?b)) \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" proof - have "(e ?c \\<^sub>B \\<^sub>B[g, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B d ?a \\<^sub>B e ?a) = e ?c \\<^sub>B \\<^sub>B[g, f, d ?a \\<^sub>B e ?a] \\<^sub>B ((\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B d ?a \\<^sub>B e ?a)" using hseq\<^sub>B B.whisker_left by auto also have "... = e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a \\<^sub>B e ?a]" using hseq\<^sub>B B.assoc_naturality [of "\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b))" f "d ?a \\<^sub>B e ?a"] by auto also have "... = (e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a \\<^sub>B e ?a])" using hseq\<^sub>B B.whisker_left by auto finally have "(e ?c \\<^sub>B \\<^sub>B[g, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f) \\<^sub>B d ?a \\<^sub>B e ?a) = (e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a \\<^sub>B e ?a])" by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B \\<^sub>B[e ?c, (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \\<^sub>B (d ?b \\<^sub>B e ?b)) \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" proof - have "(e ?c \\<^sub>B (g \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B (e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)) = e ?c \\<^sub>B (g \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)" using hseq\<^sub>B B.whisker_left by auto also have "... = e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B (\\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a)" using hseq\<^sub>B B.interchange by auto also have "... = e ?c \\<^sub>B (g \\<^sub>B \\<^sub>B[g]) \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a)" using B.comp_assoc by simp also have "... = e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))" using hseq\<^sub>B B.comp_arr_dom B.comp_cod_arr by auto finally have "(e ?c \\<^sub>B (g \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a)))) \\<^sub>B (e ?c \\<^sub>B (\\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a)) = e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))" by simp thus ?thesis by simp qed also have "... = (e ?c \\<^sub>B \\<^sub>B[g] \\<^sub>B (g \\<^sub>B B.inv (\ ?b)) \\<^sub>B \\<^sub>B[f] \\<^sub>B (f \\<^sub>B B.inv (\ ?a))) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" proof - let ?LHS = "(e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B \\<^sub>B[e ?c, (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a \\<^sub>B e ?a] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B (g \\<^sub>B d ?b \\<^sub>B e ?b) \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, (g \\<^sub>B (d ?b \\<^sub>B e ?b)) \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b \\<^sub>B e ?b, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a])) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" let ?RHS = "\\<^sub>B[e ?c, g \\<^sub>B d ?b \\<^sub>B e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, e ?b] \\<^sub>B f \\<^sub>B d ?a \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B e ?b) \\<^sub>B (f \\<^sub>B d ?a \\<^sub>B e ?a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[P g, e ?b, f \\<^sub>B d ?a \\<^sub>B e ?a] \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b, f, d ?a \\<^sub>B e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B[e ?b \\<^sub>B f, d ?a, e ?a]) \\<^sub>B (P g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?b, f, d ?a] \\<^sub>B e ?a) \\<^sub>B \\<^sub>B[P g, P f, e ?a]" have "?LHS = ?RHS" proof - let ?LHSt = "(\<^bold>\e ?c\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>\e ?b\<^bold>\, \<^bold>\f\<^bold>\, \<^bold>\d ?a\<^bold>\ \<^bold>\ \<^bold>\e ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?c\<^bold>\, (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>\e ?b\<^bold>\) \<^bold>\ \<^bold>\f\<^bold>\, \<^bold>\d ?a\<^bold>\ \<^bold>\ \<^bold>\e ?a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?c\<^bold>\ \<^bold>\ (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>\e ?b\<^bold>\) \<^bold>\ \<^bold>\f\<^bold>\, \<^bold>\d ?a\<^bold>\, \<^bold>\e ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e ?c\<^bold>\, (\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>\e ?b\<^bold>\)) \<^bold>\ \<^bold>\f\<^bold>\, \<^bold>\d ?a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\e ?a\<^bold>\) \<^bold>\ ((\<^bold>\e ?c\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>\e ?b\<^bold>\, \<^bold>\f\<^bold>\, \<^bold>\d ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\e ?a\<^bold>\) \<^bold>\ (((\<^bold>\e ?c\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>\e ?b\<^bold>\, \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\\<^bold>])) \<^bold>\ \<^bold>\e ?a\<^bold>\) \<^bold>\ ((\<^bold>\e ?c\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\d ?b\<^bold>\, \<^bold>\e ?b\<^bold>\, \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\e ?a\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\e ?c\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>P\<^bold>R\<^bold>J f\<^bold>] \<^bold>\ \<^bold>\e ?a\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\e ?c\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\, \<^bold>P\<^bold>R\<^bold>J f\<^bold>] \<^bold>\ \<^bold>\e ?a\<^bold>\) \<^bold>\ ((\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e ?c\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\ \<^bold>] \<^bold>\ \<^bold>P\<^bold>R\<^bold>J f) \<^bold>\ \<^bold>\e ?a\<^bold>\)" let ?RHSt = "\<^bold>\\<^bold>[\<^bold>\e ?c\<^bold>\, \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>\e ?b\<^bold>\, \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\ \<^bold>\ \<^bold>\e ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\e ?c\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>\e ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\ \<^bold>\ \<^bold>\e ?a\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\e ?c\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\, \<^bold>\e ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\ \<^bold>\ \<^bold>\e ?a\<^bold>\) \<^bold>\ ((\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e ?c\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\e ?b\<^bold>\) \<^bold>\ ( \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\ \<^bold>\ \<^bold>\e ?a\<^bold>\)) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>P\<^bold>R\<^bold>J g, \<^bold>\e ?b\<^bold>\, \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\ \<^bold>\ \<^bold>\e ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>P\<^bold>R\<^bold>J g \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?b\<^bold>\, \<^bold>\f\<^bold>\, \<^bold>\d ?a\<^bold>\ \<^bold>\ \<^bold>\e ?a\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>P\<^bold>R\<^bold>J g \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?b\<^bold>\ \<^bold>\ \<^bold>\f\<^bold>\, \<^bold>\d ?a\<^bold>\, \<^bold>\e ?a\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>P\<^bold>R\<^bold>J g \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e ?b\<^bold>\, \<^bold>\f\<^bold>\, \<^bold>\d ?a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\e ?a\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>P\<^bold>R\<^bold>J g, \<^bold>P\<^bold>R\<^bold>J f, \<^bold>\e ?a\<^bold>\\<^bold>]" have "?LHS = \?LHSt\" using f g fg hseq\<^sub>B B.\_def B.\'.map_ide_simp B.VVV.ide_char B.VVV.arr_char B.VV.ide_char B.VV.arr_char P_def by auto also have "... = \?RHSt\" using hseq\<^sub>B by (intro EV.eval_eqI, auto) also have "... = ?RHS" using f g fg hseq\<^sub>B B.\_def B.\'.map_ide_simp B.VVV.ide_char B.VVV.arr_char B.VV.ide_char B.VV.arr_char P_def by auto finally show ?thesis by blast qed thus ?thesis using hseq\<^sub>B by auto qed finally show ?thesis by simp qed also have "... = \\<^sub>1 (g \ f) \ (PoE.cmp (g, f) \ e (src f))" proof - have "\\<^sub>1 (g \ f) \ (PoE.cmp (g, f) \ e (src f)) = \\<^sub>1 (g \ f) \ (P (E (g \ f)) \ P (\\<^sub>E (g, f)) \ \\<^sub>P (E g, E f) \ e (src f))" using f g fg PoE.cmp_def VV.arr_char VV.dom_char by simp also have "... = \\<^sub>1 (g \ f) \ (P (E (g \ f)) \ P (g \\<^sub>B f) \ \\<^sub>P (E g, E f) \ e (src f))" using f g fg emb.cmp_def VV.arr_char by simp also have "... = \\<^sub>1 (g \ f) \ (P (E (g \ f)) \ \\<^sub>P (E g, E f) \ e (src f))" using f g fg hseq\<^sub>B comp_cod_arr emb.map_def hseq_char' prj.cmp_simps'(1,5) by auto also have "... = \\<^sub>1 (g \ f) \ (P (g \\<^sub>B f) \ \\<^sub>P (g, f) \ e (src f))" using hseq\<^sub>B hseq emb.map_def hcomp_char hseqI' by auto also have "... = \\<^sub>1 (g \ f) \ (\\<^sub>P (g, f) \ (P g \ P f) \ e (src f))" using hseq\<^sub>B B.VV.arr_char B.VV.dom_char B.VV.cod_char \\<^sub>P.naturality [of "(g, f)"] P.FF_def by auto also have "... = \\<^sub>1 (g \ f) \ (\\<^sub>P (g, f) \ e (src f))" proof - have "\\<^sub>P (g, f) \ (P g \ P f) = \\<^sub>P (g, f)" using f g fg comp_arr_dom [of "\\<^sub>P (g, f)" "P g \ P f"] VV.arr_char by (metis (no_types, lifting) prj.cmp_simps'(1) \\<^sub>P_simps(4) hseq\<^sub>B) thus ?thesis by simp qed also have "... = \\<^sub>1 (g \ f) \\<^sub>B (\\<^sub>P (g, f) \ e (src f))" proof - (* TODO: For some reason, this requires an epic struggle. *) have "seq (\\<^sub>1 (g \ f)) (\\<^sub>P (g, f) \ e (src f))" proof - have "cod (\\<^sub>P (g, f)) = P (g \ f)" using f g fg \\<^sub>P_simps(5) [of g f] VV.arr_char by (metis (no_types, lifting) hcomp_eqI hseq hseq\<^sub>B hseqI') moreover have "hseq (\\<^sub>P (g, f)) (e (src f))" using f g fg hseq\<^sub>B 1 B.VV.arr_char by auto ultimately show ?thesis using f g fg hseq\<^sub>B 1 B.VV.arr_char P\<^sub>0_props prj.cmp_simps emb.map_def \\<^sub>P_simps(5) [of g f] hcomp_eqI B.VV.dom_char B.VV.cod_char P.FF_def by auto qed thus ?thesis using comp_eqI by simp qed also have "... = \\<^sub>1 (g \\<^sub>B f) \\<^sub>B (\\<^sub>P (g, f) \\<^sub>B e (src f))" using f g fg hseq\<^sub>B 1 B.VV.arr_char hcomp_eqI by simp also have "... = ((e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f]) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, ?a] \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a)) \\<^sub>B (\\<^sub>P (g, f) \\<^sub>B e ?a)" unfolding unit\<^sub>1_def using hseq\<^sub>B 1 comp_char by simp also have "... = ((e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f]) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, ?a] \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a)) \\<^sub>B (((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a] \\<^sub>B (B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f)) \\<^sub>B e ?a)" proof - have "B.VV.ide (g, f)" using f g fg ide_char src_def trg_def B.VV.ide_char by auto hence "\\<^sub>P (g, f) = CMP g f" unfolding \\<^sub>P_def using f g fg ide_char CMP.map_simp_ide by simp also have "... = (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a] \\<^sub>B (B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f)" using hseq\<^sub>B CMP_def by auto finally have "\\<^sub>P (g, f) = (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a] \\<^sub>B (B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f)" by blast thus ?thesis by simp qed also have "... = ((e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f]) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, src f] \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a)) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" using f g fg hseq\<^sub>B src_def trg_def B.whisker_left B.comp_assoc by simp also have "... = (e ?c \\<^sub>B \\<^sub>B[g \\<^sub>B f]) \\<^sub>B \\<^sub>B[e ?c, g \\<^sub>B f, src f] \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B f) \\<^sub>B B.inv (\ ?a)) \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g \\<^sub>B f, d ?a, e ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g \\<^sub>B f, d ?a] \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" proof - have "B.arr ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f))" using hseq\<^sub>B ide_char P.preserves_ide P_def by auto hence "(e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a = ((e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B e ?a) \\<^sub>B ((e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B e ?a) \\<^sub>B (\\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B e ?a) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B e ?a)" using hseq\<^sub>B B.whisker_right by fastforce thus ?thesis using B.comp_assoc by simp qed finally show ?thesis using f ide_char src_def by simp qed finally show "(e (trg g) \ I\<^sub>S.cmp (g, f)) \ \[e (trg g), I\<^sub>S.map g, I\<^sub>S.map f] \ (\\<^sub>1 g \ I\<^sub>S.map f) \ inv \[PoE.map g, e (src g), I\<^sub>S.map f] \ (PoE.map g \ \\<^sub>1 f) \ \[PoE.map g, PoE.map f, e (src f)] = \\<^sub>1 (g \ f) \ (PoE.cmp (g, f) \ e (src f))" by blast qed qed abbreviation (input) counit\<^sub>0 where "counit\<^sub>0 \ d" definition counit\<^sub>1 where "counit\<^sub>1 f = \\<^sub>B[d (trg\<^sub>B f), e (trg\<^sub>B f), f \\<^sub>B d (src\<^sub>B f)] \\<^sub>B (\ (trg\<^sub>B f) \\<^sub>B f \\<^sub>B d (src\<^sub>B f)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d (src\<^sub>B f)]" + abbreviation (input) \\<^sub>0 + where "\\<^sub>0 \ counit\<^sub>0" + abbreviation (input) \\<^sub>1 where "\\<^sub>1 \ counit\<^sub>1" lemma counit\<^sub>1_in_hom [intro]: assumes "B.ide f" shows "\\\<^sub>1 f : f \\<^sub>B d (src\<^sub>B f) \\<^sub>B d (trg\<^sub>B f) \\<^sub>B e (trg\<^sub>B f) \\<^sub>B f \\<^sub>B d (src\<^sub>B f)\" using assms B.iso_is_arr by (unfold counit\<^sub>1_def, intro B.comp_in_homI' B.seqI B.hseqI' B.hcomp_in_vhom) auto lemma counit\<^sub>1_simps [simp]: assumes "B.ide f" shows "B.arr (\\<^sub>1 f)" and "src\<^sub>B (\\<^sub>1 f) = P\<^sub>0 (src\<^sub>B f)" and "trg\<^sub>B (\\<^sub>1 f) = trg\<^sub>B f" and "B.dom (\\<^sub>1 f) = f \\<^sub>B d (src\<^sub>B f)" and "B.cod (\\<^sub>1 f) = d (trg\<^sub>B f) \\<^sub>B e (trg\<^sub>B f) \\<^sub>B f \\<^sub>B d (src\<^sub>B f)" and "B.iso (\\<^sub>1 f)" using assms counit\<^sub>1_in_hom apply auto using B.vconn_implies_hpar(1) apply fastforce using B.vconn_implies_hpar(2) apply fastforce unfolding counit\<^sub>1_def apply (intro B.isos_compose) by auto lemma technical: assumes "B.ide f" and "B.ide g" and "src\<^sub>B g = trg\<^sub>B f" shows "(\\<^sub>1 g \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d (src\<^sub>B g), P f] \\<^sub>B (g \\<^sub>B \\<^sub>1 f) = (\\<^sub>B[d (trg\<^sub>B g), e (trg\<^sub>B g), g \\<^sub>B d (src\<^sub>B g)] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d (trg\<^sub>B g) \\<^sub>B e (trg\<^sub>B g), g, d (src\<^sub>B g)] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(d (trg\<^sub>B g) \\<^sub>B e (trg\<^sub>B g)) \\<^sub>B g, d (src\<^sub>B g), P f] \\<^sub>B (((d (trg\<^sub>B g) \\<^sub>B e (trg\<^sub>B g)) \\<^sub>B g) \\<^sub>B d (src\<^sub>B g) \\<^sub>B P f) \\<^sub>B (((d (trg\<^sub>B g) \\<^sub>B e (trg\<^sub>B g)) \\<^sub>B g) \\<^sub>B \\<^sub>B[d (src\<^sub>B g), e (src\<^sub>B g), f \\<^sub>B d (src\<^sub>B f)]) \\<^sub>B ((\ (trg\<^sub>B g) \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B (\ (src\<^sub>B g) \\<^sub>B f \\<^sub>B d (src\<^sub>B f))) \\<^sub>B (g \\<^sub>B \\<^sub>B[src\<^sub>B g, f, d (src\<^sub>B f)] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d (src\<^sub>B f)))" proof - let ?a = "src\<^sub>B f" let ?b = "src\<^sub>B g" let ?c = "trg\<^sub>B g" have "(\\<^sub>1 g \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \\<^sub>B (g \\<^sub>B \\<^sub>1 f) = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B (\ ?c \\<^sub>B g \\<^sub>B d ?b) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \\<^sub>B (g \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a] \\<^sub>B (\ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a])" using assms counit\<^sub>1_def by simp also have "... = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B ((\ ?c \\<^sub>B g \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \\<^sub>B (g \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (g \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a])" proof - have "\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B (\ ?c \\<^sub>B g \\<^sub>B d ?b) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b] \\<^sub>B P f = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B ((\ ?c \\<^sub>B g \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g \\<^sub>B d ?b] \\<^sub>B P f)" using assms B.iso_is_arr B.whisker_right P_def by auto moreover have "g \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a] \\<^sub>B (\ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a] = (g \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (g \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>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 "... = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B ((\ ?c \\<^sub>B g \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B (\\<^sub>B[?c, g, d ?b] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \\<^sub>B (g \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (g \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a])" using assms B.lunit_hcomp [of g "d ?b"] B.comp_assoc by simp also have "... = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (((\ ?c \\<^sub>B g \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B (\\<^sub>B[?c, g, d ?b] \\<^sub>B P f)) \\<^sub>B ((\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \\<^sub>B (g \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (g \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a])" using assms B.whisker_right P_def B.comp_assoc by simp also have "... = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B (((\ ?c \\<^sub>B g) \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B (((\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f]) \\<^sub>B (g \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (g \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a])" proof - have "((\ ?c \\<^sub>B g \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B (\\<^sub>B[?c, g, d ?b] \\<^sub>B P f) = (\ ?c \\<^sub>B g \\<^sub>B d ?b) \\<^sub>B \\<^sub>B[?c, g, d ?b] \\<^sub>B P f" using assms B.whisker_right P_def B.iso_is_arr by simp also have "... = \\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B d ?b) \\<^sub>B P f" using assms B.assoc_naturality [of "\ ?c" g "d ?b"] B.iso_is_arr by simp also have "... = (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B (((\ ?c \\<^sub>B g) \\<^sub>B d ?b) \\<^sub>B P f)" using assms B.whisker_right P_def B.iso_is_arr by simp finally have "((\ ?c \\<^sub>B g \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B (\\<^sub>B[?c, g, d ?b] \\<^sub>B P f) = (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B (((\ ?c \\<^sub>B g) \\<^sub>B d ?b) \\<^sub>B P f)" by blast thus ?thesis using B.comp_assoc by simp qed also have "... = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B ((((\ ?c \\<^sub>B g) \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?c \\<^sub>B g, d ?b, P f]) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (g \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (g \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a])" proof - have "((\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] = \\<^sub>B\<^sup>-\<^sup>1[?c \\<^sub>B g, d ?b, P f] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B d ?b \\<^sub>B P f)" using assms B.assoc'_naturality [of "\\<^sub>B\<^sup>-\<^sup>1[g]" "d ?b" "P f"] by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(d ?c \\<^sub>B e ?c) \\<^sub>B g, d ?b, P f] \\<^sub>B (((\ ?c \\<^sub>B g) \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B d ?b \\<^sub>B P f)) \\<^sub>B (g \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (g \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a])" proof - have "(((\ ?c \\<^sub>B g) \\<^sub>B d ?b) \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?c \\<^sub>B g, d ?b, P f] = \\<^sub>B\<^sup>-\<^sup>1[(d ?c \\<^sub>B e ?c) \\<^sub>B g, d ?b, P f] \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B d ?b \\<^sub>B P f)" using assms B.assoc'_naturality [of "\ ?c \\<^sub>B g" "d ?b" "P f"] B.iso_is_arr by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(d ?c \\<^sub>B e ?c) \\<^sub>B g, d ?b, P f] \\<^sub>B (((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (g \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a])) \\<^sub>B (g \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a])" using assms B.whisker_right B.iso_is_arr P_def B.comp_assoc by simp also have "... = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(d ?c \\<^sub>B e ?c) \\<^sub>B g, d ?b, P f] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (g \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a])" proof - have "((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (g \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) = (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a])" using assms B.comp_arr_dom B.comp_cod_arr B.iso_is_arr P_def B.interchange [of "(\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g]" g] B.interchange [of "(d ?c \\<^sub>B e ?c) \\<^sub>B g" "(\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g]"] by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(d ?c \\<^sub>B e ?c) \\<^sub>B g, d ?b, P f] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a]))" proof - have "((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (g \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) = (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a)" using assms B.comp_arr_dom B.comp_cod_arr B.iso_is_arr P_def B.interchange [of "(\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g]" g] B.interchange [of "(d ?c \\<^sub>B e ?c) \\<^sub>B g" "(\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g]"] by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(d ?c \\<^sub>B e ?c) \\<^sub>B g, d ?b, P f] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B (\ ?b \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a))" proof - have "((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a]) = (\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B (\ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a]" using assms B.comp_arr_dom B.iso_is_arr P_def B.interchange [of "(\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g]" g "\ ?b \\<^sub>B f \\<^sub>B d ?a" "\\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a]"] by simp also have "... = (\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B (\ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)" using assms B.lunit_hcomp by simp also have "... = ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B (\ ?b \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a))" using assms B.comp_arr_dom B.iso_is_arr P_def B.interchange [of "(\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g]" g] by simp finally have "((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[f \\<^sub>B d ?a]) = ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B (\ ?b \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a))" by blast thus ?thesis by simp qed finally show ?thesis by blast qed sublocale counit: pseudonatural_equivalence V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \E \ P\ EoP.cmp I\<^sub>C.map I\<^sub>C.cmp counit\<^sub>0 counit\<^sub>1 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 V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \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 : src\<^sub>B ((E \ P) a) \\<^sub>B src\<^sub>B (I\<^sub>C.map a)\" unfolding emb.map_def by fastforce show "\f. B.ide f \ B.iso (\\<^sub>1 f)" by simp show "\f. B.ide f \ \\\<^sub>1 f : I\<^sub>C.map f \\<^sub>B d (src\<^sub>B f) \\<^sub>B d (trg\<^sub>B f) \\<^sub>B (E \ P) f\" unfolding counit\<^sub>1_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 \ \ \\<^sub>1 (B.cod \) \\<^sub>B (I\<^sub>C.map \ \\<^sub>B d (src\<^sub>B \)) = (d (trg\<^sub>B \) \\<^sub>B (E \ P) \) \\<^sub>B \\<^sub>1 (B.dom \)" proof - fix \ assume \: "B.arr \" let ?a = "src\<^sub>B \" let ?b = "trg\<^sub>B \" let ?f = "B.dom \" let ?g = "B.cod \" have "\\<^sub>1 ?g \\<^sub>B (I\<^sub>C.map \ \\<^sub>B d ?a) = (\\<^sub>B[d ?b, e ?b, ?g \\<^sub>B d ?a] \\<^sub>B (\ ?b \\<^sub>B ?g \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g \\<^sub>B d ?a]) \\<^sub>B (\ \\<^sub>B d ?a)" using \ counit\<^sub>1_def P_def emb.map_def arr_char P\<^sub>0_props(1) by simp also have "... = (d ?b \\<^sub>B e ?b \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B \\<^sub>B[d ?b, e ?b, ?f \\<^sub>B d ?a] \\<^sub>B (\ ?b \\<^sub>B ?f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?f \\<^sub>B d ?a]" proof - have "(\\<^sub>B[d ?b, e ?b, ?g \\<^sub>B d ?a] \\<^sub>B (\ ?b \\<^sub>B ?g \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g \\<^sub>B d ?a]) \\<^sub>B (\ \\<^sub>B d ?a) = \\<^sub>B[d ?b, e ?b, ?g \\<^sub>B d ?a] \\<^sub>B (\ ?b \\<^sub>B ?g \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?g \\<^sub>B d ?a] \\<^sub>B (\ \\<^sub>B d ?a)" using B.comp_assoc by simp also have "... = \\<^sub>B[d ?b, e ?b, ?g \\<^sub>B d ?a] \\<^sub>B ((\ ?b \\<^sub>B ?g \\<^sub>B d ?a) \\<^sub>B (?b \\<^sub>B \ \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?f \\<^sub>B d ?a]" using \ B.lunit'_naturality [of "\ \\<^sub>B d ?a"] B.comp_assoc by simp also have "... = (\\<^sub>B[d ?b, e ?b, ?g \\<^sub>B d ?a] \\<^sub>B ((d ?b \\<^sub>B e ?b) \\<^sub>B \ \\<^sub>B d ?a)) \\<^sub>B (\ ?b \\<^sub>B ?f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?f \\<^sub>B d ?a]" proof - have "(\ ?b \\<^sub>B ?g \\<^sub>B d ?a) \\<^sub>B (?b \\<^sub>B \ \\<^sub>B d ?a) = \ ?b \\<^sub>B ?b \\<^sub>B ?g \\<^sub>B \ \\<^sub>B d ?a \\<^sub>B d ?a" using \ B.iso_is_arr B.interchange [of "\ ?b" ?b "(?g \\<^sub>B d ?a)" "\ \\<^sub>B d ?a"] B.interchange [of "?g" \ "d ?a" "d ?a"] by simp also have "... = (d ?b \\<^sub>B e ?b) \\<^sub>B \ ?b \\<^sub>B \ \\<^sub>B ?f \\<^sub>B d ?a \\<^sub>B d ?a" using \ B.comp_arr_dom B.comp_cod_arr apply simp by (metis B.arrI equivalence_data_in_hom\<^sub>B(7) equivalence_data_simps\<^sub>B(17-18) B.obj_trg) also have "... = ((d ?b \\<^sub>B e ?b) \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B (\ ?b \\<^sub>B ?f \\<^sub>B d ?a)" using \ B.iso_is_arr B.interchange [of "d ?b \\<^sub>B e ?b" "\ ?b" "\ \\<^sub>B d ?a" "?f \\<^sub>B d ?a"] B.interchange [of \ "?f" "d ?a" "d ?a"] by simp finally have "(\ ?b \\<^sub>B ?g \\<^sub>B d ?a) \\<^sub>B (?b \\<^sub>B \ \\<^sub>B d ?a) = ((d ?b \\<^sub>B e ?b) \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B (\ ?b \\<^sub>B ?f \\<^sub>B d ?a)" by blast thus ?thesis using B.comp_assoc by simp qed also have "... = (d ?b \\<^sub>B e ?b \\<^sub>B \ \\<^sub>B d ?a) \\<^sub>B \\<^sub>B[d ?b, e ?b, ?f \\<^sub>B d ?a] \\<^sub>B (\ ?b \\<^sub>B ?f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[?f \\<^sub>B d ?a]" using \ B.assoc_naturality [of "d ?b" "e ?b" "\ \\<^sub>B d ?a"] B.comp_assoc by simp finally show ?thesis by simp qed also have "... = (d ?b \\<^sub>B (E \ P) \) \\<^sub>B \\<^sub>1 ?f" using \ counit\<^sub>1_def P_def emb.map_def arr_char P\<^sub>0_props(1) P_simps\<^sub>B(1) by (simp add: P\<^sub>0_props(6)) finally show "\\<^sub>1 (?g) \\<^sub>B (I\<^sub>C.map \ \\<^sub>B d ?a) = (d ?b \\<^sub>B (E \ P) \) \\<^sub>B \\<^sub>1 ?f" by blast qed show "\f g. \B.ide f; B.ide g; src\<^sub>B g = trg\<^sub>B f\ \ (d (trg\<^sub>B g) \\<^sub>B EoP.cmp (g, f)) \\<^sub>B \\<^sub>B[d (trg\<^sub>B g), (E \ P) g, (E \ P) f] \\<^sub>B (\\<^sub>1 g \\<^sub>B (E \ P) f) \\<^sub>B B.inv \\<^sub>B[I\<^sub>C.map g, d (src\<^sub>B g), (E \ P) f] \\<^sub>B (I\<^sub>C.map g \\<^sub>B \\<^sub>1 f) \\<^sub>B \\<^sub>B[I\<^sub>C.map g, I\<^sub>C.map f, d (src\<^sub>B f)] = \\<^sub>1 (g \\<^sub>B f) \\<^sub>B (I\<^sub>C.cmp (g, f) \\<^sub>B d (src\<^sub>B f))" proof - fix f g assume f: "B.ide f" and g: "B.ide g" and fg: "src\<^sub>B g = trg\<^sub>B f" let ?a = "src\<^sub>B f" let ?b = "src\<^sub>B g" let ?c = "trg\<^sub>B g" have "(d ?c \\<^sub>B EoP.cmp (g, f)) \\<^sub>B \\<^sub>B[d ?c, (E \ P) g, (E \ P) f] \\<^sub>B (\\<^sub>1 g \\<^sub>B (E \ P) f) \\<^sub>B B.inv \\<^sub>B[I\<^sub>C.map g, d ?b, (E \ P) f] \\<^sub>B (I\<^sub>C.map g \\<^sub>B \\<^sub>1 f) \\<^sub>B \\<^sub>B[I\<^sub>C.map g, I\<^sub>C.map f, d ?a] = (d ?c \\<^sub>B P (g \\<^sub>B f) \\<^sub>B CMP.map (g, f) \\<^sub>B (P g \\<^sub>B P f)) \\<^sub>B \\<^sub>B[d ?c, P g, P f] \\<^sub>B (\\<^sub>1 g \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \\<^sub>B (g \\<^sub>B \\<^sub>1 f) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "src\<^sub>B (e ?c \\<^sub>B g \\<^sub>B d ?b) = trg\<^sub>B (e ?b \\<^sub>B f \\<^sub>B d ?a)" using f g fg src_def trg_def arr_char P_simps\<^sub>B(1) by (simp add: P\<^sub>0_props(1,6) arr_char) moreover have "B.inv \\<^sub>B[g, d ?b, P f] = \\<^sub>B\<^sup>-\<^sup>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 \\<^sub>P_def emb.cmp_def EoP.cmp_def B.VV.arr_char B.VV.dom_char VV.arr_char by simp qed also have "... = (d ?c \\<^sub>B CMP.map (g, f)) \\<^sub>B \\<^sub>B[d ?c, P g, P f] \\<^sub>B ((\\<^sub>1 g \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, d ?b, P f] \\<^sub>B (g \\<^sub>B \\<^sub>1 f)) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "\CMP.map (g, f) : P g \ P f \ P (g \\<^sub>B f)\" using f g fg VV.arr_char VV.cod_char P.FF_def \\<^sub>P_def \\<^sub>P_in_hom(2) by auto hence "\CMP.map (g, f) : P g \\<^sub>B P f \\<^sub>B P (g \\<^sub>B f)\" using in_hom_char hcomp_char by auto hence "P (g \\<^sub>B f) \\<^sub>B CMP.map (g, f) \\<^sub>B (P g \\<^sub>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 \\<^sub>B CMP.map (g, f)) \\<^sub>B \\<^sub>B[d ?c, P g, P f] \\<^sub>B (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(d ?c \\<^sub>B e ?c) \\<^sub>B g, d ?b, P f] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" using f g fg technical B.comp_assoc by simp also have "... = (d ?c \\<^sub>B (e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B P f] \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, P f] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f)) \\<^sub>B \\<^sub>B[d ?c, P g, P f] \\<^sub>B (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(d ?c \\<^sub>B e ?c) \\<^sub>B g, d ?b, P f] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" using f g fg CMP.map_simp_ide CMP_def B.VV.ide_char B.VV.arr_char B.whisker_left B.whisker_left B.comp_assoc by simp also have "... = (d ?c \\<^sub>B e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B (d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B ((d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (d ?c \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B P f]) \\<^sub>B (d ?c \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, P f]) \\<^sub>B (d ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B[d ?c, P g, P f] \\<^sub>B (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(d ?c \\<^sub>B e ?c) \\<^sub>B g, d ?b, P f] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a])) \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>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 \\<^sub>B e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B ((d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B[d ?c, e ?c, g \\<^sub>B (d ?b \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[d ?c \\<^sub>B e ?c, g, (d ?b \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "(d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (d ?c \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B P f]) \\<^sub>B (d ?c \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, P f]) \\<^sub>B (d ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B[d ?c, P g, P f] \\<^sub>B (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(d ?c \\<^sub>B e ?c) \\<^sub>B g, d ?b, P f] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) = \\<^sub>B[d ?c, e ?c, g \\<^sub>B (d ?b \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B \\<^sub>B[d ?c \\<^sub>B e ?c, g, (d ?b \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B d ?a]" proof - have "(d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d ?b, e ?b, f \\<^sub>B d ?a]) \\<^sub>B (d ?c \\<^sub>B \\<^sub>B[e ?c, g, d ?b \\<^sub>B P f]) \\<^sub>B (d ?c \\<^sub>B \\<^sub>B[e ?c \\<^sub>B g, d ?b, P f]) \\<^sub>B (d ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B[d ?c, P g, P f] \\<^sub>B (\\<^sub>B[d ?c, e ?c, g \\<^sub>B d ?b] \\<^sub>B P f) \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, d ?b] \\<^sub>B P f) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(d ?c \\<^sub>B e ?c) \\<^sub>B g, d ?b, P f] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B d ?b \\<^sub>B P f) \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B \\<^sub>B[d ?b, e ?b, f \\<^sub>B d ?a]) = \(\<^bold>\d ?c\<^bold>\ \<^bold>\ \<^bold>\e ?c\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\d ?b\<^bold>\, \<^bold>\e ?b\<^bold>\, \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\d ?c\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?c\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>P\<^bold>R\<^bold>J f\<^bold>]) \<^bold>\ (\<^bold>\d ?c\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\e ?c\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\, \<^bold>P\<^bold>R\<^bold>J f\<^bold>]) \<^bold>\ (\<^bold>\d ?c\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\e ?c\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>P\<^bold>R\<^bold>J f) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\d ?c\<^bold>\, \<^bold>P\<^bold>R\<^bold>J g, \<^bold>P\<^bold>R\<^bold>J f\<^bold>] \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\d ?c\<^bold>\, \<^bold>\e ?c\<^bold>\, \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\d ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>P\<^bold>R\<^bold>J f) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\d ?c\<^bold>\ \<^bold>\ \<^bold>\e ?c\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>P\<^bold>R\<^bold>J f) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[(\<^bold>\d ?c\<^bold>\ \<^bold>\ \<^bold>\e ?c\<^bold>\) \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\d ?b\<^bold>\, \<^bold>P\<^bold>R\<^bold>J f\<^bold>] \<^bold>\ (((\<^bold>\d ?c\<^bold>\ \<^bold>\ \<^bold>\e ?c\<^bold>\) \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\ \<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>P\<^bold>R\<^bold>J f) \<^bold>\ (((\<^bold>\d ?c\<^bold>\ \<^bold>\ \<^bold>\e ?c\<^bold>\) \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\d ?b\<^bold>\, \<^bold>\e ?b\<^bold>\, \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\\<^bold>])\" using f g fg B.\_def B.\'.map_ide_simp B.VVV.ide_char B.VVV.arr_char B.VV.ide_char B.VV.arr_char P_def B.\_ide_simp by auto also have "... = \\<^bold>\\<^bold>[\<^bold>\d ?c\<^bold>\, \<^bold>\e ?c\<^bold>\, \<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>\e ?b\<^bold>\) \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\d ?c\<^bold>\ \<^bold>\ \<^bold>\e ?c\<^bold>\, \<^bold>\g\<^bold>\, (\<^bold>\d ?b\<^bold>\ \<^bold>\ \<^bold>\e ?b\<^bold>\) \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\\<^bold>]\" using f g fg by (intro EV.eval_eqI, auto) also have "... = \\<^sub>B[d ?c, e ?c, g \\<^sub>B (d ?b \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B \\<^sub>B[d ?c \\<^sub>B e ?c, g, (d ?b \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B d ?a]" using f g fg B.\_def B.\'.map_ide_simp B.VVV.ide_char B.VVV.arr_char B.VV.ide_char B.VV.arr_char 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 \\<^sub>B e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[d ?c, e ?c, g \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B[d ?c \\<^sub>B e ?c, g, (d ?b \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B d ?a]) \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "(d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B[d ?c, e ?c, g \\<^sub>B (d ?b \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B d ?a] = \\<^sub>B[d ?c, e ?c, g \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B ((d ?c \\<^sub>B e ?c) \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a)" using f g fg B.assoc_naturality [of "d ?c" "e ?c" "g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a"] by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (d ?c \\<^sub>B e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[d ?c, e ?c, g \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B \\<^sub>B[d ?c \\<^sub>B e ?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B ((((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "((d ?c \\<^sub>B e ?c) \\<^sub>B g \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B[d ?c \\<^sub>B e ?c, g, (d ?b \\<^sub>B e ?b) \\<^sub>B f \\<^sub>B d ?a] = \\<^sub>B[d ?c \\<^sub>B e ?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a)" using f g fg B.assoc_naturality [of "d ?c \\<^sub>B e ?c" g "B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a"] by simp thus ?thesis using B.comp_assoc by simp qed also have "... = (d ?c \\<^sub>B e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[d ?c, e ?c, g \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B \\<^sub>B[d ?c \\<^sub>B e ?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "(((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) = (\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a" proof - have "(((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B \ ?b \\<^sub>B f \\<^sub>B d ?a) = ((d ?c \\<^sub>B e ?c) \\<^sub>B g) \\<^sub>B (\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B (B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (\ ?b \\<^sub>B f \\<^sub>B d ?a)" using f g fg B.interchange [of "(d ?c \\<^sub>B e ?c) \\<^sub>B g" "(\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g]" "B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a" "\ ?b \\<^sub>B f \\<^sub>B d ?a"] by fastforce also have "... = (\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B (B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (\ ?b \\<^sub>B f \\<^sub>B d ?a)" using f g fg B.comp_cod_arr [of "(\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g]" "(d ?c \\<^sub>B e ?c) \\<^sub>B g"] by (auto simp add: B.iso_is_arr) also have "... = (\ ?c \\<^sub>B g) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a" proof - have "(B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (\ ?b \\<^sub>B f \\<^sub>B d ?a) = B.inv (\ ?b) \\<^sub>B \ ?b \\<^sub>B f \\<^sub>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 \\<^sub>B f \\<^sub>B d ?a" using f g fg chosen_adjoint_equivalence(5) B.comp_inv_arr' by simp finally have "(B.inv (\ ?b) \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (\ ?b \\<^sub>B f \\<^sub>B d ?a) = ?b \\<^sub>B f \\<^sub>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 \\<^sub>B e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[d ?c, e ?c, g \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (\\<^sub>B[d ?c \\<^sub>B e ?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B ((\ ?c \\<^sub>B g) \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" using f g fg B.comp_assoc B.whisker_right [of "?b \\<^sub>B f \\<^sub>B d ?a" "\ ?c \\<^sub>B g" "\\<^sub>B\<^sup>-\<^sup>1[g]"] by fastforce also have "... = (d ?c \\<^sub>B e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B ((d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[d ?c, e ?c, g \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a]) \\<^sub>B (\ ?c \\<^sub>B g \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B \\<^sub>B[?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" using f g fg B.iso_is_arr B.comp_assoc B.assoc_naturality [of "\ ?c" g "?b \\<^sub>B f \\<^sub>B d ?a"] by simp also have "... = (d ?c \\<^sub>B e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B \\<^sub>B[d ?c, e ?c, g \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B (\ ?c \\<^sub>B g \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "(d ?c \\<^sub>B e ?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[d ?c, e ?c, g \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a] = \\<^sub>B[d ?c, e ?c, g \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B ((d ?c \\<^sub>B e ?c) \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a])" using f g fg B.iso_is_arr B.assoc_naturality [of "d ?c" "e ?c" "g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]"] by simp thus ?thesis using B.comp_assoc by simp qed also have "... = ((d ?c \\<^sub>B e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B \\<^sub>B[d ?c, e ?c, g \\<^sub>B f \\<^sub>B d ?a]) \\<^sub>B (\ ?c \\<^sub>B g \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "((d ?c \\<^sub>B e ?c) \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B (\ ?c \\<^sub>B g \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) = (\ ?c \\<^sub>B g \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a])" using f g fg B.comp_arr_dom B.comp_cod_arr B.iso_is_arr B.interchange [of "d ?c \\<^sub>B e ?c" "\ ?c" "g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]" "g \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a"] B.interchange [of "\ ?c" ?c "g \\<^sub>B f \\<^sub>B d ?a" "g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]"] by auto thus ?thesis using B.comp_assoc by simp qed also have "... = \\<^sub>B[d ?c, e ?c, (g \\<^sub>B f) \\<^sub>B d ?a] \\<^sub>B (((d ?c \\<^sub>B e ?c) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (\ ?c \\<^sub>B g \\<^sub>B f \\<^sub>B d ?a)) \\<^sub>B (?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "(d ?c \\<^sub>B e ?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B \\<^sub>B[d ?c, e ?c, g \\<^sub>B f \\<^sub>B d ?a] = \\<^sub>B[d ?c, e ?c, (g \\<^sub>B f) \\<^sub>B d ?a] \\<^sub>B ((d ?c \\<^sub>B e ?c) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a])" using f g fg B.assoc_naturality [of "d ?c" "e ?c" "\\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]"] by simp thus ?thesis using B.comp_assoc by simp qed also have "... = \\<^sub>B[d ?c, e ?c, (g \\<^sub>B f) \\<^sub>B d ?a] \\<^sub>B (\ ?c \\<^sub>B (g \\<^sub>B f) \\<^sub>B d ?a) \\<^sub>B (?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "((d ?c \\<^sub>B e ?c) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (\ ?c \\<^sub>B g \\<^sub>B f \\<^sub>B d ?a) = (\ ?c \\<^sub>B (g \\<^sub>B f) \\<^sub>B d ?a) \\<^sub>B (?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>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 \\<^sub>B e ?c" "\ ?c" "\\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]" "g \\<^sub>B f \\<^sub>B d ?a"] B.interchange [of "\ ?c" ?c "(g \\<^sub>B f) \\<^sub>B d ?a" "\\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]"] by auto thus ?thesis using B.comp_assoc by simp qed also have "... = \\<^sub>1 (g \\<^sub>B f) \\<^sub>B (I\<^sub>C.cmp (g, f) \\<^sub>B d ?a)" proof - have "\\<^sub>1 (g \\<^sub>B f) \\<^sub>B (I\<^sub>C.cmp (g, f) \\<^sub>B d ?a) = \\<^sub>B[d ?c, e ?c, (g \\<^sub>B f) \\<^sub>B d ?a] \\<^sub>B (\ ?c \\<^sub>B (g \\<^sub>B f) \\<^sub>B d ?a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[(g \\<^sub>B f) \\<^sub>B d ?a] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B d ?a)" using f g fg counit\<^sub>1_def B.comp_assoc by simp also have "... = \\<^sub>B[d ?c, e ?c, (g \\<^sub>B f) \\<^sub>B d ?a] \\<^sub>B (\ ?c \\<^sub>B (g \\<^sub>B f) \\<^sub>B d ?a) \\<^sub>B (?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "\\<^sub>B\<^sup>-\<^sup>1[(g \\<^sub>B f) \\<^sub>B d ?a] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B d ?a) = (?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" proof - have "\\<^sub>B\<^sup>-\<^sup>1[(g \\<^sub>B f) \\<^sub>B d ?a] \\<^sub>B ((g \\<^sub>B f) \\<^sub>B d ?a) = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[(\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\f\<^bold>\) \<^bold>\ \<^bold>\d ?a\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\f\<^bold>\) \<^bold>\ \<^bold>\d ?a\<^bold>\)\" using f g fg B.\_ide_simp by auto also have "... = \(\<^bold>\?c\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\f\<^bold>\, \<^bold>\d ?a\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\?c\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\?c\<^bold>\\<^sub>0, \<^bold>\g\<^bold>\, \<^bold>\?b\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\\<^bold>] \<^bold>\ \<^bold>\?b\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\d ?a\<^bold>\) \<^bold>\ (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\?b\<^bold>\\<^sub>0, \<^bold>\f\<^bold>\, \<^bold>\d ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\\<^bold>] \<^bold>\ \<^bold>\d ?a\<^bold>\)) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\f\<^bold>\, \<^bold>\d ?a\<^bold>\\<^bold>]\" using f g fg by (intro EV.eval_eqI, auto) also have "... = (?c \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[g, f, d ?a]) \\<^sub>B (?c \\<^sub>B g \\<^sub>B \\<^sub>B[f \\<^sub>B d ?a]) \\<^sub>B \\<^sub>B[?c, g, ?b \\<^sub>B f \\<^sub>B d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[g] \\<^sub>B ?b \\<^sub>B f \\<^sub>B d ?a) \\<^sub>B (g \\<^sub>B \\<^sub>B[?b, f, d ?a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[f] \\<^sub>B d ?a)) \\<^sub>B \\<^sub>B[g, f, d ?a]" using f g fg B.\_def B.\'.map_ide_simp B.VVV.ide_char B.VVV.arr_char B.VV.ide_char B.VV.arr_char 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 \\<^sub>B EoP.cmp (g, f)) \\<^sub>B \\<^sub>B[d ?c, (E \ P) g, (E \ P) f] \\<^sub>B (\\<^sub>1 g \\<^sub>B (E \ P) f) \\<^sub>B B.inv \\<^sub>B[I\<^sub>C.map g, d ?b, (E \ P) f] \\<^sub>B (I\<^sub>C.map g \\<^sub>B \\<^sub>1 f) \\<^sub>B \\<^sub>B[I\<^sub>C.map g, I\<^sub>C.map f, d ?a] = \\<^sub>1 (g \\<^sub>B f) \\<^sub>B (I\<^sub>C.cmp (g, f) \\<^sub>B d ?a)" by blast qed show "\a. B.obj a \ (d a \\<^sub>B EoP.unit a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a] = \\<^sub>1 a \\<^sub>B (I\<^sub>C.unit a \\<^sub>B d a)" proof - fix a assume a: "B.obj a" interpret adjoint_equivalence_in_bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B \e a\ \d a\ \\ a\ \\ a\ using a chosen_adjoint_equivalence by simp have 0: "src\<^sub>B a = a \ trg\<^sub>B a = a" using a by auto have 1: "obj (P\<^sub>0 a)" using a P\<^sub>0_props(1) by simp have "(d a \\<^sub>B EoP.unit a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a] = (d a \\<^sub>B ((e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) \\<^sub>B src\<^sub>B (a \\<^sub>B d (src\<^sub>B a))) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a]" proof - have "B.hseq (e (trg\<^sub>B a)) (a \\<^sub>B d (src\<^sub>B a))" using a by (elim B.objE, intro B.hseqI') auto moreover have "arr (src\<^sub>B (d (src\<^sub>B a)))" using a arr_char P\<^sub>0_props(1) obj_char B.obj_simps(1) by auto moreover have "arr (src\<^sub>B (a \\<^sub>B d (src\<^sub>B a)))" using a arr_char P\<^sub>0_props(1) obj_char calculation(1) by fastforce moreover have "src\<^sub>B (a \\<^sub>B d (src\<^sub>B a)) \ Obj \ trg\<^sub>B (e (trg\<^sub>B a)) \ Obj" using a B.obj_simps P\<^sub>0_props obj_char arr_char by simp moreover have "P\<^sub>0 a \\<^sub>B P\<^sub>0 a \ Obj" using 1 arr_char P\<^sub>0_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 (P\<^sub>0 (src\<^sub>B a)) = P\<^sub>0 (src\<^sub>B a)" using a 0 1 emb.unit_char' P.map\<^sub>0_def src_def by simp ultimately show ?thesis using a emb.map_def EoP.unit_char' prj_unit_char emb.unit_char' P.map\<^sub>0_def P_def src_def arr_char obj_char by simp qed also have "... = (d a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a) \\<^sub>B P\<^sub>0 a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a]" using a 1 B.comp_assoc B.obj_simps by auto also have "... = (d a \\<^sub>B (e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B B.inv (\ a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a]" using a B.comp_arr_dom by simp also have "... = (d a \\<^sub>B e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (d a \\<^sub>B B.inv (\ a)) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a]" using a B.whisker_left B.comp_assoc by simp also have "... = ((d a \\<^sub>B e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B \\<^sub>B[d a, e a, d a]) \\<^sub>B (\ a \\<^sub>B d a)" using a triangle_right B.comp_assoc B.invert_side_of_triangle(1) [of "\\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a]" "d a \\<^sub>B \ a"] by simp also have "... = \\<^sub>B[d a, e a, a \\<^sub>B d a] \\<^sub>B ((d a \\<^sub>B e a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B (\ a \\<^sub>B d a)" proof - have "(d a \\<^sub>B e a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]) \\<^sub>B \\<^sub>B[d a, e a, d a] = \\<^sub>B[d a, e a, a \\<^sub>B d a] \\<^sub>B ((d a \\<^sub>B e a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])" using a B.assoc_naturality [of "d a" "e a" "\\<^sub>B\<^sup>-\<^sup>1[d a]"] by simp thus ?thesis using B.comp_assoc by simp qed also have "... = \\<^sub>B[d a, e a, a \\<^sub>B d a] \\<^sub>B (\ a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])" using a B.interchange [of "d a \\<^sub>B e a" "\ a" "\\<^sub>B\<^sup>-\<^sup>1[d a]" "d a"] B.comp_arr_dom B.comp_cod_arr by simp also have "... = \\<^sub>B[d a, e a, a \\<^sub>B d a] \\<^sub>B (\ a \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B[a, a, d a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[a] \\<^sub>B d a)" proof - have "\\<^sub>B[a, a, d a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[a] \\<^sub>B d a) = a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]" proof - (* TODO: I wanted to prove this directly, but missed some necessary trick. *) have "\\<^sub>B[a, a, d a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[a] \\<^sub>B d a) = \\<^bold>\\<^bold>[\<^bold>\a\<^bold>\\<^sub>0, \<^bold>\a\<^bold>\\<^sub>0, \<^bold>\d a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\a\<^bold>\\<^sub>0\<^bold>] \<^bold>\ \<^bold>\d a\<^bold>\)\" using a ide_char B.\_def B.\'.map_ide_simp B.VVV.ide_char B.VVV.arr_char B.VV.ide_char B.VV.arr_char P_def B.\_ide_simp by auto also have "... = \\<^bold>\a\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\d a\<^bold>\\<^bold>]\" using a ide_char by (intro EV.eval_eqI, auto) also have "... = a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a]" using a ide_char B.\_def B.\'.map_ide_simp B.VVV.ide_char B.VVV.arr_char B.VV.ide_char B.VV.arr_char P_def B.\_ide_simp by auto finally show ?thesis by blast qed hence "\\<^sub>B[d a, e a, a \\<^sub>B d a] \\<^sub>B (\ a \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B[a, a, d a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[a] \\<^sub>B d a) = \\<^sub>B[d a, e a, a \\<^sub>B d a] \\<^sub>B (\ a \\<^sub>B a \\<^sub>B d a) \\<^sub>B (a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])" by simp also have "... = \\<^sub>B[d a, e a, a \\<^sub>B d a] \\<^sub>B (\ a \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a])" using a B.interchange [of "\ a" a "a \\<^sub>B d a" "\\<^sub>B\<^sup>-\<^sup>1[d a]"] B.comp_arr_dom B.comp_cod_arr by simp finally show ?thesis by simp qed also have "... = \\<^sub>1 a \\<^sub>B (I\<^sub>C.unit a \\<^sub>B d a)" proof - have "\\<^sub>1 a \\<^sub>B (I\<^sub>C.unit a \\<^sub>B d a) = \\<^sub>B[d a, e a, a \\<^sub>B d a] \\<^sub>B (\ a \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a \\<^sub>B d a] \\<^sub>B (a \\<^sub>B d a)" using a 0 counit\<^sub>1_def I\<^sub>C.unit_char' B.comp_assoc by simp also have "... = \\<^sub>B[d a, e a, a \\<^sub>B d a] \\<^sub>B (\ a \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[a \\<^sub>B d a]" proof - have "B.arr \\<^sub>B\<^sup>-\<^sup>1[a \\<^sub>B d a]" using a by (metis B.ide_hcomp B.lunit'_simps(1) B.objE equivalence_data_simps\<^sub>B(8) ide_right) moreover have "B.dom \\<^sub>B\<^sup>-\<^sup>1[a \\<^sub>B d a] = a \\<^sub>B d a" using a by (metis B.ide_hcomp B.lunit'_simps(4) B.objE antipar(1) equivalence_data_simps\<^sub>B(5) ide_right) ultimately show ?thesis using a B.comp_arr_dom [of "\\<^sub>B\<^sup>-\<^sup>1[a \\<^sub>B d a]" "a \\<^sub>B d a"] by simp qed also have "... = \\<^sub>B[d a, e a, a \\<^sub>B d a] \\<^sub>B (\ a \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B[a, a, d a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[a] \\<^sub>B d a)" using a B.lunit_hcomp(4) [of a "d a"] by auto finally have "\\<^sub>1 a \\<^sub>B (I\<^sub>C.unit a \\<^sub>B d a) = \\<^sub>B[d a, e a, a \\<^sub>B d a] \\<^sub>B (\ a \\<^sub>B a \\<^sub>B d a) \\<^sub>B \\<^sub>B[a, a, d a] \\<^sub>B (\\<^sub>B\<^sup>-\<^sup>1[a] \\<^sub>B d a)" by simp thus ?thesis by simp qed finally show "(d a \\<^sub>B EoP.unit a) \\<^sub>B \\<^sub>B\<^sup>-\<^sup>1[d a] \\<^sub>B \\<^sub>B[d a] = \\<^sub>1 a \\<^sub>B (I\<^sub>C.unit a \\<^sub>B d a)" by blast qed qed - sublocale equivalence_of_bicategories V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \ \\<^sub>B src trg - E \\<^sub>E P \\<^sub>P unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1 + interpretation equivalence_of_bicategories V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \ \\<^sub>B src trg + E \\<^sub>E P \\<^sub>P unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1 .. lemma induces_equivalence: shows "equivalence_of_bicategories V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B comp hcomp \ \\<^sub>B src trg E \\<^sub>E P \\<^sub>P unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1" .. 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 map\<^sub>0 (Collect C.obj) (Collect D.obj)" begin abbreviation F\<^sub>0 where "F\<^sub>0 \ map\<^sub>0" definition G\<^sub>0 where "G\<^sub>0 = inv_into (Collect C.obj) F\<^sub>0" lemma G\<^sub>0_props: shows "D.obj b \ C.obj (G\<^sub>0 b) \ F\<^sub>0 (G\<^sub>0 b) = b" and "C.obj a \ D.obj (F\<^sub>0 a) \ G\<^sub>0 (F\<^sub>0 a) = a" proof - have surj: "F\<^sub>0 ` (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 (G\<^sub>0 b) \ F\<^sub>0 (G\<^sub>0 b) = b" proof show "C.obj (G\<^sub>0 b)" unfolding G\<^sub>0_def using b by (metis inv_into_into mem_Collect_eq surj) show "F\<^sub>0 (G\<^sub>0 b) = b" unfolding G\<^sub>0_def using b by (simp add: f_inv_into_f surj) qed next have bij: "bij_betw G\<^sub>0 (Collect D.obj) (Collect C.obj)" using bij_on_obj G\<^sub>0_def bij_betw_inv_into by auto have surj: "G\<^sub>0 ` (Collect D.obj) = Collect C.obj" using bij_on_obj G\<^sub>0_def by (metis bij_betw_def bij_betw_inv_into) assume a: "C.obj a" show "D.obj (F\<^sub>0 a) \ G\<^sub>0 (F\<^sub>0 a) = a" using a bij surj G\<^sub>0_def bij_betw_imp_inj_on bij_on_obj by force qed text \ We extend \G\<^sub>0\ to all arrows of \D\ using chosen adjoint equivalences, which extend \F\, between \hom\<^sub>C (a, a')\ and \hom\<^sub>D (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.G\<^sub>1 V\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D src\<^sub>D trg\<^sub>D F (G\<^sub>0 (src\<^sub>D \)) (G\<^sub>0 (trg\<^sub>D \)) \ else C.null)" lemma G_in_hom [intro]: assumes "D.arr \" shows "\G \ : G\<^sub>0 (src\<^sub>D \) \\<^sub>C G\<^sub>0 (trg\<^sub>D \)\" and "\G \ : G (D.dom \) \\<^sub>C G (D.cod \)\" proof - have 1: "src\<^sub>D \ = F\<^sub>0 (G\<^sub>0 (src\<^sub>D \)) \ trg\<^sub>D \ = F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \))" using assms G\<^sub>0_props by simp interpret hom\<^sub>C: subcategory V\<^sub>C \\\. \\ : G\<^sub>0 (src\<^sub>D \) \\<^sub>C G\<^sub>0 (trg\<^sub>D \)\\ using assms C.hhom_is_subcategory by simp interpret hom\<^sub>D: subcategory V\<^sub>D \\\'. \\' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D \)) \\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \))\\ using assms 1 D.in_hhom_def D.hhom_is_subcategory by simp interpret Faa': equivalence_pseudofunctor_at_hom V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \G\<^sub>0 (src\<^sub>D \)\ \G\<^sub>0 (trg\<^sub>D \)\ using assms G\<^sub>0_props(1) by unfold_locales auto have 2: "hom\<^sub>D.arr \" using assms 1 hom\<^sub>D.arr_char by simp show "\G \ : G (D.dom \) \\<^sub>C G (D.cod \)\" unfolding G_def using assms 2 hom\<^sub>C.arr_char hom\<^sub>C.dom_char hom\<^sub>C.cod_char hom\<^sub>D.dom_char hom\<^sub>D.cod_char Faa'.\\.F.preserves_arr Faa'.\\.F.preserves_dom Faa'.\\.F.preserves_cod by (intro C.in_homI) auto thus "C.in_hhom (G \) (G\<^sub>0 (src\<^sub>D \)) (G\<^sub>0 (trg\<^sub>D \))" proof - have "hom\<^sub>C.arr (Faa'.G\<^sub>1 \)" using 2 by simp thus ?thesis using G_def assms hom\<^sub>C.arrE by presburger qed qed lemma G_simps [simp]: assumes "D.arr \" shows "C.arr (G \)" and "src\<^sub>C (G \) = G\<^sub>0 (src\<^sub>D \)" and "trg\<^sub>C (G \) = G\<^sub>0 (trg\<^sub>D \)" and "C.dom (G \) = G (D.dom \)" and "C.cod (G \) = G (D.cod \)" using assms G_in_hom by auto interpretation G: "functor" V\<^sub>D V\<^sub>C 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 \ \ src\<^sub>D \ = src\<^sub>D \ \ trg\<^sub>D \ = trg\<^sub>D \ \ src\<^sub>D \ = F\<^sub>0 (G\<^sub>0 (src\<^sub>D \)) \ trg\<^sub>D \ = F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \)) \ src\<^sub>D \ = F\<^sub>0 (G\<^sub>0 (src\<^sub>D \)) \ trg\<^sub>D \ = F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \))" using \\ G\<^sub>0_props D.vseq_implies_hpar by auto interpret hom\<^sub>C: subcategory V\<^sub>C \\\. \\ : G\<^sub>0 (src\<^sub>D \) \\<^sub>C G\<^sub>0 (trg\<^sub>D \)\\ using 1 C.hhom_is_subcategory by simp interpret hom\<^sub>D: subcategory V\<^sub>D \\\'. \\' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D \)) \\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \))\\ using 1 D.hhom_is_subcategory by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def) interpret hom\<^sub>D': subcategory V\<^sub>D \\\'. D.arr \' \ src\<^sub>D \' = src\<^sub>D \ \ trg\<^sub>D \' = trg\<^sub>D \\ using 1 D.hhom_is_subcategory by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def) interpret Faa': equivalence_pseudofunctor_at_hom V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \G\<^sub>0 (src\<^sub>D \)\ \G\<^sub>0 (trg\<^sub>D \)\ using \\ 1 G\<^sub>0_props(1) by unfold_locales auto have 2: "hom\<^sub>D.seq \ \" using \\ 1 hom\<^sub>D.arr_char hom\<^sub>D.comp_def by fastforce have "G (hom\<^sub>D.comp \ \) = hom\<^sub>C.comp (G \) (G \)" unfolding G_def using 1 2 G\<^sub>0_props Faa'.G\<^sub>1_props hom\<^sub>D.arr_char Faa'.\\.F.as_nat_trans.preserves_comp_2 by auto thus "G (\ \\<^sub>D \) = G \ \\<^sub>C G \" using \\ 1 2 G_def hom\<^sub>C.comp_simp hom\<^sub>D.comp_simp D.src_vcomp D.trg_vcomp Faa'.\\.F.preserves_arr by metis qed lemma functor_G: shows "functor V\<^sub>D V\<^sub>C G" .. interpretation G: faithful_functor V\<^sub>D V\<^sub>C G proof fix f f' assume par: "D.par f f'" assume eq: "G f = G f'" have 1: "src\<^sub>D f = src\<^sub>D f' \ trg\<^sub>D f = trg\<^sub>D f'" using par by (metis D.src_dom D.trg_dom) interpret hom\<^sub>C: subcategory V\<^sub>C \\\. \\ : G\<^sub>0 (src\<^sub>D f) \\<^sub>C G\<^sub>0 (trg\<^sub>D f)\\ using par C.hhom_is_subcategory by simp interpret hom\<^sub>D: subcategory V\<^sub>D \\\'. \\' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D f)) \\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D 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 V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \G\<^sub>0 (src\<^sub>D f)\ \G\<^sub>0 (trg\<^sub>D f)\ using par G\<^sub>0_props(1) by unfold_locales auto interpret F': equivalence_pseudofunctor_at_hom V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \G\<^sub>0 (src\<^sub>D f')\ \G\<^sub>0 (trg\<^sub>D f')\ using par G\<^sub>0_props(1) by unfold_locales auto have 2: "hom\<^sub>D.par f f'" using par 1 hom\<^sub>D.arr_char hom\<^sub>D.dom_char hom\<^sub>D.cod_char G\<^sub>0_props by simp have "F.G\<^sub>1 f = F.G\<^sub>1 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 V\<^sub>D V\<^sub>C V\<^sub>D G F .. interpretation FG: faithful_functor V\<^sub>D V\<^sub>D "F o G" using faithful_functor_axioms G.faithful_functor_axioms faithful_functors_compose by blast interpretation GF: composite_functor V\<^sub>C V\<^sub>D V\<^sub>C F G .. interpretation GF: faithful_functor V\<^sub>C V\<^sub>C "G o F" using faithful_functor_axioms G.faithful_functor_axioms faithful_functors_compose by blast interpretation G: weak_arrow_of_homs V\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C src\<^sub>C trg\<^sub>C G proof have *: "\b. D.obj b \ C.isomorphic (G b) (src\<^sub>C (G b))" proof - fix b assume b: "D.obj b" interpret hom\<^sub>C: subcategory V\<^sub>C \\\. \\ : G\<^sub>0 b \\<^sub>C G\<^sub>0 b\\ using b C.hhom_is_subcategory by simp interpret hom\<^sub>D: subcategory V\<^sub>D \\\'. \\' : F\<^sub>0 (G\<^sub>0 b) \\<^sub>D F\<^sub>0 (G\<^sub>0 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 V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \G\<^sub>0 b\ \G\<^sub>0 b\ using b G\<^sub>0_props(1) by unfold_locales auto have 1: "C.in_hhom (G\<^sub>0 b) (G\<^sub>0 b) (G\<^sub>0 b)" using b G\<^sub>0_props by force text \ Using the unit constraints of \F\ and the fact that \F\<^sub>0 (G\<^sub>0 b) = b\, we obtain an isomorphism \\?\ : b \\<^sub>D F (G\<^sub>0 b)\\, which is also an isomorphism in \hom\<^sub>D\. \ let ?\ = "unit (G\<^sub>0 b)" have \: "\?\ : b \\<^sub>D F (G\<^sub>0 b)\ \ D.iso ?\" using b G\<^sub>0_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: "hom\<^sub>D.arr ?\ \ hom\<^sub>D.arr (D.inv ?\)" - proof - - have "D.in_hhom ?\ (F\<^sub>0 (G\<^sub>0 b)) (F\<^sub>0 (G\<^sub>0 b))" - using b \_in_hhom G\<^sub>0_props by simp - thus ?thesis - using \ by auto - qed + by (metis D.in_hhomE D.inv_in_hhom G\<^sub>0_props(1) \ \_in_hhom hom\<^sub>D.arr_char) have \': "hom\<^sub>D.in_hom ?\ b (Faa'.F\<^sub>1 (G\<^sub>0 b))" proof show "hom\<^sub>D.arr ?\" using 2 by simp show "hom\<^sub>D.dom ?\ = b" using 2 \ hom\<^sub>D.dom_char by auto show "hom\<^sub>D.cod ?\ = Faa'.F\<^sub>1 (G\<^sub>0 b)" proof - have "\?\ : b \\<^sub>D Faa'.F\<^sub>1 (G\<^sub>0 b)\ \ D.iso ?\" unfolding Faa'.F\<^sub>1_def using b \ 1 G\<^sub>0_props by auto thus ?thesis using 2 \ hom\<^sub>D.cod_char by auto qed qed text \ Transposing \?\\ via the adjunction \\\\ yields an isomorphism from \G b\ to \G\<^sub>0 b\ in \hom\<^sub>C\, hence in \C\. \ have **: "hom\<^sub>C.isomorphic (G b) (G\<^sub>0 b)" proof - have "hom\<^sub>C.in_hom (Faa'.\\.\ (G\<^sub>0 b) ?\) (Faa'.G\<^sub>1 b) (G\<^sub>0 b)" - proof - - have "hom\<^sub>D.in_hom ?\ b (Faa'.F\<^sub>1 (G\<^sub>0 b))" - using \' by simp - moreover have "hom\<^sub>C.ide (G\<^sub>0 b)" - using b 1 hom\<^sub>C.ide_char G\<^sub>0_props hom\<^sub>C.arr_char by auto - ultimately show ?thesis - using \ Faa'.\\.\_in_hom hom\<^sub>D.in_hom_char hom\<^sub>D.arr_char by simp - qed + by (metis "1" C.ide_src C.in_hhom_def Faa'.\\.\_in_hom \' hom\<^sub>C.ideI) moreover have "hom\<^sub>C.iso (Faa'.\\.\ (G\<^sub>0 b) ?\)" proof (unfold Faa'.\\.\_def) have "hom\<^sub>C.iso_in_hom (hom\<^sub>C.comp (Faa'.\ (G\<^sub>0 b)) (Faa'.G\<^sub>1 ?\)) (Faa'.G\<^sub>1 b) (G\<^sub>0 b)" proof (intro hom\<^sub>C.comp_iso_in_hom) show "hom\<^sub>C.iso_in_hom (Faa'.G\<^sub>1 (unit (G\<^sub>0 b))) (Faa'.G\<^sub>1 b) (Faa'.\\.FG.map (G\<^sub>0 b))" using \ \' 1 2 hom\<^sub>D.iso_char Faa'.\\.F.preserves_iso hom\<^sub>D.iso_char Faa'.\\.F.preserves_iso hom\<^sub>C.in_hom_char hom\<^sub>C.arr_char by auto show "hom\<^sub>C.iso_in_hom (Faa'.\ (G\<^sub>0 b)) (Faa'.\\.FG.map (G\<^sub>0 b)) (G\<^sub>0 b)" - proof - show "hom\<^sub>C.iso (Faa'.\ (G\<^sub>0 b))" - proof - - have "hom\<^sub>C.ide (G\<^sub>0 b)" - using b 1 G\<^sub>0_props hom\<^sub>C.ide_char hom\<^sub>C.arr_char by auto - thus ?thesis - using b G\<^sub>0_props Faa'.\\.\.components_are_iso by simp - qed - show "hom\<^sub>C.in_hom (Faa'.\ (G\<^sub>0 b)) (Faa'.\\.FG.map (G\<^sub>0 b)) (G\<^sub>0 b)" - using \ \' 1 - by (metis C.ide_src C.in_hhom_def Faa'.\\.\.preserves_hom hom\<^sub>C.arrI - hom\<^sub>C.ideI hom\<^sub>C.ide_in_hom hom\<^sub>C.map_simp) - qed + using 1 C.ide_src C.ide_trg C.in_hhom_def Faa'.\\.\.components_are_iso + Faa'.\\.\.preserves_hom hom\<^sub>C.arrI hom\<^sub>C.ideI hom\<^sub>C.ide_in_hom + hom\<^sub>C.map_simp hom\<^sub>C.iso_in_hom_def + by metis qed thus "hom\<^sub>C.iso (hom\<^sub>C.comp (Faa'.\ (G\<^sub>0 b)) (Faa'.G\<^sub>1 (unit (G\<^sub>0 b))))" by auto qed ultimately show ?thesis unfolding G_def using b hom\<^sub>C.isomorphic_def D.obj_def D.obj_simps(3) by auto qed hence "hom\<^sub>C.isomorphic (G b) (src\<^sub>C (G b))" using b G_in_hom(1) by auto moreover have "\f g. hom\<^sub>C.isomorphic f g \ C.isomorphic f g" using hom\<^sub>C.iso_char hom\<^sub>C.isomorphic_def C.isomorphic_def hom\<^sub>C.in_hom_char hom\<^sub>C.arr_char by auto ultimately show "C.isomorphic (G b) (src\<^sub>C (G b))" by simp qed fix \ assume \: "D.arr \" show "C.isomorphic (G (src\<^sub>D \)) (src\<^sub>C (G \))" using \ * by force show "C.isomorphic (G (trg\<^sub>D \)) (trg\<^sub>C (G \))" using \ * by force qed lemma weak_arrow_of_homs_G: shows "weak_arrow_of_homs V\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C src\<^sub>C trg\<^sub>C G" .. sublocale H\<^sub>DoGG: composite_functor D.VV.comp C.VV.comp V\<^sub>C G.FF \\\\. fst \\ \\<^sub>C snd \\\ .. sublocale GoH\<^sub>D: composite_functor D.VV.comp V\<^sub>D V\<^sub>C \\\\. fst \\ \\<^sub>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.\ V\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D src\<^sub>D trg\<^sub>D F (G\<^sub>0 (src\<^sub>D \)) (G\<^sub>0 (trg\<^sub>D \)) \ else D.null)" lemma \_in_hom: assumes "D.arr \" shows [intro]: "\\ \ : src\<^sub>D \ \\<^sub>D trg\<^sub>D \\" and [intro]: "\\ \ : D.dom \ \\<^sub>D F (G (D.cod \))\" and "D.ide \ \ D.iso (\ \)" proof - interpret hom\<^sub>C: subcategory V\<^sub>C \\\. \\ : G\<^sub>0 (src\<^sub>D \) \\<^sub>C G\<^sub>0 (trg\<^sub>D \)\\ using assms C.hhom_is_subcategory by simp interpret hom\<^sub>D: subcategory V\<^sub>D \\\'. \\' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D \)) \\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \))\\ using assms D.hhom_is_subcategory by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def) interpret Faa': equivalence_pseudofunctor_at_hom V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \G\<^sub>0 (src\<^sub>D \)\ \G\<^sub>0 (trg\<^sub>D \)\ using assms G\<^sub>0_props(1) by unfold_locales auto have 1: "\\ \ : D.dom \ \\<^sub>D F (G (D.cod \))\ \ (D.ide \ \ D.iso (\ \))" proof - have "src\<^sub>D \ = F\<^sub>0 (G\<^sub>0 (src\<^sub>D \)) \ trg\<^sub>D \ = F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \))" using assms G\<^sub>0_props by simp hence "hom\<^sub>D.arr \" using assms hom\<^sub>D.arr_char hom\<^sub>D.ide_char by simp hence "hom\<^sub>D.in_hom (Faa'.\ \) (D.dom \) (F (G (D.cod \))) \ (D.ide \ \ hom\<^sub>D.iso (Faa'.\ \))" using assms Faa'.\\.\.preserves_hom [of \ "D.dom \" "D.cod \"] hom\<^sub>D.map_def G_def Faa'.F\<^sub>1_def apply simp by (simp add: D.arr_iff_in_hom hom\<^sub>D.arr_char hom\<^sub>D.cod_closed hom\<^sub>D.dom_closed hom\<^sub>D.ide_char hom\<^sub>D.in_hom_char) thus ?thesis unfolding \_def using hom\<^sub>D.in_hom_char hom\<^sub>D.iso_char by auto qed show "\\ \ : D.dom \ \\<^sub>D F (G (D.cod \))\" using 1 by simp thus "D.in_hhom (\ \) (src\<^sub>D \) (trg\<^sub>D \)" 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 "src\<^sub>D (\ \) = src\<^sub>D \" and "trg\<^sub>D (\ \) = trg\<^sub>D \" 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 \) \\<^sub>D \ = \ \" and "F (G \) \\<^sub>D \ (D.dom \) = \ \" and "\ \\<^sub>D D.inv (\ (D.dom \)) = D.inv (\ (D.cod \)) \\<^sub>D F (G \)" proof - interpret hom\<^sub>C: subcategory V\<^sub>C \\\'. \\' : G\<^sub>0 (src\<^sub>D \) \\<^sub>C G\<^sub>0 (trg\<^sub>D \)\\ using assms C.hhom_is_subcategory by simp interpret hom\<^sub>D: subcategory V\<^sub>D \\\'. \\' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D \)) \\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \))\\ using D.hhom_is_subcategory by simp interpret Faa': equivalence_pseudofunctor_at_hom V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \G\<^sub>0 (src\<^sub>D \)\ \G\<^sub>0 (trg\<^sub>D \)\ using assms G\<^sub>0_props by unfold_locales auto show 1: "\ (D.cod \) \\<^sub>D \ = \ \" unfolding \_def using assms Faa'.\\.\.is_natural_2 hom\<^sub>D.comp_char G_def hom\<^sub>D.cod_simp G_in_hom(1) hom\<^sub>C.arr_char hom\<^sub>D.arr_char hom\<^sub>D.cod_closed apply simp by (metis (no_types, lifting) D.ext Faa'.\\.F.preserves_reflects_arr Faa'.\\.\.preserves_reflects_arr) show 2: "F (G \) \\<^sub>D \ (D.dom \) = \ \" unfolding \_def using assms D.src_dom D.src_cod D.trg_dom D.trg_cod Faa'.\\.\.is_natural_1 hom\<^sub>D.comp_char G_def Faa'.F\<^sub>1_def hom\<^sub>D.dom_simp hom\<^sub>D.cod_simp apply simp by (metis (no_types, lifting) D.not_arr_null Faa'.\\.\.is_extensional \_def \_simps(1) hom\<^sub>D.null_char) show "\ \\<^sub>D D.inv (\ (D.dom \)) = D.inv (\ (D.cod \)) \\<^sub>D F (G \)" using assms 1 2 by (simp add: D.invert_opposite_sides_of_square) qed text \ The fact that \G\<^sub>0\ is chosen to be right-inverse to \F\ implies that \\\ is an ordinary natural isomorphism (with respect to vertical composition) from \I\<^sub>D\ to \FG\. \ interpretation \: natural_isomorphism V\<^sub>D V\<^sub>D 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.\ V\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D src\<^sub>D trg\<^sub>D F (src\<^sub>C \) (trg\<^sub>C \) \ else C.null)" lemma \_in_hom: assumes "C.arr \" shows [intro]: "\\ \ : src\<^sub>C \ \\<^sub>C trg\<^sub>C \\" and [intro]: "\\ \ : G (F (C.dom \)) \\<^sub>C C.cod \\" and "C.ide \ \ C.iso (\ \)" proof - interpret hom\<^sub>C: subcategory V\<^sub>C \\\'. \\' : src\<^sub>C \ \\<^sub>C trg\<^sub>C \\\ using C.hhom_is_subcategory by simp interpret hom\<^sub>D: subcategory V\<^sub>D \\\'. \\' : F\<^sub>0 (src\<^sub>C \) \\<^sub>D F\<^sub>0 (trg\<^sub>C \)\\ using assms D.hhom_is_subcategory by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def) interpret Faa': equivalence_pseudofunctor_at_hom V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \G\<^sub>0 (F\<^sub>0 (src\<^sub>C \))\ \G\<^sub>0 (F\<^sub>0 (trg\<^sub>C \))\ using assms G\<^sub>0_props by unfold_locales simp_all have 1: "\\ \ : G (F (C.dom \)) \\<^sub>C C.cod \\ \ (C.ide \ \ C.iso (\ \))" proof - have "hom\<^sub>C.arr \" using assms hom\<^sub>C.arr_char by simp hence "hom\<^sub>C.in_hom \ (C.dom \) (C.cod \)" using assms hom\<^sub>C.in_hom_char hom\<^sub>C.dom_char hom\<^sub>C.cod_char by auto hence "hom\<^sub>C.in_hom (Faa'.\ \) (G (F (C.dom \))) (C.cod \) \ (C.ide \ \ hom\<^sub>C.iso (Faa'.\ \))" using assms Faa'.\\.\.preserves_hom [of \ "C.dom \" "C.cod \"] Faa'.\\.\.components_are_iso hom\<^sub>C.map_def G_def Faa'.F\<^sub>1_def G\<^sub>0_props by (simp add: hom\<^sub>C.ideI) thus ?thesis unfolding \_def using assms hom\<^sub>C.in_hom_char hom\<^sub>C.iso_char G\<^sub>0_props(2) by simp qed show "\\ \ : G (F (C.dom \)) \\<^sub>C C.cod \\" using 1 by simp thus "C.in_hhom (\ \) (src\<^sub>C \) (trg\<^sub>C \)" 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 "src\<^sub>C (\ \) = src\<^sub>C \" and "trg\<^sub>C (\ \) = trg\<^sub>C \" 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 \) \\<^sub>C G (F \) = \ \" and "\ \\<^sub>C \ (C.dom \) = \ \" and "G (F \) \\<^sub>C C.inv (\ (C.dom \)) = C.inv (\ (C.cod \)) \\<^sub>C \" proof - interpret hom\<^sub>C: subcategory V\<^sub>C \\\'. \\' : src\<^sub>C \ \\<^sub>C trg\<^sub>C \\\ using assms C.hhom_is_subcategory by simp interpret hom\<^sub>D: subcategory V\<^sub>D \\\. \\ : F\<^sub>0 (src\<^sub>C \) \\<^sub>D F\<^sub>0 (trg\<^sub>C \)\\ using assms D.hhom_is_subcategory by (simp add: D.in_hhom_def subcategory_axioms_def subcategory_def) interpret Faa': equivalence_pseudofunctor_at_hom V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \src\<^sub>C \\ \trg\<^sub>C \\ using assms by unfold_locales simp_all show 1: "\ (C.cod \) \\<^sub>C G (F \) = \ \" unfolding \_def using assms C.src_dom C.src_cod C.trg_dom C.trg_cod Faa'.\\.\.is_natural_2 hom\<^sub>C.comp_char G_def Faa'.F\<^sub>1_def G\<^sub>0_props hom\<^sub>C.dom_char hom\<^sub>C.cod_char apply simp by (metis (no_types, lifting) C.in_hhomI Faa'.\\.\.preserves_reflects_arr hom\<^sub>C.arr_char hom\<^sub>C.not_arr_null hom\<^sub>C.null_char) show 2: "\ \\<^sub>C \ (C.dom \) = \ \" unfolding \_def using assms C.src_dom C.src_cod C.trg_dom C.trg_cod Faa'.\\.\.is_natural_1 hom\<^sub>C.comp_char G_def Faa'.F\<^sub>1_def G\<^sub>0_props hom\<^sub>C.dom_char hom\<^sub>C.cod_char apply simp by (metis (no_types, lifting) C.in_hhomI Faa'.\\.\.preserves_reflects_arr hom\<^sub>C.arr_char hom\<^sub>C.not_arr_null hom\<^sub>C.null_char) show "G (F \) \\<^sub>C C.inv (\ (C.dom \)) = C.inv (\ (C.cod \)) \\<^sub>C \" using assms 1 2 C.invert_opposite_sides_of_square by simp qed interpretation \: natural_isomorphism V\<^sub>C V\<^sub>C GF.map C.map \ using \_simps \_naturality \_def by unfold_locales auto interpretation GFG: composite_functor V\<^sub>D V\<^sub>C V\<^sub>C G \GF.map\ .. interpretation FGF: composite_functor V\<^sub>C V\<^sub>D V\<^sub>D F \FG.map\ .. interpretation Go\: natural_transformation V\<^sub>D V\<^sub>C 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 V\<^sub>D V\<^sub>C G GFG.map (G \ \)" using \.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms horizontal_composite [of V\<^sub>D V\<^sub>D D.map FG.map \ V\<^sub>C G G G] by simp qed interpretation \oF: natural_transformation V\<^sub>C V\<^sub>D F FGF.map \\ \ F\ using \.natural_transformation_axioms as_nat_trans.natural_transformation_axioms horizontal_composite [of V\<^sub>C V\<^sub>D F F F V\<^sub>D D.map FG.map \] by simp interpretation \oG: natural_transformation V\<^sub>D V\<^sub>C GFG.map G \\ \ G\ using \.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms horizontal_composite [of V\<^sub>D V\<^sub>C G G G V\<^sub>C GF.map C.map \] by simp interpretation Fo\: natural_transformation V\<^sub>C V\<^sub>D 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 V\<^sub>C V\<^sub>D FGF.map F (F \ \)" using \.natural_transformation_axioms as_nat_trans.natural_transformation_axioms horizontal_composite [of V\<^sub>C V\<^sub>C GF.map C.map \ V\<^sub>D F F F] by simp qed interpretation \oG_Go\: vertical_composite V\<^sub>D V\<^sub>C G GFG.map G \G \ \\ \\ \ G\ .. interpretation Fo\_\oF: vertical_composite V\<^sub>C V\<^sub>D F FGF.map F \\ \ F\ \F \ \\ .. text \ Bijectivity results in an ordinary adjunction between the vertical categories. \ lemma adjunction_\\: shows "unit_counit_adjunction V\<^sub>C V\<^sub>D 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 hom\<^sub>C: subcategory V\<^sub>C \\\'. \\' : G\<^sub>0 (src\<^sub>D \) \\<^sub>C G\<^sub>0 (trg\<^sub>D \)\\ using \ C.hhom_is_subcategory by simp interpret hom\<^sub>D: subcategory V\<^sub>D \\\'. \\' : F\<^sub>0 (G\<^sub>0 (src\<^sub>D \)) \\<^sub>D F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \))\\ using D.hhom_is_subcategory by simp interpret Faa': equivalence_pseudofunctor_at_hom V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \G\<^sub>0 (src\<^sub>D \)\ \G\<^sub>0 (trg\<^sub>D \)\ using \ G\<^sub>0_props by unfold_locales simp_all show "\oG_Go\.map \ = G \" proof - have 1: "C.arr (Faa'.G\<^sub>1 (D.cod \))" using \ hom\<^sub>D.arr_char G_def by (metis Faa'.\\.F.preserves_reflects_arr G_in_hom(1) hom\<^sub>C.arr_char hom\<^sub>C.inclusion hom\<^sub>D.cod_closed) have 2: "D.arr (Faa'.\ \)" using \ hom\<^sub>D.arr_char by (metis D.in_hhomI D.obj_src D.obj_trg Faa'.\\.\.preserves_reflects_arr G\<^sub>0_props(1) hom\<^sub>D.inclusion) have 3: "src\<^sub>C (Faa'.G\<^sub>1 (D.cod \)) = G\<^sub>0 (src\<^sub>D \)" using \ by (metis D.arr_cod D.src_cod D.trg_cod G_def G_simps(2)) have 4: "trg\<^sub>C (Faa'.G\<^sub>1 (D.cod \)) = G\<^sub>0 (trg\<^sub>D \)" using \ by (metis D.arr_cod D.src_cod D.trg_cod G_def G_simps(3)) have 5: "hom\<^sub>D.arr \" using \ hom\<^sub>D.arr_char G\<^sub>0_props by simp have "\oG_Go\.map \ = \ (G (D.cod \)) \\<^sub>C G (\ \)" using \ \oG_Go\.map_def by simp also have "... = Faa'.\ (Faa'.G\<^sub>1 (D.cod \)) \\<^sub>C Faa'.G\<^sub>1 (Faa'.\ \)" using \ 1 2 3 4 \_def \_def G_def G_simps(2-3) \_simps(2-3) by auto also have "... = hom\<^sub>C.comp ((Faa'.\ \ Faa'.G\<^sub>1) (hom\<^sub>D.cod \)) ((Faa'.G\<^sub>1 \ Faa'.\) \)" proof - have "Faa'.\ (Faa'.G\<^sub>1 (D.cod \)) \\<^sub>C Faa'.G\<^sub>1 (Faa'.\ \) = Faa'.\\.\FoF\.map \" proof - have "C.seq (Faa'.\ (Faa'.G\<^sub>1 (hom\<^sub>D.cod \))) (Faa'.G\<^sub>1 (Faa'.\ \))" proof - have "hom\<^sub>C.seq (Faa'.\ (Faa'.G\<^sub>1 (hom\<^sub>D.cod \))) (Faa'.G\<^sub>1 (Faa'.\ \))" using \ 5 hom\<^sub>C.arr_char hom\<^sub>C.seq_char Faa'.\\.F.preserves_arr apply (intro hom\<^sub>C.seqI) apply auto using Faa'.\\.\.preserves_reflects_arr hom\<^sub>C.inclusion hom\<^sub>D.arr_cod by presburger thus ?thesis using hom\<^sub>C.seq_char by simp qed moreover have "D.in_hhom (Faa'.\ \) (F\<^sub>0 (G\<^sub>0 (src\<^sub>D \))) (F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \)))" using \ 5 Faa'.\\.\.preserves_reflects_arr hom\<^sub>D.arrE by blast moreover have "D.in_hhom (D.cod \) (F\<^sub>0 (G\<^sub>0 (src\<^sub>D \))) (F\<^sub>0 (G\<^sub>0 (trg\<^sub>D \)))" using \ 5 hom\<^sub>D.arrE hom\<^sub>D.cod_closed by blast ultimately show ?thesis using \ 5 Faa'.\\.\FoF\.map_def hom\<^sub>D.arr_char hom\<^sub>C.comp_char hom\<^sub>D.cod_char 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 hom\<^sub>C: subcategory V\<^sub>C \\\'. \\' : src\<^sub>C \ \\<^sub>C trg\<^sub>C \\\ using \ C.hhom_is_subcategory by simp interpret hom\<^sub>D: subcategory V\<^sub>D \\\. \\ : F\<^sub>0 (src\<^sub>C \) \\<^sub>D F\<^sub>0 (trg\<^sub>C \)\\ using \ D.in_hhom_def D.hhom_is_subcategory by simp interpret Faa': equivalence_pseudofunctor_at_hom V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \src\<^sub>C \\ \trg\<^sub>C \\ using \ by unfold_locales auto show "Fo\_\oF.map \ = F \" proof - have 1: "hom\<^sub>C.arr \" using \ hom\<^sub>C.arr_char by simp have "Fo\_\oF.map \ = F (\ (C.cod \)) \\<^sub>D \ (F \)" using \ Fo\_\oF.map_def by simp also have "... = Faa'.F\<^sub>1 (Faa'.\ (C.cod \)) \\<^sub>D Faa'.\ (Faa'.F\<^sub>1 \)" unfolding \_def \_def G_def using \ G\<^sub>0_props Faa'.F\<^sub>1_def by auto also have "... = hom\<^sub>D.comp ((Faa'.F\<^sub>1 \ Faa'.\) (hom\<^sub>C.cod \)) ((Faa'.\ \ Faa'.F\<^sub>1) \)" proof - have 2: "hom\<^sub>C.arr (C.cod \)" using 1 hom\<^sub>C.cod_char [of \] hom\<^sub>C.arr_cod hom\<^sub>C.dom_char hom\<^sub>C.cod_char by simp moreover have "D.seq (Faa'.F\<^sub>1 (Faa'.\ (C.cod \))) (Faa'.\ (Faa'.F\<^sub>1 \))" proof - have "hom\<^sub>D.seq (Faa'.F\<^sub>1 (Faa'.\ (C.cod \))) (Faa'.\ (Faa'.F\<^sub>1 \))" using \ 1 2 hom\<^sub>D.arr_char hom\<^sub>D.seq_char Faa'.\\.G.preserves_arr hom\<^sub>C.cod_simp hom\<^sub>C.dom_cod apply (intro hom\<^sub>D.seqI) by auto metis thus ?thesis using hom\<^sub>D.seq_char by simp qed ultimately show ?thesis using 1 hom\<^sub>D.comp_char hom\<^sub>C.dom_char hom\<^sub>C.cod_char 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 hom\<^sub>C.arr_char Faa'.\\.\FoF\.map_def hom\<^sub>D.arr_char hom\<^sub>C.comp_char hom\<^sub>D.comp_char Faa'.F\<^sub>1_def by simp finally show ?thesis by simp qed qed ultimately show "Fo\_\oF.map \ = F \" by auto qed qed interpretation \\: unit_counit_adjunction V\<^sub>C V\<^sub>D 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 \\<^sub>G\<^sub>0 where "\\<^sub>G\<^sub>0 f f' = C.inv (\ (G f \\<^sub>C G f') \\<^sub>C G (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')))" lemma \\<^sub>G\<^sub>0_in_hom: assumes "D.ide f" and "D.ide f'" and "src\<^sub>D f = trg\<^sub>D f'" shows "\\\<^sub>G\<^sub>0 f f' : G f \\<^sub>C G f' \\<^sub>C G (f \\<^sub>D f')\" proof (unfold \\<^sub>G\<^sub>0_def, intro C.inv_in_hom) show 1: "\\ (G f \\<^sub>C G f') \\<^sub>C G (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')) : G (f \\<^sub>D f') \\<^sub>C G f \\<^sub>C G f'\" proof show "\\ (G f \\<^sub>C G f') : G (F (G f \\<^sub>C G f')) \\<^sub>C G f \\<^sub>C G f'\" using assms \_in_hom by auto show "\G (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')) : G (f \\<^sub>D f') \\<^sub>C G (F (G f \\<^sub>C G f'))\" proof - have "\\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f') : f \\<^sub>D f' \\<^sub>D F (G f \\<^sub>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 \\<^sub>C G f') \\<^sub>C G (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')))" proof (intro C.isos_compose) show "C.iso (G (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')))" proof - have "D.iso (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f'))" using assms \_in_hom by (intro D.isos_compose D.seqI) auto thus ?thesis by simp qed show "C.iso (\ (G f \\<^sub>C G f'))" using assms \_in_hom by simp show "C.seq (\ (G f \\<^sub>C G f')) (G (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')))" using 1 by auto qed qed lemma iso_\\<^sub>G\<^sub>0: assumes "D.ide f" and "D.ide f'" and "src\<^sub>D f = trg\<^sub>D f'" shows "C.iso (\\<^sub>G\<^sub>0 f f')" proof (unfold \\<^sub>G\<^sub>0_def, intro C.iso_inv_iso C.isos_compose) show 1: "C.iso (G (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')))" using assms \_in_hom D.in_hhom_def cmp_simps'(1) cmp_simps(4) by auto show "C.iso (\ (G f \\<^sub>C G f'))" using assms \_in_hom by simp show "C.seq (\ (G f \\<^sub>C G f')) (G (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')))" using assms 1 by force qed lemma \\<^sub>G\<^sub>0_naturality: assumes "D.arr \" and "D.arr \'" and "src\<^sub>D \ = trg\<^sub>D \'" shows "\\<^sub>G\<^sub>0 (D.cod \) (D.cod \') \\<^sub>C (G \ \\<^sub>C G \') = G (\ \\<^sub>D \') \\<^sub>C \\<^sub>G\<^sub>0 (D.dom \) (D.dom \')" proof - let ?X = "\ (G (D.cod \) \\<^sub>C G (D.cod \')) \\<^sub>C G (\ (G (D.cod \), G (D.cod \')) \\<^sub>D (\ (D.cod \) \\<^sub>D \ (D.cod \')))" have iso_X: "C.iso ?X" using assms C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def by (intro C.isos_compose) auto have "\\<^sub>G\<^sub>0 (D.cod \) (D.cod \') \\<^sub>C (G \ \\<^sub>C G \') = C.inv ?X \\<^sub>C (G \ \\<^sub>C G \')" unfolding \\<^sub>G\<^sub>0_def by simp also have "... = (C.inv (G (\ (G (D.cod \), G (D.cod \')) \\<^sub>D (\ (D.cod \) \\<^sub>D \ (D.cod \')))) \\<^sub>C C.inv (\ (G (D.cod \) \\<^sub>C G (D.cod \')))) \\<^sub>C (G \ \\<^sub>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 \'))) \\<^sub>C G (\ (D.cod \) \\<^sub>D \ (D.cod \'))) \\<^sub>C C.inv (\ (G (D.cod \) \\<^sub>C G (D.cod \')))) \\<^sub>C (G \ \\<^sub>C G \')" proof - have "G (\ (G (D.cod \), G (D.cod \')) \\<^sub>D (\ (D.cod \) \\<^sub>D \ (D.cod \'))) = G (\ (G (D.cod \), G (D.cod \'))) \\<^sub>C G (\ (D.cod \) \\<^sub>D \ (D.cod \'))" using assms iso_X \_in_hom by fast thus ?thesis by simp qed also have "... = C.inv (G (\ (D.cod \) \\<^sub>D \ (D.cod \'))) \\<^sub>C C.inv (G (\ (G (D.cod \), G (D.cod \')))) \\<^sub>C C.inv (\ (G (D.cod \) \\<^sub>C G (D.cod \'))) \\<^sub>C (G \ \\<^sub>C G \')" proof - have "C.iso (G (\ (G (D.cod \), G (D.cod \'))) \\<^sub>C G (\ (D.cod \) \\<^sub>D \ (D.cod \')))" using assms iso_X \_in_hom C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def by (intro C.isos_compose C.seqI) auto hence "C.inv (G (\ (G (D.cod \), G (D.cod \'))) \\<^sub>C G (\ (D.cod \) \\<^sub>D \ (D.cod \'))) = C.inv (G (\ (D.cod \) \\<^sub>D \ (D.cod \'))) \\<^sub>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 \)) \\<^sub>D D.inv (\ (D.cod \'))) \\<^sub>C C.inv (G (\ (G (D.cod \), G (D.cod \')))) \\<^sub>C C.inv (\ (G (D.cod \) \\<^sub>C G (D.cod \'))) \\<^sub>C (G \ \\<^sub>C G \')" proof - have "C.inv (G (\ (D.cod \) \\<^sub>D \ (D.cod \'))) = G (D.inv (\ (D.cod \)) \\<^sub>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 \)) \\<^sub>D D.inv (\ (D.cod \'))) \\<^sub>C (C.inv (G (\ (G (D.cod \), G (D.cod \')))) \\<^sub>C G (F (G \ \\<^sub>C G \'))) \\<^sub>C C.inv (\ (G (D.dom \) \\<^sub>C G (D.dom \')))" proof - have "C.inv (\ (G (D.cod \) \\<^sub>C G (D.cod \'))) \\<^sub>C (G \ \\<^sub>C G \') = G (F (G \ \\<^sub>C G \')) \\<^sub>C C.inv (\ (G (D.dom \) \\<^sub>C G (D.dom \')))" proof - have "G (D.dom \) \\<^sub>C G (D.dom \') = C.dom (G \ \\<^sub>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 \)) \\<^sub>D D.inv (\ (D.cod \'))) \\<^sub>C G (F (G \) \\<^sub>D F (G \'))) \\<^sub>C G (D.inv (\ (C.dom (G \), C.dom (G \')))) \\<^sub>C C.inv (\ (G (D.dom \) \\<^sub>C G (D.dom \')))" proof - have "C.inv (G (\ (G (D.cod \), G (D.cod \')))) \\<^sub>C G (F (G \ \\<^sub>C G \')) = G (D.inv (\ (C.cod (G \), C.cod (G \'))) \\<^sub>D F (G \ \\<^sub>C G \'))" using assms by (simp add: G.preserves_inv) also have "... = G ((F (G \) \\<^sub>D F (G \')) \\<^sub>D D.inv (\ (C.dom (G \), C.dom (G \'))))" using assms C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def \.inv_naturality [of "(G \, G \')"] by simp also have "... = G (F (G \) \\<^sub>D F (G \')) \\<^sub>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 (\ \\<^sub>D \') \\<^sub>C G (D.inv (\ (D.dom \)) \\<^sub>D D.inv (\ (D.dom \'))) \\<^sub>C G (D.inv (\ (C.dom (G \), C.dom (G \')))) \\<^sub>C C.inv (\ (G (D.dom \) \\<^sub>C G (D.dom \')))" proof - have "G (D.inv (\ (D.cod \)) \\<^sub>D D.inv (\ (D.cod \'))) \\<^sub>C G (F (G \) \\<^sub>D F (G \')) = G ((D.inv (\ (D.cod \)) \\<^sub>D D.inv (\ (D.cod \'))) \\<^sub>D (F (G \) \\<^sub>D F (G \')))" using assms by simp also have "... = G (D.inv (\ (D.cod \)) \\<^sub>D F (G \) \\<^sub>D D.inv (\ (D.cod \')) \\<^sub>D F (G \'))" using assms D.interchange by simp also have "... = G (\ \\<^sub>D D.inv (\ (D.dom \)) \\<^sub>D \' \\<^sub>D D.inv (\ (D.dom \')))" using assms \_naturality by simp also have "... = G ((\ \\<^sub>D \') \\<^sub>D (D.inv (\ (D.dom \)) \\<^sub>D D.inv (\ (D.dom \'))))" using assms D.interchange by simp also have "... = G (\ \\<^sub>D \') \\<^sub>C G (D.inv (\ (D.dom \)) \\<^sub>D D.inv (\ (D.dom \')))" using assms by simp finally show ?thesis using C.comp_assoc by simp qed also have "... = G (\ \\<^sub>D \') \\<^sub>C \\<^sub>G\<^sub>0 (D.dom \) (D.dom \')" proof - have "G (D.inv (\ (D.dom \)) \\<^sub>D D.inv (\ (D.dom \'))) \\<^sub>C G (D.inv (\ (C.dom (G \), C.dom (G \')))) \\<^sub>C C.inv (\ (G (D.dom \) \\<^sub>C G (D.dom \'))) = C.inv (G (\ (D.dom \) \\<^sub>D \ (D.dom \'))) \\<^sub>C C.inv (G (\ (C.dom (G \), C.dom (G \')))) \\<^sub>C C.inv (\ (G (D.dom \) \\<^sub>C G (D.dom \')))" proof - have "G (D.inv (\ (D.dom \)) \\<^sub>D D.inv (\ (D.dom \'))) = C.inv (G (\ (D.dom \) \\<^sub>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 \) \\<^sub>C G (D.dom \')) \\<^sub>C G (\ (C.dom (G \), C.dom (G \'))) \\<^sub>C G (\ (D.dom \) \\<^sub>D \ (D.dom \')))" proof - have "C.iso (\ (G (D.dom \) \\<^sub>C G (D.dom \')) \\<^sub>C G (\ (C.dom (G \), C.dom (G \'))) \\<^sub>C G (\ (D.dom \) \\<^sub>D \ (D.dom \')))" using assms \_in_hom \_in_hom cmp_simps'(1,4) C.VV.cod_char 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 \) \\<^sub>C G (D.dom \')) \\<^sub>C G (\ (C.dom (G \), C.dom (G \')) \\<^sub>D (\ (D.dom \) \\<^sub>D \ (D.dom \'))))" proof - have "G (\ (C.dom (G \), C.dom (G \'))) \\<^sub>C G (\ (D.dom \) \\<^sub>D \ (D.dom \')) = G (\ (C.dom (G \), C.dom (G \')) \\<^sub>D (\ (D.dom \) \\<^sub>D \ (D.dom \')))" using assms cmp_simps'(1,4) by auto thus ?thesis by simp qed also have "... = \\<^sub>G\<^sub>0 (D.dom \) (D.dom \')" unfolding \\<^sub>G\<^sub>0_def using assms by simp finally show ?thesis by simp qed finally show ?thesis by simp qed interpretation \\<^sub>G: transformation_by_components D.VV.comp V\<^sub>C H\<^sub>DoGG.map GoH\<^sub>D.map \\fg. \\<^sub>G\<^sub>0 (fst fg) (snd fg)\ using D.VV.ide_char D.VV.arr_char \\<^sub>G\<^sub>0_in_hom \\<^sub>G\<^sub>0_naturality G.FF_def D.VV.dom_char D.VV.cod_char by unfold_locales auto interpretation \\<^sub>G: natural_isomorphism D.VV.comp V\<^sub>C H\<^sub>DoGG.map GoH\<^sub>D.map \\<^sub>G.map using \\<^sub>G.map_simp_ide D.VV.ide_char D.VV.arr_char iso_\\<^sub>G\<^sub>0 by unfold_locales simp abbreviation \\<^sub>G where "\\<^sub>G \ \\<^sub>G.map" lemma \\<^sub>G_in_hom [intro]: assumes "D.arr \" and "D.arr \'" and "src\<^sub>D \ = trg\<^sub>D \'" shows "C.in_hhom (\\<^sub>G.map (\, \')) (src\<^sub>C (G (D.dom \'))) (trg\<^sub>C (G (D.cod \)))" and "\\\<^sub>G.map (\, \') : G (D.dom \) \\<^sub>C G (D.dom \') \\<^sub>C G (D.cod \ \\<^sub>D D.cod \')\" proof - show "\\\<^sub>G.map (\, \') : G (D.dom \) \\<^sub>C G (D.dom \') \\<^sub>C G (D.cod \ \\<^sub>D D.cod \')\" proof - have "\\<^sub>G.map (\, \') = \\<^sub>G\<^sub>0 (D.cod \) (D.cod \') \\<^sub>C (G \ \\<^sub>C G \')" using assms D.VV.arr_char \\<^sub>G.map_def \\<^sub>G\<^sub>0_in_hom G.FF_def D.VV.cod_char by simp moreover have "\\\<^sub>G\<^sub>0 (D.cod \) (D.cod \') \\<^sub>C (G \ \\<^sub>C G \') : G (D.dom \) \\<^sub>C G (D.dom \') \\<^sub>C G (D.cod \ \\<^sub>D D.cod \')\" using assms \\<^sub>G\<^sub>0_in_hom by (intro C.comp_in_homI) auto ultimately show ?thesis by simp qed thus "C.in_hhom (\\<^sub>G.map (\, \')) (src\<^sub>C (G (D.dom \'))) (trg\<^sub>C (G (D.cod \)))" using assms C.vconn_implies_hpar(1-2) by auto qed lemma \\<^sub>G_simps [simp]: assumes "D.arr \" and "D.arr \'" and "src\<^sub>D \ = trg\<^sub>D \'" shows "C.arr (\\<^sub>G.map (\, \'))" and "src\<^sub>C (\\<^sub>G.map (\, \')) = src\<^sub>C (G (D.dom \'))" and "trg\<^sub>C (\\<^sub>G.map (\, \')) = trg\<^sub>C (G (D.cod \))" and "C.dom (\\<^sub>G.map (\, \')) = G (D.dom \) \\<^sub>C G (D.dom \')" and "C.cod (\\<^sub>G.map (\, \')) = G (D.cod \ \\<^sub>D D.cod \')" using assms \\<^sub>G_in_hom apply auto by fast+ lemma \\<^sub>G_map_simp_ide: assumes "D.ide f" and "D.ide f'" and "src\<^sub>D f = trg\<^sub>D f'" shows "\\<^sub>G.map (f, f') = G (D.inv (\ f) \\<^sub>D D.inv (\ f')) \\<^sub>C G (D.inv (\ (G f, G f'))) \\<^sub>C C.inv (\ (G f \\<^sub>C G f'))" proof - have "\\<^sub>G.map (f, f') = C.inv (\ (G f \\<^sub>C G f') \\<^sub>C G (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')))" using assms \\<^sub>G.map_simp_ide D.VV.ide_char D.VV.arr_char \\<^sub>G\<^sub>0_def G.FF_def by auto also have "... = C.inv (\ (G f \\<^sub>C G f') \\<^sub>C G (\ (G f, G f')) \\<^sub>C G (\ f \\<^sub>D \ f'))" using assms D.VV.ide_char D.VV.arr_char cmp_simps(1,4) by simp also have "... = C.inv (G (\ f \\<^sub>D \ f')) \\<^sub>C C.inv (G (\ (G f, G f'))) \\<^sub>C C.inv (\ (G f \\<^sub>C G f'))" proof - have "C.iso (\ (G f \\<^sub>C G f') \\<^sub>C G (\ (G f, G f')) \\<^sub>C G (\ f \\<^sub>D \ f'))" using assms C.VV.ide_char C.VV.arr_char 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) \\<^sub>D D.inv (\ f')) \\<^sub>C G (D.inv (\ (G f, G f'))) \\<^sub>C C.inv (\ (G f \\<^sub>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 "src\<^sub>D f = trg\<^sub>D f'" shows "\ (f \\<^sub>D f') = F (\\<^sub>G.map (f, f')) \\<^sub>D \ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')" proof - have 1: "\F (\\<^sub>G.map (f, f')) \\<^sub>D \ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f') : f \\<^sub>D f' \\<^sub>D F (G (f \\<^sub>D f'))\" using assms by fastforce have 2: "\\ (G f \\<^sub>C G f') \\<^sub>C G (\ (G f, G f')) \\<^sub>C G (\ f \\<^sub>D \ f') : G (f \\<^sub>D f') \\<^sub>C G f \\<^sub>C G f'\" using assms G.preserves_hom cmp_in_hom(2) by (intro C.comp_in_homI) auto have "F (\\<^sub>G.map (f, f')) \\<^sub>D \ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f') = F (C.inv (\ (G f \\<^sub>C G f') \\<^sub>C G (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')))) \\<^sub>D \ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')" using assms \\<^sub>G.map_simp_ide D.VV.ide_char D.VV.arr_char \\<^sub>G\<^sub>0_def by simp also have "... = F (C.inv (G (\ f \\<^sub>D \ f')) \\<^sub>C C.inv (G (\ (G f, G f'))) \\<^sub>C C.inv (\ (G f \\<^sub>C G f'))) \\<^sub>D \ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')" proof - have "C.inv (\ (G f \\<^sub>C G f') \\<^sub>C G (\ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f'))) = C.inv (\ (G f \\<^sub>C G f') \\<^sub>C G (\ (G f, G f')) \\<^sub>C G (\ f \\<^sub>D \ f'))" using assms 1 by force also have "... = C.inv (G (\ f \\<^sub>D \ f')) \\<^sub>C C.inv (G (\ (G f, G f'))) \\<^sub>C C.inv (\ (G f \\<^sub>C G f'))" proof - have "C.iso (\ (G f \\<^sub>C G f') \\<^sub>C G (\ (G f, G f')) \\<^sub>C G (\ f \\<^sub>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 \\<^sub>D \ f'))) \\<^sub>D F (C.inv (G (\ (G f, G f')))) \\<^sub>D F (C.inv (\ (G f \\<^sub>C G f'))) \\<^sub>D \ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')" proof - have "C.arr ((C.inv (G (\ f \\<^sub>D \ f')) \\<^sub>C C.inv (G (\ (G f, G f'))) \\<^sub>C C.inv (\ (G f \\<^sub>C G f'))))" using assms 1 2 cmp_components_are_iso C.VV.ide_char C.VV.arr_char 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 \\<^sub>D \ f'))) \\<^sub>D D.inv (F (G (\ (G f, G f')))) \\<^sub>D D.inv (F (\ (G f \\<^sub>C G f'))) \\<^sub>D \ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')" using assms by (simp add: preserves_inv) also have "... = D.inv ((\ (F (G f) \\<^sub>D F (G f'))) \\<^sub>D (\ f \\<^sub>D \ f') \\<^sub>D D.inv (\ (f \\<^sub>D f'))) \\<^sub>D D.inv (\ (F (G f \\<^sub>C G f')) \\<^sub>D \ (G f, G f') \\<^sub>D D.inv (\ (F (G f) \\<^sub>D F (G f')))) \\<^sub>D D.inv (F (\ (G f \\<^sub>C G f'))) \\<^sub>D \ (G f, G f') \\<^sub>D (\ f \\<^sub>D \ f')" proof - have *: "\\. D.arr \ \ \ (D.cod \) \\<^sub>D \ \\<^sub>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 \\<^sub>D \ f')) = \ (F (G f) \\<^sub>D F (G f')) \\<^sub>D (\ f \\<^sub>D \ f') \\<^sub>D D.inv (\ (f \\<^sub>D f'))" using assms * [of "\ f \\<^sub>D \ f'"] by simp moreover have "F (G (\ (G f, G f'))) = \ (F (G f \\<^sub>C G f')) \\<^sub>D \ (G f, G f') \\<^sub>D D.inv (\ (F (G f) \\<^sub>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 \\<^sub>D f') \\<^sub>D D.inv (\ f \\<^sub>D \ f') \\<^sub>D ((D.inv (\ (F (G f) \\<^sub>D F (G f')))) \\<^sub>D (\ (F (G f) \\<^sub>D F (G f')))) \\<^sub>D (D.inv (\ (G f, G f')) \\<^sub>D ((D.inv (\ (F (G f \\<^sub>C G f')))) \\<^sub>D D.inv (F (\ (G f \\<^sub>C G f')))) \\<^sub>D \ (G f, G f')) \\<^sub>D (\ f \\<^sub>D \ f')" proof - have "D.inv ((\ (F (G f) \\<^sub>D F (G f'))) \\<^sub>D (\ f \\<^sub>D \ f') \\<^sub>D D.inv (\ (f \\<^sub>D f'))) = \ (f \\<^sub>D f') \\<^sub>D D.inv (\ f \\<^sub>D \ f') \\<^sub>D D.inv (\ (F (G f) \\<^sub>D F (G f')))" proof - have "D.iso (\ (F (G f) \\<^sub>D F (G f')) \\<^sub>D (\ f \\<^sub>D \ f') \\<^sub>D D.inv (\ (f \\<^sub>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 \\<^sub>C G f')) \\<^sub>D \ (G f, G f') \\<^sub>D D.inv (\ (F (G f) \\<^sub>D F (G f')))) = \ (F (G f) \\<^sub>D F (G f')) \\<^sub>D D.inv (\ (G f, G f')) \\<^sub>D D.inv (\ (F (G f \\<^sub>C G f')))" proof - have "D.iso (\ (F (G f \\<^sub>C G f')) \\<^sub>D \ (G f, G f') \\<^sub>D D.inv (\ (F (G f) \\<^sub>D F (G f'))))" proof - have 1: "D.seq (\ (G f, G f')) (D.inv (\ (F (G f) \\<^sub>D F (G f'))))" using assms by (intro D.seqI) auto moreover have "D.seq (\ (F (G f \\<^sub>C G f'))) (\ (G f, G f') \\<^sub>D D.inv (\ (F (G f) \\<^sub>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 \\<^sub>D f') \\<^sub>D D.inv (\ f \\<^sub>D \ f') \\<^sub>D (D.inv (\ (F (G f) \\<^sub>D F (G f'))) \\<^sub>D (\ (F (G f) \\<^sub>D F (G f')))) \\<^sub>D (D.inv (\ (G f, G f')) \\<^sub>D \ (G f, G f')) \\<^sub>D (\ f \\<^sub>D \ f')" proof - have "(D.inv (\ (F (G f \\<^sub>C G f')))) \\<^sub>D D.inv (F (\ (G f \\<^sub>C G f'))) \\<^sub>D \ (G f, G f') = F (G f \\<^sub>C G f') \\<^sub>D \ (G f, G f')" proof - have "(D.inv (\ (F (G f \\<^sub>C G f')))) \\<^sub>D D.inv (F (\ (G f \\<^sub>C G f'))) = D.inv (F (\ (G f \\<^sub>C G f')) \\<^sub>D \ (F (G f \\<^sub>C G f')))" using assms by (simp add: D.inv_comp) also have "... = F (G f \\<^sub>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 \\<^sub>C G f')))) \\<^sub>D D.inv (F (\ (G f \\<^sub>C G f'))) \\<^sub>D \ (G f, G f') = \ (G f, G f')" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = \ (f \\<^sub>D f') \\<^sub>D D.inv (\ f \\<^sub>D \ f') \\<^sub>D (D.inv (\ (F (G f) \\<^sub>D F (G f'))) \\<^sub>D (\ (F (G f) \\<^sub>D F (G f')))) \\<^sub>D (\ f \\<^sub>D \ f')" using assms D.comp_cod_arr D.comp_inv_arr' cmp_components_are_iso by simp also have "... = \ (f \\<^sub>D f') \\<^sub>D D.inv (\ f \\<^sub>D \ f') \\<^sub>D (\ f \\<^sub>D \ f')" using assms \_in_hom D.comp_inv_arr' D.comp_cod_arr by simp also have "... = \ (f \\<^sub>D f')" using assms D.comp_inv_arr' [of "\ f \\<^sub>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 "src\<^sub>C g = trg\<^sub>C g'" shows "\ (g \\<^sub>C g') = (\ g \\<^sub>C \ g') \\<^sub>C C.inv (\\<^sub>G.map (F g, F g')) \\<^sub>C C.inv (G (\ (g, g')))" proof - have 1: "\\ (G (F g) \\<^sub>C G (F g')) \\<^sub>C G (\ (G (F g), G (F g')) \\<^sub>D (\ (F g) \\<^sub>D \ (F g'))) : G (F g \\<^sub>D F g') \\<^sub>C G (F g) \\<^sub>C G (F g')\" using assms by fastforce have "(\ g \\<^sub>C \ g') \\<^sub>C C.inv (\\<^sub>G.map (F g, F g')) \\<^sub>C C.inv (G (\ (g, g'))) = (\ g \\<^sub>C \ g') \\<^sub>C C.inv (C.inv (\ (G (F g) \\<^sub>C G (F g')) \\<^sub>C G (\ (G (F g), G (F g')) \\<^sub>D (\ (F g) \\<^sub>D \ (F g'))))) \\<^sub>C C.inv (G (\ (g, g')))" using assms \\<^sub>G.map_simp_ide D.VV.ide_char D.VV.arr_char \\<^sub>G\<^sub>0_def by simp also have "... = ((\ g \\<^sub>C \ g') \\<^sub>C \ (G (F g) \\<^sub>C G (F g'))) \\<^sub>C G (\ (G (F g), G (F g')) \\<^sub>D (\ (F g) \\<^sub>D \ (F g'))) \\<^sub>C C.inv (G (\ (g, g')))" proof - have "C.iso (\ (G (F g) \\<^sub>C G (F g')) \\<^sub>C G (\ (G (F g), G (F g')) \\<^sub>D (\ (F g) \\<^sub>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 \\<^sub>C g') \\<^sub>C (G (F (\ g \\<^sub>C \ g')) \\<^sub>C G (\ (G (F g), G (F g')) \\<^sub>D (\ (F g) \\<^sub>D \ (F g')))) \\<^sub>C C.inv (G (\ (g, g')))" proof - have "(\ g \\<^sub>C \ g') \\<^sub>C \ (G (F g) \\<^sub>C G (F g')) = \ (g \\<^sub>C g') \\<^sub>C G (F (\ g \\<^sub>C \ g'))" using assms \_naturality [of "\ g \\<^sub>C \ g'"] by simp thus ?thesis using C.comp_assoc by simp qed also have "... = \ (g \\<^sub>C g') \\<^sub>C G ((F (\ g \\<^sub>C \ g') \\<^sub>D \ (G (F g), G (F g'))) \\<^sub>D (\ (F g) \\<^sub>D \ (F g'))) \\<^sub>C C.inv (G (\ (g, g')))" using assms C.VV.ide_char C.VV.arr_char FF_def D.comp_assoc by auto also have "... = \ (g \\<^sub>C g') \\<^sub>C G (\ (g, g') \\<^sub>D (F (\ g) \\<^sub>D F (\ g')) \\<^sub>D (\ (F g) \\<^sub>D \ (F g'))) \\<^sub>C C.inv (G (\ (g, g')))" proof - have "F (\ g \\<^sub>C \ g') \\<^sub>D \ (G (F g), G (F g')) = \ (g, g') \\<^sub>D (F (\ g) \\<^sub>D F (\ g'))" using assms C.VV.arr_char D.VV.ide_char D.VV.arr_char FF_def \.naturality [of "(\ g, \ g')"] C.VV.dom_char C.VV.cod_char by simp thus ?thesis using D.comp_assoc by simp qed also have "... = \ (g \\<^sub>C g') \\<^sub>C G (\ (g, g') \\<^sub>D (F (\ g) \\<^sub>D \ (F g) \\<^sub>D F (\ g') \\<^sub>D \ (F g'))) \\<^sub>C C.inv (G (\ (g, g')))" using assms D.interchange by simp also have "... = \ (g \\<^sub>C g') \\<^sub>C G (\ (g, g') \\<^sub>D (F g \\<^sub>D F g')) \\<^sub>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 \\<^sub>C g') \\<^sub>C G (\ (g, g')) \\<^sub>C C.inv (G (\ (g, g')))" using assms D.comp_arr_dom cmp_simps(1) cmp_simps(4) by auto also have "... = \ (g \\<^sub>C g')" using assms D.comp_arr_dom C.comp_arr_inv' C.VV.ide_char C.VV.arr_char 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 (\ \\<^sub>D \') = \\<^sub>G.map (D.cod \, D.cod \') \\<^sub>C (G \ \\<^sub>C G \') \\<^sub>C C.inv (\\<^sub>G.map (D.dom \, D.dom \'))" proof - have "G (\ \\<^sub>D \') \\<^sub>C \\<^sub>G.map (D.dom \, D.dom \') = \\<^sub>G.map (D.cod \, D.cod \') \\<^sub>C (G \ \\<^sub>C G \')" using assms \\<^sub>G.naturality D.VV.arr_char D.VV.cod_char G.FF_def D.VV.dom_char by auto moreover have "C.seq (\\<^sub>G.map (D.cod \, D.cod \')) (G \ \\<^sub>C G \')" proof show "\G \ \\<^sub>C G \' : G (D.dom \) \\<^sub>C G (D.dom \') \\<^sub>C G (D.cod \) \\<^sub>C G (D.cod \')\" using assms G_in_hom(2) C.hcomp_in_vhom by auto show "\\\<^sub>G.map (D.cod \, D.cod \') : G (D.cod \) \\<^sub>C G (D.cod \') \\<^sub>C G (D.cod \ \\<^sub>D D.cod \')\" using assms \\<^sub>G\<^sub>0_in_hom \\<^sub>G.map_simp_ide D.VV.ide_char D.VV.arr_char by auto qed moreover have "C.iso (\\<^sub>G.map (D.dom \, D.dom \'))" using assms \\<^sub>G.components_are_iso D.VV.ide_char D.VV.arr_char by auto ultimately show ?thesis using assms \\<^sub>G.components_are_iso C.comp_assoc C.invert_side_of_triangle(2) [of "\\<^sub>G.map (D.cod \, D.cod \') \\<^sub>C (G \ \\<^sub>C G \')" "G (\ \\<^sub>D \')" "\\<^sub>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 "src\<^sub>D f = trg\<^sub>D g" and "src\<^sub>D g = trg\<^sub>D h" shows "F (G \\<^sub>D[f, g, h] \\<^sub>C \\<^sub>G.map (f \\<^sub>D g, h) \\<^sub>C (\\<^sub>G.map (f, g) \\<^sub>C G h)) = (\ (f \\<^sub>D g \\<^sub>D h) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h))) \\<^sub>D \\<^sub>D[F (G f), F (G g), F (G h)] \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h))" proof - have "F (G \\<^sub>D[f, g, h] \\<^sub>C \\<^sub>G.map (f \\<^sub>D g, h) \\<^sub>C (\\<^sub>G.map (f, g) \\<^sub>C G h)) = F (G \\<^sub>D[f, g, h] \\<^sub>C (G (D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h))) \\<^sub>C G (D.inv (\ (G (f \\<^sub>D g), G h))) \\<^sub>C C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h)) \\<^sub>C (G (D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>C G (D.inv (\ (G f, G g))) \\<^sub>C C.inv (\ (G f \\<^sub>C G g)) \\<^sub>C G h))" using assms \\<^sub>G_map_simp_ide C.comp_assoc by simp also have "... = F (G \\<^sub>D[f, g, h] \\<^sub>C (G (D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h))) \\<^sub>C G (D.inv (\ (G (f \\<^sub>D g), G h))) \\<^sub>C C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h)) \\<^sub>C (G (D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>C G h) \\<^sub>C (G (D.inv (\ (G f, G g))) \\<^sub>C G h) \\<^sub>C (C.inv (\ (G f \\<^sub>C G g)) \\<^sub>C G h))" using assms C.whisker_right D.comp_assoc by simp also have "... = F (G \\<^sub>D[f, g, h]) \\<^sub>D F (G (D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h))) \\<^sub>D F (G (D.inv (\ (G (f \\<^sub>D g), G h)))) \\<^sub>D F (C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h))) \\<^sub>D F (G (D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>C G h) \\<^sub>D F (G (D.inv (\ (G f, G g))) \\<^sub>C G h) \\<^sub>D F (C.inv (\ (G f \\<^sub>C G g)) \\<^sub>C G h)" proof - have "C.arr (G \\<^sub>D[f, g, h] \\<^sub>C (G (D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h))) \\<^sub>C G (D.inv (\ (G (f \\<^sub>D g), G h))) \\<^sub>C C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h)) \\<^sub>C (G (D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>C G h) \\<^sub>C (G (D.inv (\ (G f, G g))) \\<^sub>C G h) \\<^sub>C (C.inv (\ (G f \\<^sub>C G g)) \\<^sub>C G h))" using assms \\<^sub>G_simps G.FF_def G\<^sub>0_props by auto thus ?thesis by (metis C.seqE as_nat_trans.preserves_comp_2) qed also have "... = F (G \\<^sub>D[f, g, h]) \\<^sub>D F (G (D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h))) \\<^sub>D F (G (D.inv (\ (G (f \\<^sub>D g), G h)))) \\<^sub>D F (C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h))) \\<^sub>D (\ (G (f \\<^sub>D g), G h) \\<^sub>D (F (G (D.inv (\ f) \\<^sub>D D.inv (\ g))) \\<^sub>D F (G h)) \\<^sub>D ((D.inv (\ (G (F (G f) \\<^sub>D F (G g)), G h))) \\<^sub>D (\ (G (F (G f) \\<^sub>D F (G g)), G h)) \\<^sub>D (F (G (D.inv (\ (G f, G g)))) \\<^sub>D F (G h))) \\<^sub>D ((D.inv (\ (G (F (G f \\<^sub>C G g)), G h))) \\<^sub>D (\ (G (F (G f \\<^sub>C G g)), G h)) \\<^sub>D (F (C.inv (\ (G f \\<^sub>C G g))) \\<^sub>D F (G h))) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h)))" using assms G\<^sub>0_props preserves_hcomp FF_def D.comp_assoc by auto also have "... = F (G \\<^sub>D[f, g, h]) \\<^sub>D F (G (D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h))) \\<^sub>D F (G (D.inv (\ (G (f \\<^sub>D g), G h)))) \\<^sub>D F (C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h))) \\<^sub>D (\ (G (f \\<^sub>D g), G h) \\<^sub>D (F (G (D.inv (\ f) \\<^sub>D D.inv (\ g))) \\<^sub>D F (G h)) \\<^sub>D (F (G (D.inv (\ (G f, G g)))) \\<^sub>D F (G h)) \\<^sub>D (F (C.inv (\ (G f \\<^sub>C G g))) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h)))" proof - have "((D.inv (\ (G (F (G f) \\<^sub>D F (G g)), G h))) \\<^sub>D (\ (G (F (G f) \\<^sub>D F (G g)), G h))) \\<^sub>D (F (G (D.inv (\ (G f, G g)))) \\<^sub>D F (G h)) = F (G (D.inv (\ (G f, G g)))) \\<^sub>D F (G h)" using assms \.components_are_iso C.VV.ide_char C.VV.arr_char G\<^sub>0_props FF_def D.comp_inv_arr' D.comp_cod_arr by simp moreover have "(D.inv (\ (G (F (G f \\<^sub>C G g)), G h)) \\<^sub>D \ (G (F (G f \\<^sub>C G g)), G h)) \\<^sub>D (F (C.inv (\ (G f \\<^sub>C G g))) \\<^sub>D F (G h)) = F (C.inv (\ (G f \\<^sub>C G g))) \\<^sub>D F (G h)" using assms D.comp_cod_arr \.components_are_iso C.VV.ide_char C.VV.arr_char G\<^sub>0_props FF_def D.comp_inv_arr' by simp ultimately show ?thesis using D.comp_assoc by simp qed also have "... = \ (f \\<^sub>D g \\<^sub>D h) \\<^sub>D \\<^sub>D[f, g, h] \\<^sub>D ((D.inv (\ ((f \\<^sub>D g) \\<^sub>D h))) \\<^sub>D (\ ((f \\<^sub>D g) \\<^sub>D h)) \\<^sub>D (D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h))) \\<^sub>D (((D.inv (\ (F (G (f \\<^sub>D g)) \\<^sub>D F (G h)))) \\<^sub>D (\ (F (G (f \\<^sub>D g)) \\<^sub>D F (G h)))) \\<^sub>D D.inv (\ (G (f \\<^sub>D g), G h))) \\<^sub>D D.inv (\ (F (G (f \\<^sub>D g) \\<^sub>C G h))) \\<^sub>D F (C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h))) \\<^sub>D (\ (G (f \\<^sub>D g), G h) \\<^sub>D ((\ (f \\<^sub>D g) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D D.inv (\ (F (G f) \\<^sub>D F (G g)))) \\<^sub>D F (G h)) \\<^sub>D ((\ (F (G f) \\<^sub>D F (G g)) \\<^sub>D D.inv (\ (G f, G g)) \\<^sub>D D.inv (\ (F (G f \\<^sub>C G g)))) \\<^sub>D F (G h)) \\<^sub>D (F (C.inv (\ (G f \\<^sub>C G g))) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h)))" proof - have "\\. D.arr \ \ F (G \) = \ (D.cod \) \\<^sub>D \ \\<^sub>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 \\<^sub>D g \\<^sub>D h) \\<^sub>D \\<^sub>D[f, g, h] \\<^sub>D (D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h)) \\<^sub>D D.inv (\ (G (f \\<^sub>D g), G h)) \\<^sub>D D.inv (\ (F (G (f \\<^sub>D g) \\<^sub>C G h))) \\<^sub>D F (C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h))) \\<^sub>D (\ (G (f \\<^sub>D g), G h) \\<^sub>D ((\ (f \\<^sub>D g) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D D.inv (\ (F (G f) \\<^sub>D F (G g))) \\<^sub>D F (G h)) \\<^sub>D (\ (F (G f) \\<^sub>D F (G g)) \\<^sub>D D.inv (\ (G f, G g)) \\<^sub>D D.inv (\ (F (G f \\<^sub>C G g))) \\<^sub>D F (G h))) \\<^sub>D (F (C.inv (\ (G f \\<^sub>C G g))) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h)))" proof - have "(D.inv (\ ((f \\<^sub>D g) \\<^sub>D h)) \\<^sub>D \ ((f \\<^sub>D g) \\<^sub>D h)) \\<^sub>D (D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h)) = D.inv (\ (f \\<^sub>D g)) \\<^sub>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 \\<^sub>D g)) \\<^sub>D F (G h)))) \\<^sub>D \ (F (G (f \\<^sub>D g)) \\<^sub>D F (G h))) \\<^sub>D D.inv (\ (G (f \\<^sub>D g), G h)) = D.inv (\ (G (f \\<^sub>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 \\<^sub>D g \\<^sub>D h) \\<^sub>D \\<^sub>D[f, g, h] \\<^sub>D ((D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h))) \\<^sub>D (D.inv (\ (G (f \\<^sub>D g), G h)) \\<^sub>D D.inv (\ (F (G (f \\<^sub>D g) \\<^sub>C G h)))) \\<^sub>D F (C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h))) \\<^sub>D (\ (G (f \\<^sub>D g), G h) \\<^sub>D (\ (f \\<^sub>D g) \\<^sub>D F (G h)) \\<^sub>D ((D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D F (G h)) \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h)) \\<^sub>D (((D.inv (\ (F (G f \\<^sub>C G g))) \\<^sub>D F (G h)) \\<^sub>D (F (C.inv (\ (G f \\<^sub>C G g))) \\<^sub>D F (G h))) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h))))" proof - have "(\ (f \\<^sub>D g) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D D.inv (\ (F (G f) \\<^sub>D F (G g))) \\<^sub>D F (G h)) \\<^sub>D (\ (F (G f) \\<^sub>D F (G g)) \\<^sub>D D.inv (\ (G f, G g)) \\<^sub>D D.inv (\ (F (G f \\<^sub>C G g))) \\<^sub>D F (G h)) = (\ (f \\<^sub>D g) \\<^sub>D F (G h)) \\<^sub>D ((D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D F (G h)) \\<^sub>D (((D.inv (\ (F (G f) \\<^sub>D F (G g))) \\<^sub>D F (G h)) \\<^sub>D (\ (F (G f) \\<^sub>D F (G g)) \\<^sub>D F (G h))) \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h))) \\<^sub>D (D.inv (\ (F (G f \\<^sub>C G g))) \\<^sub>D F (G h))" using assms D.comp_assoc D.whisker_right by simp also have "... = (\ (f \\<^sub>D g) \\<^sub>D F (G h)) \\<^sub>D ((D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D F (G h)) \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h)) \\<^sub>D (D.inv (\ (F (G f \\<^sub>C G g))) \\<^sub>D F (G h))" proof - have "((D.inv (\ (F (G f) \\<^sub>D F (G g))) \\<^sub>D F (G h)) \\<^sub>D (\ (F (G f) \\<^sub>D F (G g)) \\<^sub>D F (G h))) \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h)) = D.inv (\ (G f, G g)) \\<^sub>D F (G h)" using assms D.comp_inv_arr' D.comp_cod_arr D.interchange [of "D.inv (\ (F (G f) \\<^sub>D F (G g)))" "\ (F (G f) \\<^sub>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 \\<^sub>D g \\<^sub>D h) \\<^sub>D \\<^sub>D[f, g, h] \\<^sub>D ((D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h))) \\<^sub>D (D.inv (\ (G (f \\<^sub>D g), G h)) \\<^sub>D ((D.inv (\ (F (G (f \\<^sub>D g) \\<^sub>C G h)))) \\<^sub>D F (C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h)))) \\<^sub>D \ (G (f \\<^sub>D g), G h)) \\<^sub>D (\ (f \\<^sub>D g) \\<^sub>D F (G h)) \\<^sub>D ((D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D F (G h)) \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h))" proof - have "((D.inv (\ (F (G f \\<^sub>C G g))) \\<^sub>D F (G h)) \\<^sub>D (F (C.inv (\ (G f \\<^sub>C G g))) \\<^sub>D F (G h))) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h)) = D.inv (\ (G f \\<^sub>C G g, G h))" proof - have "((D.inv (\ (F (G f \\<^sub>C G g))) \\<^sub>D F (G h)) \\<^sub>D (F (C.inv (\ (G f \\<^sub>C G g))) \\<^sub>D F (G h))) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h)) = ((D.inv (\ (F (G f \\<^sub>C G g))) \\<^sub>D F (G h)) \\<^sub>D (D.inv (F (\ (G f \\<^sub>C G g))) \\<^sub>D F (G h))) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h))" using assms by (simp add: preserves_inv) also have "... = (D.inv (\ (F (G f \\<^sub>C G g))) \\<^sub>D D.inv (F (\ (G f \\<^sub>C G g))) \\<^sub>D F (G h) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h))" using assms \.components_are_iso D.interchange [of "D.inv (\ (F (G f \\<^sub>C G g)))" "D.inv (F (\ (G f \\<^sub>C G g)))" "F (G h)" "F (G h)"] by simp also have "... = (D.inv (F (\ (G f \\<^sub>C G g)) \\<^sub>D \ (F (G f \\<^sub>C G g))) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h))" proof - have "D.iso (F (\ (G f \\<^sub>C G g)) \\<^sub>D \ (F (G f \\<^sub>C G g)))" using assms by auto hence "D.inv (\ (F (G f \\<^sub>C G g))) \\<^sub>D D.inv (F (\ (G f \\<^sub>C G g))) = D.inv (F (\ (G f \\<^sub>C G g)) \\<^sub>D \ (F (G f \\<^sub>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 \\<^sub>C G g) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h))" using assms \\.triangle_G Fo\_\oF.map_def [of "G f \\<^sub>C G g"] by simp also have "... = D.inv (\ (G f \\<^sub>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 \\<^sub>D g \\<^sub>D h) \\<^sub>D \\<^sub>D[f, g, h] \\<^sub>D ((D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h))) \\<^sub>D ((D.inv (\ (G (f \\<^sub>D g), G h)) \\<^sub>D \ (G (f \\<^sub>D g), G h)) \\<^sub>D (\ (f \\<^sub>D g) \\<^sub>D F (G h))) \\<^sub>D ((D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D F (G h)) \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h))" proof - have "((D.inv (\ (F (G (f \\<^sub>D g) \\<^sub>C G h))) \\<^sub>D F (C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h))))) \\<^sub>D (\ (G (f \\<^sub>D g), G h)) = \ (G (f \\<^sub>D g), G h)" proof - have "D.inv (\ (F (G (f \\<^sub>D g) \\<^sub>C G h))) \\<^sub>D F (C.inv (\ (G (f \\<^sub>D g) \\<^sub>C G h))) = D.inv (\ (F (G (f \\<^sub>D g) \\<^sub>C G h))) \\<^sub>D D.inv (F (\ (G (f \\<^sub>D g) \\<^sub>C G h)))" using assms preserves_inv by simp also have "... = D.inv (F (\ (G (f \\<^sub>D g) \\<^sub>C G h)) \\<^sub>D \ (F (G (f \\<^sub>D g) \\<^sub>C G h)))" proof - have "D.iso (F (\ (G (f \\<^sub>D g) \\<^sub>C G h)) \\<^sub>D \ (F (G (f \\<^sub>D g) \\<^sub>C G h)))" using assms by auto thus ?thesis using assms D.inv_comp_right(1) by simp qed also have "... = F (G (f \\<^sub>D g) \\<^sub>C G h)" using assms \\.triangle_G Fo\_\oF.map_def [of "G (f \\<^sub>D g) \\<^sub>C G h"] by simp finally show ?thesis using assms D.comp_cod_arr [of "\ (G (f \\<^sub>D g), G h)" "F (G (f \\<^sub>D g) \\<^sub>C G h)"] by fastforce qed thus ?thesis using D.comp_assoc by simp qed also have "... = \ (f \\<^sub>D g \\<^sub>D h) \\<^sub>D \\<^sub>D[f, g, h] \\<^sub>D (((D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h)) \\<^sub>D (\ (f \\<^sub>D g) \\<^sub>D F (G h))) \\<^sub>D ((D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D F (G h))) \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h))" proof - have "((D.inv (\ (G (f \\<^sub>D g), G h)) \\<^sub>D \ (G (f \\<^sub>D g), G h)) \\<^sub>D (\ (f \\<^sub>D g) \\<^sub>D F (G h))) = \ (f \\<^sub>D g) \\<^sub>D F (G h)" using assms D.comp_inv_arr' \.components_are_iso C.VV.ide_char C.VV.arr_char G\<^sub>0_props FF_def D.comp_cod_arr by simp thus ?thesis using D.comp_assoc by simp qed also have "... = \ (f \\<^sub>D g \\<^sub>D h) \\<^sub>D (\\<^sub>D[f, g, h] \\<^sub>D ((D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D D.inv (\ h))) \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h))" proof - have "((D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h)) \\<^sub>D (\ (f \\<^sub>D g) \\<^sub>D F (G h))) \\<^sub>D ((D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D F (G h)) = (D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D D.inv (\ h)" proof - have "(D.inv (\ (f \\<^sub>D g)) \\<^sub>D D.inv (\ h)) \\<^sub>D (\ (f \\<^sub>D g) \\<^sub>D F (G h)) = (f \\<^sub>D g) \\<^sub>D D.inv (\ h)" using assms D.comp_inv_arr' D.comp_arr_dom D.interchange [of "D.inv (\ (f \\<^sub>D g))" "\ (f \\<^sub>D g)" "D.inv (\ h)" "F (G h)"] D.invert_side_of_triangle(2) by simp moreover have "((f \\<^sub>D g) \\<^sub>D D.inv (\ h)) \\<^sub>D ((D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D F (G h)) = (D.inv (\ f) \\<^sub>D D.inv (\ g)) \\<^sub>D D.inv (\ h)" using assms D.comp_cod_arr D.comp_arr_dom D.interchange [of "f \\<^sub>D g" "D.inv (\ f) \\<^sub>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 \\<^sub>D g \\<^sub>D h) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h)) \\<^sub>D \\<^sub>D[F (G f), F (G g), F (G h)]) \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>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 "src\<^sub>D f = trg\<^sub>D g" and "src\<^sub>D g = trg\<^sub>D h" shows "F (\\<^sub>G.map (f, g \\<^sub>D h) \\<^sub>C (G f \\<^sub>C \\<^sub>G.map (g, h))) = (\ (f \\<^sub>D g \\<^sub>D h) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h))) \\<^sub>D (F (G f) \\<^sub>D D.inv (\ (G g, G h))) \\<^sub>D D.inv (\ (G f, G g \\<^sub>C G h))" proof - have "F (\\<^sub>G.map (f, g \\<^sub>D h) \\<^sub>C (G f \\<^sub>C \\<^sub>G.map (g, h))) = F (G (D.inv (\ f) \\<^sub>D D.inv (\ (g \\<^sub>D h))) \\<^sub>C G (D.inv (\ (G f, G (g \\<^sub>D h)))) \\<^sub>C C.inv (\ (G f \\<^sub>C G (g \\<^sub>D h))) \\<^sub>C (G f \\<^sub>C G (D.inv (\ g) \\<^sub>D D.inv (\ h)) \\<^sub>C G (D.inv (\ (G g, G h))) \\<^sub>C C.inv (\ (G g \\<^sub>C G h))))" using assms \\<^sub>G_map_simp_ide C.comp_assoc by simp also have "... = F (G (D.inv (\ f) \\<^sub>D D.inv (\ (g \\<^sub>D h))) \\<^sub>C G (D.inv (\ (G f, G (g \\<^sub>D h)))) \\<^sub>C C.inv (\ (G f \\<^sub>C G (g \\<^sub>D h))) \\<^sub>C (G f \\<^sub>C G (D.inv (\ g) \\<^sub>D D.inv (\ h))) \\<^sub>C (G f \\<^sub>C G (D.inv (\ (G g, G h)))) \\<^sub>C (G f \\<^sub>C C.inv (\ (G g \\<^sub>C G h))))" using assms C.whisker_left by simp also have "... = F (G (D.inv (\ f) \\<^sub>D D.inv (\ (g \\<^sub>D h)))) \\<^sub>D F (G (D.inv (\ (G f, G (g \\<^sub>D h))))) \\<^sub>D F (C.inv (\ (G f \\<^sub>C G (g \\<^sub>D h)))) \\<^sub>D F (G f \\<^sub>C G (D.inv (\ g) \\<^sub>D D.inv (\ h))) \\<^sub>D F (G f \\<^sub>C G (D.inv (\ (G g, G h)))) \\<^sub>D F (G f \\<^sub>C C.inv (\ (G g \\<^sub>C G h)))" proof - have "C.arr (G (D.inv (\ f) \\<^sub>D D.inv (\ (g \\<^sub>D h))) \\<^sub>C G (D.inv (\ (G f, G (g \\<^sub>D h)))) \\<^sub>C C.inv (\ (G f \\<^sub>C G (g \\<^sub>D h))) \\<^sub>C (G f \\<^sub>C G (D.inv (\ g) \\<^sub>D D.inv (\ h))) \\<^sub>C (G f \\<^sub>C G (D.inv (\ (G g, G h)))) \\<^sub>C (G f \\<^sub>C C.inv (\ (G g \\<^sub>C G h))))" using assms G\<^sub>0_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) \\<^sub>D D.inv (\ (g \\<^sub>D h)))) \\<^sub>D F (G (D.inv (\ (G f, G (g \\<^sub>D h))))) \\<^sub>D F (C.inv (\ (G f \\<^sub>C G (g \\<^sub>D h)))) \\<^sub>D (\ (G f, G (g \\<^sub>D h)) \\<^sub>D (F (G f) \\<^sub>D F (G (D.inv (\ g) \\<^sub>D D.inv (\ h)))) \\<^sub>D ((D.inv (\ (G f, G (F (G g) \\<^sub>D F (G h))))) \\<^sub>D (\ (G f, G (F (G g) \\<^sub>D F (G h)))) \\<^sub>D (F (G f) \\<^sub>D F (G (D.inv (\ (G g, G h)))))) \\<^sub>D ((D.inv (\ (G f, G (F (G g \\<^sub>C G h))))) \\<^sub>D (\ (G f, G (F (G g \\<^sub>C G h)))) \\<^sub>D (F (G f) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h))))) \\<^sub>D D.inv (\ (G f, G g \\<^sub>C G h)))" using assms preserves_hcomp G\<^sub>0_props D.comp_assoc by simp also have "... = F (G (D.inv (\ f) \\<^sub>D D.inv (\ (g \\<^sub>D h)))) \\<^sub>D F (G (D.inv (\ (G f, G (g \\<^sub>D h))))) \\<^sub>D F (C.inv (\ (G f \\<^sub>C G (g \\<^sub>D h)))) \\<^sub>D (\ (G f, G (g \\<^sub>D h)) \\<^sub>D (F (G f) \\<^sub>D F (G (D.inv (\ g) \\<^sub>D D.inv (\ h)))) \\<^sub>D (F (G f) \\<^sub>D F (G (D.inv (\ (G g, G h))))) \\<^sub>D (F (G f) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h)))) \\<^sub>D D.inv (\ (G f, G g \\<^sub>C G h)))" proof - have "D.inv (\ (G f, G (F (G g) \\<^sub>D F (G h)))) \\<^sub>D \ (G f, G (F (G g) \\<^sub>D F (G h))) \\<^sub>D (F (G f) \\<^sub>D F (G (D.inv (\ (G g, G h))))) = (F (G f) \\<^sub>D F (G (D.inv (\ (G g, G h)))))" proof - have "D.inv (\ (G f, G (F (G g) \\<^sub>D F (G h)))) \\<^sub>D \ (G f, G (F (G g) \\<^sub>D F (G h))) = F (G f) \\<^sub>D F (G (F (G g) \\<^sub>D F (G h)))" using assms D.comp_inv_arr D.inv_is_inverse G\<^sub>0_props FF_def by simp moreover have "... = D.cod (F (G f) \\<^sub>D F (G (D.inv (\ (G g, G h)))))" using assms G\<^sub>0_props by simp moreover have "D.hseq (F (G f)) (F (G (D.inv (\ (G g, G h)))))" using assms G\<^sub>0_props by auto ultimately show ?thesis using assms D.comp_assoc D.comp_cod_arr [of "F (G f) \\<^sub>D F (G (D.inv (\ (G g, G h))))" "F (G f) \\<^sub>D F (G (F (G g) \\<^sub>D F (G h)))"] by metis qed moreover have "D.inv (\ (G f, G (F (G g \\<^sub>C G h)))) \\<^sub>D (\ (G f, G (F (G g \\<^sub>C G h)))) \\<^sub>D (F (G f) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h)))) = (F (G f) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h))))" proof - have "D.inv (\ (G f, G (F (G g \\<^sub>C G h)))) \\<^sub>D \ (G f, G (F (G g \\<^sub>C G h))) = F (G f) \\<^sub>D F (G (F (G g \\<^sub>C G h)))" using assms D.comp_inv_arr D.inv_is_inverse G\<^sub>0_props by simp moreover have "... = D.cod (F (G f) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h))))" using assms by simp moreover have "D.hseq (F (G f)) (F (C.inv (\ (G g \\<^sub>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) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h)))" "F (G f) \\<^sub>D F (G (F (G g \\<^sub>C G h)))"] by metis qed ultimately show ?thesis by simp qed also have "... = (\ (f \\<^sub>D g \\<^sub>D h) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ (g \\<^sub>D h))) \\<^sub>D ((D.inv (\ (F (G f) \\<^sub>D F (G (g \\<^sub>D h))))) \\<^sub>D \ (F (G f) \\<^sub>D F (G (g \\<^sub>D h)))) \\<^sub>D D.inv (\ (G f, G (g \\<^sub>D h)))) \\<^sub>D ((D.inv (\ (F (G f \\<^sub>C G (g \\<^sub>D h)))) \\<^sub>D F (C.inv (\ (G f \\<^sub>C G (g \\<^sub>D h))))) \\<^sub>D \ (G f, G (g \\<^sub>D h))) \\<^sub>D (F (G f) \\<^sub>D \ (g \\<^sub>D h) \\<^sub>D (D.inv (\ g) \\<^sub>D D.inv (\ h)) \\<^sub>D D.inv (\ (F (G g) \\<^sub>D F (G h)))) \\<^sub>D (F (G f) \\<^sub>D \ (F (G g) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G g, G h)) \\<^sub>D D.inv (\ (F (G g \\<^sub>C G h)))) \\<^sub>D (F (G f) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h)))) \\<^sub>D D.inv (\ (G f, G g \\<^sub>C G h))" proof - have "\\. D.arr \ \ F (G \) = \ (D.cod \) \\<^sub>D \ \\<^sub>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 \\<^sub>D g \\<^sub>D h) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ (g \\<^sub>D h))) \\<^sub>D D.inv (\ (G f, G (g \\<^sub>D h)))) \\<^sub>D \ (G f, G (g \\<^sub>D h)) \\<^sub>D (F (G f) \\<^sub>D \ (g \\<^sub>D h) \\<^sub>D (D.inv (\ g) \\<^sub>D D.inv (\ h)) \\<^sub>D D.inv (\ (F (G g) \\<^sub>D F (G h)))) \\<^sub>D (F (G f) \\<^sub>D \ (F (G g) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G g, G h)) \\<^sub>D D.inv (\ (F (G g \\<^sub>C G h)))) \\<^sub>D (F (G f) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h)))) \\<^sub>D D.inv (\ (G f, G g \\<^sub>C G h))" proof - have "D.inv (\ (F (G f) \\<^sub>D F (G (g \\<^sub>D h)))) \\<^sub>D \ (F (G f) \\<^sub>D F (G (g \\<^sub>D h))) \\<^sub>D D.inv (\ (G f, G (g \\<^sub>D h))) = D.inv (\ (G f, G (g \\<^sub>D h)))" proof - have "D.inv (\ (F (G f) \\<^sub>D F (G (g \\<^sub>D h)))) \\<^sub>D \ (F (G f) \\<^sub>D F (G (g \\<^sub>D h))) = F (G f) \\<^sub>D F (G (g \\<^sub>D h))" using assms D.comp_inv_arr D.inv_is_inverse by simp moreover have "... = D.cod (D.inv (\ (G f, G (g \\<^sub>D h))))" using assms by simp ultimately show ?thesis using assms D.comp_cod_arr [of "D.inv (\ (G f, G (g \\<^sub>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 \\<^sub>C G (g \\<^sub>D h)))) \\<^sub>D F (C.inv (\ (G f \\<^sub>C G (g \\<^sub>D h))))) \\<^sub>D \ (G f, G (g \\<^sub>D h)) = \ (G f, G (g \\<^sub>D h))" proof - have "D.inv (\ (F (G f \\<^sub>C G (g \\<^sub>D h)))) \\<^sub>D F (C.inv (\ (G f \\<^sub>C G (g \\<^sub>D h)))) = F (G f \\<^sub>C G (g \\<^sub>D h))" proof - have "D.inv (\ (F (G f \\<^sub>C G (g \\<^sub>D h)))) \\<^sub>D F (C.inv (\ (G f \\<^sub>C G (g \\<^sub>D h)))) = D.inv (\ (F (G f \\<^sub>C G (g \\<^sub>D h)))) \\<^sub>D D.inv (F (\ (G f \\<^sub>C G (g \\<^sub>D h))))" using assms by (simp add: preserves_inv) also have "... = D.inv (F (\ (G f \\<^sub>C G (g \\<^sub>D h))) \\<^sub>D \ (F (G f \\<^sub>C G (g \\<^sub>D h))))" proof - have "D.iso (F (\ (G f \\<^sub>C G (g \\<^sub>D h))) \\<^sub>D \ (F (G f \\<^sub>C G (g \\<^sub>D h))))" using assms by auto thus ?thesis using assms by (simp add: D.inv_comp_right(1)) qed also have "... = F (G f \\<^sub>C G (g \\<^sub>D h))" using assms Fo\_\oF.map_def [of "G f \\<^sub>C G (g \\<^sub>D h)"] \\.triangle_G by simp finally show ?thesis by blast qed moreover have "... = D.cod (\ (G f, G (g \\<^sub>D h)))" using assms by simp moreover have "D.arr (\ (G f, G (g \\<^sub>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 \\<^sub>D g \\<^sub>D h) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ (g \\<^sub>D h))) \\<^sub>D ((D.inv (\ (G f, G (g \\<^sub>D h)))) \\<^sub>D \ (G f, G (g \\<^sub>D h))) \\<^sub>D (F (G f) \\<^sub>D \ (g \\<^sub>D h))) \\<^sub>D (F (G f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h)) \\<^sub>D (((F (G f) \\<^sub>D D.inv (\ (F (G g) \\<^sub>D F (G h)))) \\<^sub>D (F (G f) \\<^sub>D \ (F (G g) \\<^sub>D F (G h)))) \\<^sub>D (F (G f) \\<^sub>D D.inv (\ (G g, G h)))) \\<^sub>D ((F (G f) \\<^sub>D D.inv (\ (F (G g \\<^sub>C G h)))) \\<^sub>D (F (G f) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h))))) \\<^sub>D D.inv (\ (G f, G g \\<^sub>C G h))" using assms D.whisker_left D.comp_assoc by simp also have "... = (\ (f \\<^sub>D g \\<^sub>D h) \\<^sub>D (((D.inv (\ f) \\<^sub>D D.inv (\ (g \\<^sub>D h))) \\<^sub>D (F (G f) \\<^sub>D \ (g \\<^sub>D h)))) \\<^sub>D (F (G f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h))) \\<^sub>D (F (G f) \\<^sub>D D.inv (\ (G g, G h))) \\<^sub>D D.inv (\ (G f, G g \\<^sub>C G h))" proof - have "(D.inv (\ (G f, G (g \\<^sub>D h))) \\<^sub>D \ (G f, G (g \\<^sub>D h))) \\<^sub>D (F (G f) \\<^sub>D \ (g \\<^sub>D h)) = (F (G f) \\<^sub>D \ (g \\<^sub>D h))" using assms D.comp_inv_arr' G\<^sub>0_props D.comp_cod_arr by simp moreover have "((F (G f) \\<^sub>D D.inv (\ (F (G g) \\<^sub>D F (G h)))) \\<^sub>D (F (G f) \\<^sub>D \ (F (G g) \\<^sub>D F (G h)))) \\<^sub>D (F (G f) \\<^sub>D D.inv (\ (G g, G h))) = F (G f) \\<^sub>D D.inv (\ (G g, G h))" proof - have "(F (G f) \\<^sub>D D.inv (\ (F (G g) \\<^sub>D F (G h)))) \\<^sub>D (F (G f) \\<^sub>D \ (F (G g) \\<^sub>D F (G h))) = F (G f) \\<^sub>D F (G g) \\<^sub>D F (G h)" proof - have "(F (G f) \\<^sub>D D.inv (\ (F (G g) \\<^sub>D F (G h)))) \\<^sub>D (F (G f) \\<^sub>D \ (F (G g) \\<^sub>D F (G h))) = F (G f) \\<^sub>D D.inv (\ (F (G g) \\<^sub>D F (G h))) \\<^sub>D \ (F (G g) \\<^sub>D F (G h))" using assms D.interchange [of "F (G f)" "F (G f)"] by simp also have "... = F (G f) \\<^sub>D F (G g) \\<^sub>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) \\<^sub>D D.inv (\ (G g, G h)))" using assms by simp moreover have "D.arr (F (G f) \\<^sub>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) \\<^sub>D D.inv (\ (F (G g \\<^sub>C G h)))) \\<^sub>D (F (G f) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h))))) \\<^sub>D D.inv (\ (G f, G g \\<^sub>C G h)) = D.inv (\ (G f, G g \\<^sub>C G h))" proof - have "(F (G f) \\<^sub>D D.inv (\ (F (G g \\<^sub>C G h)))) \\<^sub>D (F (G f) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h)))) = F (G f) \\<^sub>D F (G g \\<^sub>C G h)" proof - have "(F (G f) \\<^sub>D D.inv (\ (F (G g \\<^sub>C G h)))) \\<^sub>D (F (G f) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h)))) = F (G f) \\<^sub>D D.inv (\ (F (G g \\<^sub>C G h))) \\<^sub>D F (C.inv (\ (G g \\<^sub>C G h)))" using assms D.interchange [of "F (G f)" "F (G f)"] by simp also have "... = F (G f) \\<^sub>D D.inv (\ (F (G g \\<^sub>C G h))) \\<^sub>D D.inv (F (\ (G g \\<^sub>C G h)))" using assms preserves_inv by simp also have "... = F (G f) \\<^sub>D D.inv (F (\ (G g \\<^sub>C G h)) \\<^sub>D \ (F (G g \\<^sub>C G h)))" proof - have "D.iso (F (\ (G g \\<^sub>C G h)) \\<^sub>D \ (F (G g \\<^sub>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) \\<^sub>D F (G g \\<^sub>C G h)" using assms \\.triangle_G Fo\_\oF.map_def [of "G g \\<^sub>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 \\<^sub>D g \\<^sub>D h) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h))) \\<^sub>D (F (G f) \\<^sub>D D.inv (\ (G g, G h))) \\<^sub>D D.inv (\ (G f, G g \\<^sub>C G h))" proof - have "((D.inv (\ f) \\<^sub>D D.inv (\ (g \\<^sub>D h))) \\<^sub>D (F (G f) \\<^sub>D \ (g \\<^sub>D h))) \\<^sub>D (F (G f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h)) = D.inv (\ f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h)" using assms D.comp_cod_arr D.comp_arr_dom D.interchange proof - have "((D.inv (\ f) \\<^sub>D D.inv (\ (g \\<^sub>D h))) \\<^sub>D (F (G f) \\<^sub>D \ (g \\<^sub>D h))) \\<^sub>D (F (G f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h)) = (D.inv (\ f) \\<^sub>D F (G f) \\<^sub>D D.inv (\ (g \\<^sub>D h)) \\<^sub>D \ (g \\<^sub>D h)) \\<^sub>D (F (G f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h))" using assms D.interchange by simp also have "... = (D.inv (\ f) \\<^sub>D g \\<^sub>D h) \\<^sub>D (F (G f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h))" using assms D.comp_inv_arr' D.comp_arr_dom by simp also have "... = D.inv (\ f) \\<^sub>D F (G f) \\<^sub>D (g \\<^sub>D h) \\<^sub>D (D.inv (\ g) \\<^sub>D D.inv (\ h))" using assms D.interchange by simp also have "... = D.inv (\ f) \\<^sub>D D.inv (\ g) \\<^sub>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 V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C G \\<^sub>G.map proof show "\f g h. \D.ide f; D.ide g; D.ide h; src\<^sub>D f = trg\<^sub>D g; src\<^sub>D g = trg\<^sub>D h\ \ G \\<^sub>D[f, g, h] \\<^sub>C \\<^sub>G.map (f \\<^sub>D g, h) \\<^sub>C (\\<^sub>G.map (f, g) \\<^sub>C G h) = \\<^sub>G.map (f, g \\<^sub>D h) \\<^sub>C (G f \\<^sub>C \\<^sub>G.map (g, h)) \\<^sub>C \\<^sub>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: "src\<^sub>D f = trg\<^sub>D g" and gh: "src\<^sub>D g = trg\<^sub>D h" have "F (G \\<^sub>D[f, g, h] \\<^sub>C \\<^sub>G.map (f \\<^sub>D g, h) \\<^sub>C (\\<^sub>G.map (f, g) \\<^sub>C G h)) = F (\\<^sub>G.map (f, g \\<^sub>D h) \\<^sub>C (G f \\<^sub>C \\<^sub>G.map (g, h)) \\<^sub>C \\<^sub>C[G f, G g, G h])" proof - have "F (G \\<^sub>D[f, g, h] \\<^sub>C \\<^sub>G.map (f \\<^sub>D g, h) \\<^sub>C (\\<^sub>G.map (f, g) \\<^sub>C G h)) = (\ (f \\<^sub>D g \\<^sub>D h) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h))) \\<^sub>D \\<^sub>D[F (G f), F (G g), F (G h)] \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h))" using f g h fg gh coherence_LHS by simp also have "... = (\ (f \\<^sub>D g \\<^sub>D h) \\<^sub>D (D.inv (\ f) \\<^sub>D D.inv (\ g) \\<^sub>D D.inv (\ h))) \\<^sub>D ((F (G f) \\<^sub>D D.inv (\ (G g, G h))) \\<^sub>D D.inv (\ (G f, G g \\<^sub>C G h))) \\<^sub>D F \\<^sub>C[G f, G g, G h]" proof - have "\\<^sub>D[F (G f), F (G g), F (G h)] \\<^sub>D (D.inv (\ (G f, G g)) \\<^sub>D F (G h)) \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h)) = \\<^sub>D[F (G f), F (G g), F (G h)] \\<^sub>D D.inv (\ (G f \\<^sub>C G g, G h) \\<^sub>D (\ (G f, G g) \\<^sub>D F (G h)))" proof - have "D.iso (\ (G f \\<^sub>C G g, G h) \\<^sub>D (\ (G f, G g) \\<^sub>D F (G h)))" using f g h fg gh \.components_are_iso C.VV.ide_char C.VV.arr_char 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_char D.inv_comp_right(1)) qed also have "... = D.inv (\ (G f, G g \\<^sub>C G h) \\<^sub>D (F (G f) \\<^sub>D \ (G g, G h))) \\<^sub>D F \\<^sub>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_char C.VV.arr_char FF_def D.invert_opposite_sides_of_square by simp also have "... = ((F (G f) \\<^sub>D D.inv (\ (G g, G h))) \\<^sub>D D.inv (\ (G f, G g \\<^sub>C G h))) \\<^sub>D F \\<^sub>C[G f, G g, G h]" proof - have "D.iso (\ (G f, G g \\<^sub>C G h) \\<^sub>D (F (G f) \\<^sub>D \ (G g, G h)))" using f g h fg gh \.components_are_iso C.VV.ide_char C.VV.arr_char FF_def by auto thus ?thesis using f g h fg gh \.components_are_iso C.VV.ide_char C.VV.arr_char FF_def by (simp add: C.VV.arr_char D.comp_assoc D.inv_comp_right(1)) qed finally show ?thesis using D.comp_assoc by simp qed also have "... = F (\\<^sub>G.map (f, g \\<^sub>D h) \\<^sub>C (G f \\<^sub>C \\<^sub>G.map (g, h))) \\<^sub>D F \\<^sub>C[G f, G g, G h]" using f g h fg gh coherence_RHS D.comp_assoc by simp also have "... = F ((\\<^sub>G.map (f, g \\<^sub>D h) \\<^sub>C (G f \\<^sub>C \\<^sub>G.map (g, h))) \\<^sub>C \\<^sub>C[G f, G g, G h])" using f g h fg gh \\<^sub>G_simps by auto also have "... = F (\\<^sub>G.map (f, g \\<^sub>D h) \\<^sub>C (G f \\<^sub>C \\<^sub>G.map (g, h)) \\<^sub>C \\<^sub>C[G f, G g, G h])" using C.comp_assoc by simp finally show ?thesis by simp qed moreover have "C.par (G \\<^sub>D[f, g, h] \\<^sub>C \\<^sub>G.map (f \\<^sub>D g, h) \\<^sub>C (\\<^sub>G.map (f, g) \\<^sub>C G h)) (\\<^sub>G.map (f, g \\<^sub>D h) \\<^sub>C (G f \\<^sub>C \\<^sub>G.map (g, h)) \\<^sub>C \\<^sub>C[G f, G g, G h])" using f g h fg gh \\<^sub>G_simps by auto ultimately show "G \\<^sub>D[f, g, h] \\<^sub>C \\<^sub>G.map (f \\<^sub>D g, h) \\<^sub>C (\\<^sub>G.map (f, g) \\<^sub>C G h) = \\<^sub>G.map (f, g \\<^sub>D h) \\<^sub>C (G f \\<^sub>C \\<^sub>G.map (g, h)) \\<^sub>C \\<^sub>C[G f, G g, G h]" using is_faithful by blast qed qed interpretation GF: composite_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \ G \\<^sub>G.map .. interpretation FG: composite_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D G \\<^sub>G.map F \ .. interpretation I\<^sub>C: identity_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C .. interpretation I\<^sub>D: identity_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D .. 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.map\<^sub>0 a \\<^sub>D FG.map a\" using assms \_in_hom [of a] FG.map\<^sub>0_def G\<^sub>0_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 \\<^sub>D \\<^sub>D[FG.map\<^sub>0 a] = (FG.map \\<^sub>D[a] \\<^sub>D FG.cmp (a, a)) \\<^sub>D (\ a \\<^sub>D \ a)" proof - have "(FG.map \\<^sub>D[a] \\<^sub>D FG.cmp (a, a)) \\<^sub>D (\ a \\<^sub>D \ a) = F (G \\<^sub>D[a]) \\<^sub>D FG.cmp (a, a) \\<^sub>D (\ a \\<^sub>D \ a)" using assms D.comp_assoc by simp also have "... = F (G \\<^sub>D[a]) \\<^sub>D F (G (I\<^sub>D.cmp (a, a))) \\<^sub>D F (\\<^sub>G.map (a, a)) \\<^sub>D \ (G a, G a) \\<^sub>D (\ a \\<^sub>D \ a)" using assms FG.cmp_def D.comp_assoc D.VV.arr_char D.VV.dom_char D.VV.cod_char by auto also have "... = F (G \\<^sub>D[a]) \\<^sub>D F (G (I\<^sub>D.cmp (a, a))) \\<^sub>D \ (a \\<^sub>D a)" using assms \_hcomp by auto also have "... = F (G \\<^sub>D[a]) \\<^sub>D \ (a \\<^sub>D a)" using assms D.comp_cod_arr by auto also have "... = \ a \\<^sub>D \\<^sub>D[a]" using assms \_naturality [of "\\<^sub>D[a]"] by simp also have "... = \ a \\<^sub>D \\<^sub>D[FG.map\<^sub>0 a]" using assms \\\ a : FG.map\<^sub>0 a \\<^sub>D FG.map a\\ by fastforce finally show ?thesis by simp qed qed lemma \_hcomp': assumes "C.ide g" and "C.ide f" and "src\<^sub>C g = trg\<^sub>C f" shows "\ (g \\<^sub>C f) \\<^sub>C GF.cmp (g, f) = \ g \\<^sub>C \ f" proof - have "\ (g \\<^sub>C f) \\<^sub>C GF.cmp (g, f) = (\ g \\<^sub>C \ f) \\<^sub>C C.inv (\\<^sub>G.map (F g, F f)) \\<^sub>C C.inv (G (\ (g, f))) \\<^sub>C G (F (g \\<^sub>C f)) \\<^sub>C G (\ (g, f)) \\<^sub>C \\<^sub>G.map (F g, F f)" using assms \_hcomp GF.cmp_def C.VV.arr_char C.comp_cod_arr C.comp_inv_arr' C.comp_assoc C.VV.dom_char C.VV.cod_char by simp also have "... = (\ g \\<^sub>C \ f) \\<^sub>C C.inv (\\<^sub>G.map (F g, F f)) \\<^sub>C (C.inv (G (\ (g, f))) \\<^sub>C G (F (g \\<^sub>C f)) \\<^sub>C G (\ (g, f))) \\<^sub>C \\<^sub>G.map (F g, F f)" using C.comp_assoc by simp also have "... = (\ g \\<^sub>C \ f) \\<^sub>C C.inv (\\<^sub>G.map (F g, F f)) \\<^sub>C (C.inv (G (\ (g, f))) \\<^sub>C G (\ (g, f))) \\<^sub>C \\<^sub>G.map (F g, F f)" using assms C.comp_ide_arr [of "G (F (g \\<^sub>C f))" "G (\ (g, f))"] C.VV.arr_char C.VV.dom_char C.VV.cod_char by simp also have "... = (\ g \\<^sub>C \ f) \\<^sub>C C.inv (\\<^sub>G.map (F g, F f)) \\<^sub>C \\<^sub>G.map (F g, F f)" proof - have "C.inv (G (\ (g, f))) \\<^sub>C G (\ (g, f)) = G (F g \\<^sub>D F f)" using assms C.comp_inv_arr' cmp_components_are_iso C.inv_is_inverse C.VV.arr_char C.VV.ide_char cmp_simps(4) by auto moreover have "... = C.cod (\\<^sub>G.map (F g, F f))" using assms by simp ultimately have "(C.inv (G (\ (g, f))) \\<^sub>C G (\ (g, f))) \\<^sub>C \\<^sub>G.map (F g, F f) = \\<^sub>G.map (F g, F f)" using assms C.comp_cod_arr [of "\\<^sub>G.map (F g, F f)" "G (F g \\<^sub>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 \\<^sub>C \ f" using assms C.comp_inv_arr' C.comp_arr_dom [of "\ g \\<^sub>C \ f" "G (F g) \\<^sub>C G (F f)"] by simp finally show ?thesis by simp qed lemma \_inverts_GF_unit: assumes "C.obj a" shows "\ a \\<^sub>C GF.unit a = a" proof - have "\ a \\<^sub>C GF.unit a = I\<^sub>C.unit a" proof (intro I\<^sub>C.unit_eqI) show "C.obj a" by fact show 1: "\\ a \\<^sub>C GF.unit a : I\<^sub>C.map\<^sub>0 a \\<^sub>C I\<^sub>C.map a\" proof - have "src\<^sub>C (G (F a)) = src\<^sub>C (I\<^sub>C.map a)" using assms G\<^sub>0_props C.obj_def' by simp thus ?thesis using assms I\<^sub>C.map\<^sub>0_def GF.map\<^sub>0_def GF.unit_in_hom by (intro C.comp_in_homI') auto qed show "C.iso (\ a \\<^sub>C GF.unit a)" using assms \_in_hom GF.unit_char(2) by (intro C.isos_compose) auto show "(\ a \\<^sub>C GF.unit a) \\<^sub>C \\<^sub>C[I\<^sub>C.map\<^sub>0 a] = (I\<^sub>C.map \\<^sub>C[a] \\<^sub>C I\<^sub>C.cmp (a, a)) \\<^sub>C (\ a \\<^sub>C GF.unit a \\<^sub>C \ a \\<^sub>C GF.unit a)" proof - have "(I\<^sub>C.map \\<^sub>C[a] \\<^sub>C I\<^sub>C.cmp (a, a)) \\<^sub>C (\ a \\<^sub>C GF.unit a \\<^sub>C \ a \\<^sub>C GF.unit a) = \\<^sub>C[a] \\<^sub>C (a \\<^sub>C a) \\<^sub>C (\ a \\<^sub>C GF.unit a \\<^sub>C \ a \\<^sub>C GF.unit a)" using assms C.comp_assoc by simp also have "... = \\<^sub>C[a] \\<^sub>C (\ a \\<^sub>C GF.unit a \\<^sub>C \ a \\<^sub>C GF.unit a)" proof - have "C.hseq (\ a \\<^sub>C GF.unit a) (\ a \\<^sub>C GF.unit a)" using assms GF.unit_simps C.iso_is_arr \C.iso (\ a \\<^sub>C GF.unit a)\ by (intro C.hseqI') auto moreover have "a \\<^sub>C a = C.cod (\ a \\<^sub>C GF.unit a \\<^sub>C \ a \\<^sub>C GF.unit a)" proof - have "C.cod (\ a \\<^sub>C GF.unit a) = a" using assms 1 C.obj_simps by auto moreover have "C.hseq (\ a \\<^sub>C GF.unit a) (\ a \\<^sub>C GF.unit a)" using assms 1 C.src_dom [of "\ a \\<^sub>C GF.unit a"] C.trg_dom [of "\ a \\<^sub>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 "... = \\<^sub>C[a] \\<^sub>C (\ a \\<^sub>C \ a) \\<^sub>C (GF.unit a \\<^sub>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 \\<^sub>C GF.unit a)\) also have "... = \\<^sub>C[a] \\<^sub>C (\ (a \\<^sub>C a) \\<^sub>C GF.cmp (a, a)) \\<^sub>C (GF.unit a \\<^sub>C GF.unit a)" using assms \_hcomp' [of a a] by auto also have "... = (\\<^sub>C[a] \\<^sub>C \ (a \\<^sub>C a)) \\<^sub>C GF.cmp (a, a) \\<^sub>C (GF.unit a \\<^sub>C GF.unit a)" using C.comp_assoc by simp also have "... = (\ a \\<^sub>C G (F \\<^sub>C[a])) \\<^sub>C GF.cmp (a, a) \\<^sub>C (GF.unit a \\<^sub>C GF.unit a)" using assms \_naturality [of "\\<^sub>C[a]"] by simp also have "... = \ a \\<^sub>C (G (F \\<^sub>C[a]) \\<^sub>C GF.cmp (a, a)) \\<^sub>C (GF.unit a \\<^sub>C GF.unit a)" using C.comp_assoc by simp also have "... = \ a \\<^sub>C GF.unit a \\<^sub>C \\<^sub>C[GF.map\<^sub>0 a]" using assms GF.unit_char [of a] by simp also have "... = (\ a \\<^sub>C GF.unit a) \\<^sub>C \\<^sub>C[I\<^sub>C.map\<^sub>0 a]" proof - have "GF.map\<^sub>0 a = I\<^sub>C.map\<^sub>0 a" using assms G\<^sub>0_props(2) [of a] GF.map\<^sub>0_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 I\<^sub>C.unit_char' by simp qed lemma \_respects_comp: assumes "D.ide f" and "D.ide g" and "src\<^sub>D g = trg\<^sub>D f" shows "(\\<^sub>D\<^sup>-\<^sup>1[F (G (g \\<^sub>D f))] \\<^sub>D \ (g \\<^sub>D f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]) \\<^sub>D ((g \\<^sub>D f) \\<^sub>D src\<^sub>D f) = (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)] \\<^sub>D (\ g \\<^sub>D \ f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" and "(trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F (G g)] \\<^sub>D \ g \\<^sub>D \\<^sub>D[g] \\<^sub>D F (G f)) \\<^sub>D D.inv \\<^sub>D[g, src\<^sub>D g, F (G f)] \\<^sub>D (g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G f)] \\<^sub>D \ f \\<^sub>D \\<^sub>D[f]) \\<^sub>D \\<^sub>D[g, f, src\<^sub>D f] = (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)] \\<^sub>D (\ g \\<^sub>D \ f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" proof - show "(\\<^sub>D\<^sup>-\<^sup>1[F (G (g \\<^sub>D f))] \\<^sub>D \ (g \\<^sub>D f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]) \\<^sub>D ((g \\<^sub>D f) \\<^sub>D src\<^sub>D f) = (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)] \\<^sub>D (\ g \\<^sub>D \ f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F (G (g \\<^sub>D f))] \\<^sub>D \ (g \\<^sub>D f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]) \\<^sub>D ((g \\<^sub>D f) \\<^sub>D src\<^sub>D f) = \\<^sub>D\<^sup>-\<^sup>1[F (G (g \\<^sub>D f))] \\<^sub>D \ (g \\<^sub>D f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" using assms D.comp_assoc D.comp_arr_dom by simp also have 1: "... = (\\<^sub>D\<^sup>-\<^sup>1[F (G (g \\<^sub>D f))] \\<^sub>D F (\\<^sub>G.map (g, f))) \\<^sub>D \ (G g, G f) \\<^sub>D (\ g \\<^sub>D \ f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" using assms \_hcomp D.comp_assoc by simp also have "... = (trg\<^sub>D g \\<^sub>D F (\\<^sub>G.map (g, f))) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F (G g \\<^sub>C G f)] \\<^sub>D \ (G g, G f)) \\<^sub>D (\ g \\<^sub>D \ f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[F (G (g \\<^sub>D f))] \\<^sub>D F (\\<^sub>G.map (g, f)) = (trg\<^sub>D g \\<^sub>D F (\\<^sub>G.map (g, f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g \\<^sub>C G f)]" using assms G\<^sub>0_props D.lunit'_naturality [of "F (\\<^sub>G.map (g, f))"] \\<^sub>G_in_hom [of g f] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = ((trg\<^sub>D g \\<^sub>D F (\\<^sub>G.map (g, f))) \\<^sub>D (trg\<^sub>D g \\<^sub>D \ (G g, G f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)] \\<^sub>D (\ g \\<^sub>D \ f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[F (G g \\<^sub>C G f)] \\<^sub>D \ (G g, G f) = (trg\<^sub>D g \\<^sub>D \ (G g, G f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)]" using assms D.lunit'_naturality [of "\ (G g, G f)"] G\<^sub>0_props by fastforce thus ?thesis using D.comp_assoc by simp qed also have "... = (trg\<^sub>D g \\<^sub>D F (\\<^sub>G.map (g, f)) \\<^sub>D \ (G g, G f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)] \\<^sub>D (\ g \\<^sub>D \ f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" using assms 1 D.whisker_left [of "trg\<^sub>D g" "F (\\<^sub>G.map (g, f))" "\ (G g, G f)"] by force also have "... = (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)] \\<^sub>D (\ g \\<^sub>D \ f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" proof - have 1: "FG.cmp (g, f) = (F (G (g \\<^sub>D f)) \\<^sub>D F (\\<^sub>G.map (g, f))) \\<^sub>D \ (G g, G f)" using assms FG.cmp_def D.VV.arr_char D.VV.dom_char D.comp_assoc by simp also have "... = F (\\<^sub>G.map (g, f)) \\<^sub>D \ (G g, G f)" proof - have "D.cod (F (\\<^sub>G (g, f))) = F (G (g \\<^sub>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 (\\<^sub>G.map (g, f))" "F (G (g \\<^sub>D f))"] by fastforce qed finally have "FG.cmp (g, f) = F (\\<^sub>G.map (g, f)) \\<^sub>D \ (G g, G f)" by simp thus ?thesis by simp qed finally show ?thesis by simp qed show "(trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F (G g)] \\<^sub>D \ g \\<^sub>D \\<^sub>D[g] \\<^sub>D F (G f)) \\<^sub>D D.inv \\<^sub>D[g, src\<^sub>D g, F (G f)] \\<^sub>D (g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G f)] \\<^sub>D \ f \\<^sub>D \\<^sub>D[f]) \\<^sub>D \\<^sub>D[g, f, src\<^sub>D f] = (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)] \\<^sub>D (\ g \\<^sub>D \ f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" proof - have "(trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F (G g)] \\<^sub>D \ g \\<^sub>D \\<^sub>D[g] \\<^sub>D F (G f)) \\<^sub>D D.inv \\<^sub>D[g, src\<^sub>D g, F (G f)] \\<^sub>D (g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G f)] \\<^sub>D \ f \\<^sub>D \\<^sub>D[f]) \\<^sub>D \\<^sub>D[g, f, src\<^sub>D f] = (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F (G g)] \\<^sub>D F (G f)) \\<^sub>D (\ g \\<^sub>D F (G f)) \\<^sub>D (\\<^sub>D[g] \\<^sub>D F (G f)) \\<^sub>D D.inv \\<^sub>D[g, src\<^sub>D g, F (G f)] \\<^sub>D (g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G f)]) \\<^sub>D (g \\<^sub>D \ f) \\<^sub>D (g \\<^sub>D \\<^sub>D[f]) \\<^sub>D \\<^sub>D[g, f, src\<^sub>D f]" using assms D.comp_assoc D.whisker_right D.whisker_left by simp also have "... = (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D (\\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F (G g)] \\<^sub>D F (G f))) \\<^sub>D (\ g \\<^sub>D F (G f)) \\<^sub>D (\\<^sub>D[g] \\<^sub>D F (G f)) \\<^sub>D D.inv \\<^sub>D[g, src\<^sub>D g, F (G f)] \\<^sub>D (g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G f)]) \\<^sub>D (g \\<^sub>D \ f) \\<^sub>D ((g \\<^sub>D \\<^sub>D[f]) \\<^sub>D \\<^sub>D[g, f, src\<^sub>D f])" using assms D.comp_cod_arr D.VV.ide_char D.VV.arr_char D.VV.dom_char FG.FF_def G\<^sub>0_props D.comp_assoc by presburger also have "... = (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)] \\<^sub>D (\ g \\<^sub>D F (G f)) \\<^sub>D ((\\<^sub>D[g] \\<^sub>D F (G f)) \\<^sub>D D.inv \\<^sub>D[g, src\<^sub>D g, F (G f)]) \\<^sub>D (g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G f)]) \\<^sub>D (g \\<^sub>D \ f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" proof - have "(g \\<^sub>D \\<^sub>D[f]) \\<^sub>D \\<^sub>D[g, f, src\<^sub>D f] = \\<^sub>D[g \\<^sub>D f]" using assms D.runit_hcomp by simp moreover have "\\<^sub>D[trg\<^sub>D g, F (G g), F (G f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F (G g)] \\<^sub>D F (G f)) = \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)]" using assms D.lunit_hcomp [of "F (G g)" "F (G f)"] G\<^sub>0_props by simp ultimately show ?thesis using D.comp_assoc by simp qed also have "... = (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)] \\<^sub>D (\ g \\<^sub>D F (G f)) \\<^sub>D (((g \\<^sub>D \\<^sub>D[F (G f)]) \\<^sub>D (g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G f)])) \\<^sub>D (g \\<^sub>D \ f)) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" proof - have "(\\<^sub>D[g] \\<^sub>D F (G f)) \\<^sub>D D.inv \\<^sub>D[g, src\<^sub>D g, F (G f)] = g \\<^sub>D \\<^sub>D[F (G f)]" using assms D.triangle' [of g "F (G f)"] G\<^sub>0_props by simp thus ?thesis using D.comp_assoc by simp qed also have "... = (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)] \\<^sub>D ((\ g \\<^sub>D F (G f)) \\<^sub>D (g \\<^sub>D \ f)) \\<^sub>D \\<^sub>D[g \\<^sub>D f]" proof - have "((g \\<^sub>D \\<^sub>D[F (G f)]) \\<^sub>D (g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G f)])) \\<^sub>D (g \\<^sub>D \ f) = g \\<^sub>D \ f" using assms D.interchange [of g g "\\<^sub>D[F (G f)]" "\\<^sub>D\<^sup>-\<^sup>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 "... = (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G g) \\<^sub>D F (G f)] \\<^sub>D (\ g \\<^sub>D \ f) \\<^sub>D \\<^sub>D[g \\<^sub>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 \\<^sub>D FG.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[a] \\<^sub>D \\<^sub>D[a] = (\\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod a)] \\<^sub>D \ a \\<^sub>D \\<^sub>D[D.dom a]) \\<^sub>D (I\<^sub>D.unit a \\<^sub>D a)" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod a)] \\<^sub>D \ a \\<^sub>D \\<^sub>D[D.dom a]) \\<^sub>D (I\<^sub>D.unit a \\<^sub>D a) = (\\<^sub>D\<^sup>-\<^sup>1[F (G a)] \\<^sub>D \ a) \\<^sub>D \\<^sub>D[a]" using assms I\<^sub>D.lunit_coherence I\<^sub>D.unit_char' D.comp_arr_dom D.comp_assoc by auto also have "... = ((a \\<^sub>D \ a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[a]) \\<^sub>D \\<^sub>D[a]" using assms D.lunit'_naturality [of "\ a"] by auto also have "... = (a \\<^sub>D \ a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[a] \\<^sub>D \\<^sub>D[a]" using D.comp_assoc by simp also have "... = (a \\<^sub>D \ a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[a] \\<^sub>D \\<^sub>D[a]" using assms D.unitor_coincidence by simp also have "... = (a \\<^sub>D FG.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[a] \\<^sub>D \\<^sub>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 "src\<^sub>C g = trg\<^sub>C f" shows "(trg\<^sub>C g \\<^sub>C g \\<^sub>C f) \\<^sub>C \\<^sub>C[trg\<^sub>C g, g, f] \\<^sub>C (\\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>C \ g \\<^sub>C \\<^sub>C[G (F g)] \\<^sub>C f) \\<^sub>C C.inv \\<^sub>C[G (F g), src\<^sub>C g, f] \\<^sub>C (G (F g) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \ f \\<^sub>C \\<^sub>C[G (F f)]) \\<^sub>C \\<^sub>C[G (F g), G (F f), src\<^sub>C f] = \\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C (\ g \\<^sub>C \ f) \\<^sub>C \\<^sub>C[G (F g) \\<^sub>C G (F f)]" and "(\\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C \ (g \\<^sub>C f) \\<^sub>C \\<^sub>C[G (F (g \\<^sub>C f))]) \\<^sub>C (GF.cmp (g, f) \\<^sub>C src\<^sub>C f) = \\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C (\ g \\<^sub>C \ f) \\<^sub>C \\<^sub>C[G (F g) \\<^sub>C G (F f)]" proof - have "(trg\<^sub>C g \\<^sub>C g \\<^sub>C f) \\<^sub>C \\<^sub>C[trg\<^sub>C g, g, f] \\<^sub>C (\\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>C \ g \\<^sub>C \\<^sub>C[G (F g)] \\<^sub>C f) \\<^sub>C C.inv \\<^sub>C[G (F g), src\<^sub>C g, f] \\<^sub>C (G (F g) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \ f \\<^sub>C \\<^sub>C[G (F f)]) \\<^sub>C \\<^sub>C[G (F g), G (F f), src\<^sub>C f] = ((trg\<^sub>C g \\<^sub>C g \\<^sub>C f) \\<^sub>C \\<^sub>C[trg\<^sub>C g, g, f]) \\<^sub>C (\\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>C \ g \\<^sub>C \\<^sub>C[G (F g)] \\<^sub>C f) \\<^sub>C C.inv \\<^sub>C[G (F g), src\<^sub>C g, f] \\<^sub>C (G (F g) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \ f \\<^sub>C \\<^sub>C[G (F f)]) \\<^sub>C \\<^sub>C[G (F g), G (F f), src\<^sub>C f]" using assms C.comp_assoc by simp also have "... = \\<^sub>C[trg\<^sub>C g, g, f] \\<^sub>C (\\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>C \ g \\<^sub>C \\<^sub>C[G (F g)] \\<^sub>C f) \\<^sub>C C.inv \\<^sub>C[G (F g), src\<^sub>C g, f] \\<^sub>C (G (F g) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \ f \\<^sub>C \\<^sub>C[G (F f)]) \\<^sub>C \\<^sub>C[G (F g), G (F f), src\<^sub>C f]" using assms C.comp_cod_arr by simp also have "... = (\\<^sub>C[trg\<^sub>C g, g, f] \\<^sub>C (\\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>C f)) \\<^sub>C (\ g \\<^sub>C f) \\<^sub>C (\\<^sub>C[G (F g)] \\<^sub>C f) \\<^sub>C C.inv \\<^sub>C[G (F g), src\<^sub>C g, f] \\<^sub>C (G (F g) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f]) \\<^sub>C (G (F g) \\<^sub>C \ f) \\<^sub>C ((G (F g) \\<^sub>C \\<^sub>C[G (F f)]) \\<^sub>C \\<^sub>C[G (F g), G (F f), src\<^sub>C f])" using assms C.whisker_left C.whisker_right C.comp_assoc by simp also have "... = \\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C (\ g \\<^sub>C f) \\<^sub>C ((\\<^sub>C[G (F g)] \\<^sub>C f) \\<^sub>C C.inv \\<^sub>C[G (F g), src\<^sub>C g, f]) \\<^sub>C (G (F g) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f]) \\<^sub>C (G (F g) \\<^sub>C \ f) \\<^sub>C \\<^sub>C[G (F g) \\<^sub>C G (F f)]" using assms G\<^sub>0_props C.lunit_hcomp C.runit_hcomp C.comp_assoc by simp also have "... = \\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C (\ g \\<^sub>C f) \\<^sub>C (((G (F g) \\<^sub>C \\<^sub>C[f]) \\<^sub>C (G (F g) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f])) \\<^sub>C (G (F g) \\<^sub>C \ f)) \\<^sub>C \\<^sub>C[G (F g) \\<^sub>C G (F f)]" using assms G\<^sub>0_props C.triangle' C.comp_assoc by simp also have "... = \\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C ((\ g \\<^sub>C f) \\<^sub>C (G (F g) \\<^sub>C \ f)) \\<^sub>C \\<^sub>C[G (F g) \\<^sub>C G (F f)]" proof - have "(G (F g) \\<^sub>C \\<^sub>C[f]) \\<^sub>C (G (F g) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f]) = G (F g) \\<^sub>C f" using assms C.whisker_left [of "G (F g)" "\\<^sub>C[f]" "\\<^sub>C\<^sup>-\<^sup>1[f]"] C.comp_arr_inv' by simp moreover have "... = C.cod (G (F g) \\<^sub>C \ f)" using assms G\<^sub>0_props by auto moreover have "C.hseq (G (F g)) (\ f)" using assms G\<^sub>0_props by simp ultimately have "((G (F g) \\<^sub>C \\<^sub>C[f]) \\<^sub>C (G (F g) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f])) \\<^sub>C (G (F g) \\<^sub>C \ f) = (G (F g) \\<^sub>C \ f)" using assms G\<^sub>0_props C.comp_cod_arr by presburger thus ?thesis using C.comp_assoc by simp qed also have "... = \\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C (\ g \\<^sub>C \ f) \\<^sub>C \\<^sub>C[G (F g) \\<^sub>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 "(trg\<^sub>C g \\<^sub>C g \\<^sub>C f) \\<^sub>C \\<^sub>C[trg\<^sub>C g, g, f] \\<^sub>C (\\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>C \ g \\<^sub>C \\<^sub>C[G (F g)] \\<^sub>C f) \\<^sub>C C.inv \\<^sub>C[G (F g), src\<^sub>C g, f] \\<^sub>C (G (F g) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \ f \\<^sub>C \\<^sub>C[G (F f)]) \\<^sub>C \\<^sub>C[G (F g), G (F f), src\<^sub>C f] = \\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C (\ g \\<^sub>C \ f) \\<^sub>C \\<^sub>C[G (F g) \\<^sub>C G (F f)]" by simp show "(\\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C \ (g \\<^sub>C f) \\<^sub>C \\<^sub>C[G (F (g \\<^sub>C f))]) \\<^sub>C (GF.cmp (g, f) \\<^sub>C src\<^sub>C f) = ..." proof - have "(\\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C \ (g \\<^sub>C f) \\<^sub>C \\<^sub>C[G (F (g \\<^sub>C f))]) \\<^sub>C (GF.cmp (g, f) \\<^sub>C src\<^sub>C f) = \\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C \ (g \\<^sub>C f) \\<^sub>C \\<^sub>C[G (F (g \\<^sub>C f))] \\<^sub>C (GF.cmp (g, f) \\<^sub>C src\<^sub>C f)" using assms C.comp_assoc by simp also have "... = \\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C (\ (g \\<^sub>C f) \\<^sub>C GF.cmp (g, f)) \\<^sub>C \\<^sub>C[G (F g) \\<^sub>C G (F f)]" proof - have "src\<^sub>C (GF.cmp (g, f)) = src\<^sub>C f" using assms G\<^sub>0_props by simp hence "\\<^sub>C[G (F (g \\<^sub>C f))] \\<^sub>C (GF.cmp (g, f) \\<^sub>C src\<^sub>C f) = GF.cmp (g, f) \\<^sub>C \\<^sub>C[G (F g) \\<^sub>C G (F f)]" using assms C.runit_naturality [of "GF.cmp (g, f)"] C.VV.arr_char C.VV.cod_char GF.cmp_simps'(1,4-5) by simp thus ?thesis using C.comp_assoc by simp qed also have "... = \\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C (\ g \\<^sub>C \ f) \\<^sub>C \\<^sub>C[G (F g) \\<^sub>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 \\<^sub>C I\<^sub>C.unit a) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[a] \\<^sub>C \\<^sub>C[a] = (\\<^sub>C\<^sup>-\<^sup>1[C.cod a] \\<^sub>C \ a \\<^sub>C \\<^sub>C[GF.map (C.dom a)]) \\<^sub>C (GF.unit a \\<^sub>C a)" proof - have "(\\<^sub>C\<^sup>-\<^sup>1[C.cod a] \\<^sub>C \ a \\<^sub>C \\<^sub>C[GF.map (C.dom a)]) \\<^sub>C (GF.unit a \\<^sub>C a) = \\<^sub>C\<^sup>-\<^sup>1[a] \\<^sub>C \ a \\<^sub>C \\<^sub>C[G (F a)] \\<^sub>C (GF.unit a \\<^sub>C a)" using assms C.comp_assoc by auto also have "... = (\\<^sub>C\<^sup>-\<^sup>1[a] \\<^sub>C \ a) \\<^sub>C GF.unit a \\<^sub>C \\<^sub>C[a]" proof - have "src\<^sub>C (GF.unit a) = a" using assms GF.unit_simps(2) GF.map\<^sub>0_def [of a] G\<^sub>0_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 \\<^sub>C \ a) \\<^sub>C (\\<^sub>C\<^sup>-\<^sup>1[G (F a)] \\<^sub>C GF.unit a) \\<^sub>C \\<^sub>C[a]" proof - have "\\<^sub>C\<^sup>-\<^sup>1[a] \\<^sub>C \ a = (a \\<^sub>C \ a) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>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 \\<^sub>C \ a) \\<^sub>C (a \\<^sub>C GF.unit a)) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[a] \\<^sub>C \\<^sub>C[a]" proof - have "\\<^sub>C\<^sup>-\<^sup>1[G (F a)] \\<^sub>C GF.unit a = (a \\<^sub>C GF.unit a) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[a]" using assms C.lunit'_naturality [of "GF.unit a"] G\<^sub>0_props C.obj_simps by (simp add: GF.map\<^sub>0_def) thus ?thesis using C.comp_assoc by simp qed also have "... = (a \\<^sub>C \ a \\<^sub>C GF.unit a) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[a] \\<^sub>C \\<^sub>C[a]" using assms C.interchange [of a a "\ a" "GF.unit a"] by force also have "... = (a \\<^sub>C \ a \\<^sub>C GF.unit a) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[a] \\<^sub>C \\<^sub>C[a]" using assms C.unitor_coincidence by simp also have "... = (a \\<^sub>C I\<^sub>C.unit a) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[a] \\<^sub>C \\<^sub>C[a]" using assms \_inverts_GF_unit I\<^sub>C.unit_char' by simp finally show ?thesis by simp qed - abbreviation counit\<^sub>0 - where "counit\<^sub>0 \ \b. b" - - abbreviation counit\<^sub>1 - where "counit\<^sub>1 \ \g. \\<^sub>D\<^sup>-\<^sup>1[F (G g)] \\<^sub>D \ g \\<^sub>D \\<^sub>D[g]" + (* + * 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 counit\<^sub>0' + where "counit\<^sub>0' \ \b. b" + + abbreviation counit\<^sub>1' + where "counit\<^sub>1' \ \g. \\<^sub>D\<^sup>-\<^sup>1[F (G g)] \\<^sub>D \ g \\<^sub>D \\<^sub>D[g]" interpretation \: pseudonatural_equivalence V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D - FG.map FG.cmp I\<^sub>D.map I\<^sub>D.cmp counit\<^sub>0 counit\<^sub>1 + FG.map FG.cmp I\<^sub>D.map I\<^sub>D.cmp counit\<^sub>0' counit\<^sub>1' proof show "\a. D.obj a \ D.ide a" by auto show "\f. D.ide f \ D.iso (\\<^sub>D\<^sup>-\<^sup>1[F (G f)] \\<^sub>D \ f \\<^sub>D \\<^sub>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 : src\<^sub>D (FG.map a) \\<^sub>D src\<^sub>D (I\<^sub>D.map a)\" using D.obj_def G\<^sub>0_props(1) by auto show "\f. D.ide f \ \\\<^sub>D\<^sup>-\<^sup>1[F (G f)] \\<^sub>D \ f \\<^sub>D \\<^sub>D[f] : I\<^sub>D.map f \\<^sub>D src\<^sub>D f \\<^sub>D trg\<^sub>D f \\<^sub>D FG.map f\" using G\<^sub>0_props by (intro D.comp_in_homI') auto show "\\. D.arr \ \ (\\<^sub>D\<^sup>-\<^sup>1[F (G (D.cod \))] \\<^sub>D \ (D.cod \) \\<^sub>D \\<^sub>D[D.cod \]) \\<^sub>D (I\<^sub>D.map \ \\<^sub>D src\<^sub>D \) = (trg\<^sub>D \ \\<^sub>D FG.map \) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G (D.dom \))] \\<^sub>D \ (D.dom \) \\<^sub>D \\<^sub>D[D.dom \]" proof - fix \ assume \: "D.arr \" have "(\\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \)] \\<^sub>D \ (D.cod \) \\<^sub>D \\<^sub>D[D.cod \]) \\<^sub>D (I\<^sub>D.map \ \\<^sub>D src\<^sub>D \) = \\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \)] \\<^sub>D \ \ \\<^sub>D \\<^sub>D[D.dom \]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \)] \\<^sub>D \ (D.cod \) \\<^sub>D \\<^sub>D[D.cod \]) \\<^sub>D (I\<^sub>D.map \ \\<^sub>D src\<^sub>D \) = \\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \)] \\<^sub>D (\ (D.cod \) \\<^sub>D \) \\<^sub>D \\<^sub>D[D.dom \]" using \ D.runit_naturality D.comp_assoc by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \)] \\<^sub>D \ \ \\<^sub>D \\<^sub>D[D.dom \]" using \ \_naturality(1) by simp finally show ?thesis by blast qed also have "... = (trg\<^sub>D \ \\<^sub>D FG.map \) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[FG.map (D.dom \)] \\<^sub>D \ (D.dom \) \\<^sub>D \\<^sub>D[D.dom \]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \)] \\<^sub>D \ \ \\<^sub>D \\<^sub>D[D.dom \] = \\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \)] \\<^sub>D (FG.map \ \\<^sub>D \ (D.dom \)) \\<^sub>D \\<^sub>D[D.dom \]" using \ \_naturality(2) D.comp_assoc by simp also have "... = (\\<^sub>D\<^sup>-\<^sup>1[FG.map (D.cod \)] \\<^sub>D FG.map \) \\<^sub>D \ (D.dom \) \\<^sub>D \\<^sub>D[D.dom \]" using D.comp_assoc by simp also have "... = ((trg\<^sub>D \ \\<^sub>D FG.map \) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[FG.map (D.dom \)]) \\<^sub>D \ (D.dom \) \\<^sub>D \\<^sub>D[D.dom \]" using \ D.lunit'_naturality [of "FG.map \"] G\<^sub>0_props by simp also have "... = (trg\<^sub>D \ \\<^sub>D FG.map \) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[FG.map (D.dom \)] \\<^sub>D \ (D.dom \) \\<^sub>D \\<^sub>D[D.dom \]" using D.comp_assoc by simp finally show ?thesis by blast qed finally show "(\\<^sub>D\<^sup>-\<^sup>1[F (G (D.cod \))] \\<^sub>D \ (D.cod \) \\<^sub>D \\<^sub>D[D.cod \]) \\<^sub>D (I\<^sub>D.map \ \\<^sub>D src\<^sub>D \) = (trg\<^sub>D \ \\<^sub>D FG.map \) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G (D.dom \))] \\<^sub>D \ (D.dom \) \\<^sub>D \\<^sub>D[D.dom \]" by simp qed show "\f g. \D.ide f; D.ide g; src\<^sub>D g = trg\<^sub>D f\ \ (trg\<^sub>D g \\<^sub>D FG.cmp (g, f)) \\<^sub>D \\<^sub>D[trg\<^sub>D g, FG.map g, FG.map f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F (G g)] \\<^sub>D \ g \\<^sub>D \\<^sub>D[g] \\<^sub>D FG.map f) \\<^sub>D D.inv \\<^sub>D[I\<^sub>D.map g, src\<^sub>D g, FG.map f] \\<^sub>D (I\<^sub>D.map g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (G f)] \\<^sub>D \ f \\<^sub>D \\<^sub>D[f]) \\<^sub>D \\<^sub>D[I\<^sub>D.map g, I\<^sub>D.map f, src\<^sub>D f] = (\\<^sub>D\<^sup>-\<^sup>1[F (G (g \\<^sub>D f))] \\<^sub>D \ (g \\<^sub>D f) \\<^sub>D \\<^sub>D[g \\<^sub>D f]) \\<^sub>D (I\<^sub>D.cmp (g, f) \\<^sub>D src\<^sub>D f)" using \_respects_comp by simp show "\a. D.obj a \ (a \\<^sub>D FG.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[a] \\<^sub>D \\<^sub>D[a] = (\\<^sub>D\<^sup>-\<^sup>1[F (G a)] \\<^sub>D \ a \\<^sub>D \\<^sub>D[a]) \\<^sub>D (I\<^sub>D.unit a \\<^sub>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 unit\<^sub>0 - where "unit\<^sub>0 \ \a. a" - - abbreviation unit\<^sub>1 - where "unit\<^sub>1 \ \f. \\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \ f \\<^sub>C \\<^sub>C[G (F f)]" + abbreviation unit\<^sub>0' + where "unit\<^sub>0' \ \a. a" + + abbreviation unit\<^sub>1' + where "unit\<^sub>1' \ \f. \\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \ f \\<^sub>C \\<^sub>C[G (F f)]" interpretation \: pseudonatural_equivalence V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C - I\<^sub>C.map I\<^sub>C.cmp GF.map GF.cmp unit\<^sub>0 unit\<^sub>1 + I\<^sub>C.map I\<^sub>C.cmp GF.map GF.cmp unit\<^sub>0' unit\<^sub>1' proof show "\a. C.obj a \ C.ide a" by auto show "\f. C.ide f \ C.iso (\\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \ f \\<^sub>C \\<^sub>C[G (F f)])" using C.iso_runit C.iso_lunit' by (intro C.isos_compose) auto show "\a. C.obj a \ \a : src\<^sub>C (I\<^sub>C.map a) \\<^sub>C src\<^sub>C (GF.map a)\" by (simp_all add: C.obj_simps(1-3) G\<^sub>0_props(2)) show "\f. C.ide f \ \\\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \ f \\<^sub>C \\<^sub>C[G (F f)] : GF.map f \\<^sub>C src\<^sub>C f \\<^sub>C trg\<^sub>C f \\<^sub>C I\<^sub>C.map f\" using G\<^sub>0_props by auto show "\\. C.arr \ \ (\\<^sub>C\<^sup>-\<^sup>1[C.cod \] \\<^sub>C \ (C.cod \) \\<^sub>C \\<^sub>C[G (F (C.cod \))]) \\<^sub>C (GF.map \ \\<^sub>C src\<^sub>C \) = (trg\<^sub>C \ \\<^sub>C I\<^sub>C.map \) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[C.dom \] \\<^sub>C \ (C.dom \) \\<^sub>C \\<^sub>C[G (F (C.dom \))]" proof - fix \ assume \: "C.arr \" have "(\\<^sub>C\<^sup>-\<^sup>1[C.cod \] \\<^sub>C \ (C.cod \) \\<^sub>C \\<^sub>C[GF.map (C.cod \)]) \\<^sub>C (GF.map \ \\<^sub>C src\<^sub>C \) = \\<^sub>C\<^sup>-\<^sup>1[C.cod \] \\<^sub>C \ \ \\<^sub>C \\<^sub>C[GF.map (C.dom \)]" proof - have "(\\<^sub>C\<^sup>-\<^sup>1[C.cod \] \\<^sub>C \ (C.cod \) \\<^sub>C \\<^sub>C[GF.map (C.cod \)]) \\<^sub>C (GF.map \ \\<^sub>C src\<^sub>C \) = \\<^sub>C\<^sup>-\<^sup>1[C.cod \] \\<^sub>C \ (C.cod \) \\<^sub>C \\<^sub>C[GF.map (C.cod \)] \\<^sub>C (GF.map \ \\<^sub>C src\<^sub>C \)" using C.comp_assoc by simp also have "... = \\<^sub>C\<^sup>-\<^sup>1[C.cod \] \\<^sub>C (\ (C.cod \) \\<^sub>C GF.map \) \\<^sub>C \\<^sub>C[GF.map (C.dom \)]" using \ C.runit_naturality [of "GF.map \"] G\<^sub>0_props C.comp_assoc by simp also have "... = \\<^sub>C\<^sup>-\<^sup>1[C.cod \] \\<^sub>C \ \ \\<^sub>C \\<^sub>C[GF.map (C.dom \)]" using \ \_naturality(1) [of \] by simp finally show ?thesis by blast qed also have "... = (trg\<^sub>C \ \\<^sub>C I\<^sub>C.map \) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[C.dom \] \\<^sub>C \ (C.dom \) \\<^sub>C \\<^sub>C[GF.map (C.dom \)]" proof - have "... = \\<^sub>C\<^sup>-\<^sup>1[C.cod \] \\<^sub>C (\ \\<^sub>C \ (C.dom \)) \\<^sub>C \\<^sub>C[GF.map (C.dom \)]" using \ \_naturality(2) [of \] by simp also have "... = (\\<^sub>C\<^sup>-\<^sup>1[C.cod \] \\<^sub>C \) \\<^sub>C \ (C.dom \) \\<^sub>C \\<^sub>C[GF.map (C.dom \)]" using C.comp_assoc by simp also have "... = ((trg\<^sub>C \ \\<^sub>C I\<^sub>C.map \) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[C.dom \]) \\<^sub>C \ (C.dom \) \\<^sub>C \\<^sub>C[GF.map (C.dom \)]" using \ C.lunit'_naturality [of \] by simp also have "... = (trg\<^sub>C \ \\<^sub>C I\<^sub>C.map \) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[C.dom \] \\<^sub>C \ (C.dom \) \\<^sub>C \\<^sub>C[GF.map (C.dom \)]" using C.comp_assoc by simp finally show ?thesis by blast qed finally show "(\\<^sub>C\<^sup>-\<^sup>1[C.cod \] \\<^sub>C \ (C.cod \) \\<^sub>C \\<^sub>C[G (F (C.cod \))]) \\<^sub>C (GF.map \ \\<^sub>C src\<^sub>C \) = (trg\<^sub>C \ \\<^sub>C I\<^sub>C.map \) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[C.dom \] \\<^sub>C \ (C.dom \) \\<^sub>C \\<^sub>C[G (F (C.dom \))]" by simp qed show "\f g. \C.ide f; C.ide g; src\<^sub>C g = trg\<^sub>C f\ \ (trg\<^sub>C g \\<^sub>C I\<^sub>C.cmp (g, f)) \\<^sub>C \\<^sub>C[trg\<^sub>C g, I\<^sub>C.map g, I\<^sub>C.map f] \\<^sub>C (\\<^sub>C\<^sup>-\<^sup>1[g] \\<^sub>C \ g \\<^sub>C \\<^sub>C[G (F g)] \\<^sub>C I\<^sub>C.map f) \\<^sub>C C.inv \\<^sub>C[GF.map g, src\<^sub>C g, I\<^sub>C.map f] \\<^sub>C (GF.map g \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[f] \\<^sub>C \ f \\<^sub>C \\<^sub>C[G (F f)]) \\<^sub>C \\<^sub>C[GF.map g, GF.map f, src\<^sub>C f] = (\\<^sub>C\<^sup>-\<^sup>1[g \\<^sub>C f] \\<^sub>C \ (g \\<^sub>C f) \\<^sub>C \\<^sub>C[G (F (g \\<^sub>C f))]) \\<^sub>C (GF.cmp (g, f) \\<^sub>C src\<^sub>C f)" using \_respects_comp by simp show "\a. C.obj a \ (a \\<^sub>C I\<^sub>C.unit a) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[a] \\<^sub>C \\<^sub>C[a] = (\\<^sub>C\<^sup>-\<^sup>1[a] \\<^sub>C \ a \\<^sub>C \\<^sub>C[G (F a)]) \\<^sub>C (GF.unit a \\<^sub>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 V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C - F \ G \\<^sub>G.map unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1 + F \ G \\<^sub>G.map unit\<^sub>0' unit\<^sub>1' counit\<^sub>0' counit\<^sub>1' .. lemma extends_to_equivalence_of_bicategories: shows "equivalence_of_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C - F \ G \\<^sub>G.map unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1" + F \ G \\<^sub>G.map unit\<^sub>0' unit\<^sub>1' counit\<^sub>0' counit\<^sub>1'" .. 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' \ map\<^sub>0 a' = map\<^sub>0 a)}" lemma U_dense: assumes "C.obj a" shows "\a' \ U. C.equivalent_objects a a'" proof - let ?a' = "SOME a'. C.obj a' \ map\<^sub>0 a' = map\<^sub>0 a" have "\a'. C.obj a' \ map\<^sub>0 a' = map\<^sub>0 a" using assms by auto hence 1: "?a' \ U \ map\<^sub>0 ?a' = map\<^sub>0 a" using assms U_def someI_ex [of "\a'. C.obj a' \ map\<^sub>0 a' = map\<^sub>0 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 = map\<^sub>0 ` Collect C.obj" lemma V_dense: assumes "D.obj b" shows "\b'. b' \ map\<^sub>0 ` 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 map\<^sub>0 U V" proof - have "inj_on map\<^sub>0 U" using U_def by (intro inj_onI) simp moreover have "map\<^sub>0 ` U = V" proof show "map\<^sub>0 ` U \ V" using U_def V_def by auto show "V \ map\<^sub>0 ` U" proof fix b assume b: "b \ V" obtain a where a: "C.obj a \ map\<^sub>0 a = b" using b V_def by auto let ?a' = "SOME a'. C.obj a' \ map\<^sub>0 a' = map\<^sub>0 a" have 1: "C.obj ?a' \ map\<^sub>0 ?a' = b" using a someI_ex [of "\a'. C.obj a' \ map\<^sub>0 a' = map\<^sub>0 a"] by auto moreover have "?a' = (SOME a''. C.obj a'' \ map\<^sub>0 a'' = map\<^sub>0 ?a')" using a 1 by simp ultimately have "?a' \ U" unfolding U_def by simp thus "b \ map\<^sub>0 ` U" using a 1 U_def by (metis (no_types, lifting) image_eqI) qed qed ultimately show ?thesis using bij_betw_def [of map\<^sub>0 U V] by simp qed abbreviation (input) Arr\<^sub>U where "Arr\<^sub>U \ \\. C.arr \ \ src\<^sub>C \ \ U \ trg\<^sub>C \ \ U" interpretation C\<^sub>U: subbicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C Arr\<^sub>U 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 C\<^sub>U: dense_subbicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C U (* 25 sec *) + interpretation C\<^sub>U: dense_subbicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C 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 (* 15 sec *) + qed (* 25 sec *) abbreviation (input) Arr\<^sub>V where "Arr\<^sub>V \ \\. D.arr \ \ src\<^sub>D \ \ V \ trg\<^sub>D \ \ V" interpretation D\<^sub>V: subbicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D Arr\<^sub>V 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 D\<^sub>V: dense_subbicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V (* 35 sec *) + interpretation D\<^sub>V: dense_subbicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V (* 25 sec *) using V_dense D.equivalent_objects_def D.equivalent_objects_symmetric V_def - by unfold_locales metis (* 20 sec -- time is similar doing it this way. *) + by unfold_locales metis (* 35 sec -- time is similar doing it this way. *) interpretation F\<^sub>U: restricted_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ \\\. C.arr \ \ src\<^sub>C \ \ U \ trg\<^sub>C \ \ U\ .. interpretation F\<^sub>U\<^sub>V: corestricted_pseudofunctor C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\ \\<^sub>C C\<^sub>U.src C\<^sub>U.trg V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F\<^sub>U.map F\<^sub>U.cmp \\a. a \ V\ proof show "\\. C\<^sub>U.arr \ \ D\<^sub>V.arr (F\<^sub>U.map \)" proof - fix \ assume \: "C\<^sub>U.arr \" have "src\<^sub>C \ \ U \ trg\<^sub>C \ \ U" using \ C\<^sub>U.arr_char by simp moreover have "\a. a \ U \ F\<^sub>U.map\<^sub>0 a = map\<^sub>0 a" proof - fix a :: 'a assume a: "a \ U" hence 1: "C.obj a" using U_def by blast have "src\<^sub>C a = a" using a U_def by blast thus "F\<^sub>U.map\<^sub>0 a = map\<^sub>0 a" using a 1 C.obj_simps(1,3) C\<^sub>U.arr_char F\<^sub>U.map\<^sub>0_def map\<^sub>0_def by presburger qed ultimately have "F\<^sub>U.map\<^sub>0 (src\<^sub>C \) \ V \ F\<^sub>U.map\<^sub>0 (trg\<^sub>C \) \ V" using \ bij_betw_U_V bij_betw_def by (metis (no_types, lifting) image_eqI) hence "src\<^sub>D (F\<^sub>U.map \) \ V \ trg\<^sub>D (F\<^sub>U.map \) \ V" using \ F\<^sub>U.map\<^sub>0_def C\<^sub>U.src_def C\<^sub>U.trg_def F\<^sub>U.preserves_src F\<^sub>U.preserves_trg by auto thus "D\<^sub>V.arr (F\<^sub>U.map \)" using \ D\<^sub>V.arr_char [of "F\<^sub>U.map \"] by blast qed qed interpretation F\<^sub>U\<^sub>V: equivalence_pseudofunctor C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\ \\<^sub>C C\<^sub>U.src C\<^sub>U.trg D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\ \\<^sub>D D\<^sub>V.src D\<^sub>V.trg F\<^sub>U\<^sub>V.map F\<^sub>U\<^sub>V.cmp proof show "\f f'. \C\<^sub>U.par f f'; F\<^sub>U.map f = F\<^sub>U.map f'\ \ f = f'" using C\<^sub>U.arr_char C\<^sub>U.dom_char C\<^sub>U.cod_char by (metis (no_types, lifting) is_faithful) show "\a b g. \C\<^sub>U.obj a; C\<^sub>U.obj b; D\<^sub>V.in_hhom g (F\<^sub>U\<^sub>V.map\<^sub>0 a) (F\<^sub>U\<^sub>V.map\<^sub>0 b); D\<^sub>V.ide g\ \ \f. C\<^sub>U.in_hhom f a b \ C\<^sub>U.ide f \ D\<^sub>V.isomorphic (F\<^sub>U.map f) g" proof - fix a b g assume a: "C\<^sub>U.obj a" and b: "C\<^sub>U.obj b" assume g: "D\<^sub>V.in_hhom g (F\<^sub>U\<^sub>V.map\<^sub>0 a) (F\<^sub>U\<^sub>V.map\<^sub>0 b)" and ide_g: "D\<^sub>V.ide g" have 1: "\f. \f : a \\<^sub>C b\ \ C.ide f \ D.isomorphic (F f) g" proof - have "C.obj a \ C.obj b" using a b C\<^sub>U.obj_char by simp moreover have "\g : map\<^sub>0 a \\<^sub>D map\<^sub>0 b\" using a b g D\<^sub>V.in_hhom_def D\<^sub>V.arr_char D\<^sub>V.src_def D\<^sub>V.trg_def by (intro D.in_hhomI) auto moreover have "D.ide g" using ide_g D\<^sub>V.ide_char D\<^sub>V.arr_char by simp ultimately show ?thesis using locally_essentially_surjective by simp qed obtain f where f: "\f : a \\<^sub>C b\ \ C.ide f \ D.isomorphic (F f) g" using 1 by blast have 2: "C\<^sub>U.in_hhom f a b" using f a b C\<^sub>U.arr_char C\<^sub>U.obj_char C\<^sub>U.src_def C\<^sub>U.trg_def by fastforce moreover have "C\<^sub>U.ide f" using f 2 C\<^sub>U.ide_char C\<^sub>U.arr_char by auto moreover have "D\<^sub>V.isomorphic (F\<^sub>U\<^sub>V.map f) g" proof - obtain \ where \: "\\ : F f \\<^sub>D g\ \ D.iso \" using f D.isomorphic_def by auto have 3: "D\<^sub>V.in_hom \ (F\<^sub>U\<^sub>V.map f) g" proof - have "D\<^sub>V.arr \" using f g \ 2 D\<^sub>V.arr_char by (metis (no_types, lifting) D.arrI D.vconn_implies_hpar(1-4) D\<^sub>V.ideD(1) ide_g) thus ?thesis using \ 2 D\<^sub>V.dom_char D\<^sub>V.cod_char by (intro D\<^sub>V.in_homI) auto qed moreover have "D\<^sub>V.iso \" using \ D\<^sub>V.iso_char D\<^sub>V.arr_char 3 by fastforce ultimately show ?thesis using D\<^sub>V.isomorphic_def by auto qed ultimately show "\f. C\<^sub>U.in_hhom f a b \ C\<^sub>U.ide f \ D\<^sub>V.isomorphic (F\<^sub>U\<^sub>V.map f) g" by auto qed show "\f f' \. \C\<^sub>U.ide f; C\<^sub>U.ide f'; C\<^sub>U.src f = C\<^sub>U.src f'; C\<^sub>U.trg f = C\<^sub>U.trg f'; D\<^sub>V.in_hom \ (F\<^sub>U.map f) (F\<^sub>U.map f')\ \ \\. C\<^sub>U.in_hom \ f f' \ F\<^sub>U.map \ = \" proof - fix f f' \ assume f: "C\<^sub>U.ide f" and f': "C\<^sub>U.ide f'" and eq_src: "C\<^sub>U.src f = C\<^sub>U.src f'" and eq_trg: "C\<^sub>U.trg f = C\<^sub>U.trg f'" and \: "D\<^sub>V.in_hom \ (F\<^sub>U.map f) (F\<^sub>U.map f')" have 1: "C.ide f \ C.ide f'" using f f' C\<^sub>U.ide_char C\<^sub>U.arr_char by simp have 2: "\\. \\ : f \\<^sub>C f'\ \ F \ = \" proof - have "src\<^sub>C f = src\<^sub>C f' \ trg\<^sub>C f = trg\<^sub>C f'" using f f' 1 eq_src eq_trg C\<^sub>U.src_def C\<^sub>U.trg_def by simp moreover have "\\ : F f \\<^sub>D F f'\" using \ D\<^sub>V.in_hom_char D\<^sub>V.arr_char by auto ultimately show ?thesis using 1 locally_full by simp qed obtain \ where \: "\\ : f \\<^sub>C f'\ \ F \ = \" using 2 by auto have 3: "C\<^sub>U.in_hom \ f f'" using \ f f' C\<^sub>U.in_hom_char C\<^sub>U.ide_char C\<^sub>U.arr_char by auto moreover have "F\<^sub>U.map \ = \" using \ 3 by auto ultimately show "\\. C\<^sub>U.in_hom \ f f' \ F\<^sub>U.map \ = \" by auto qed show "\b. D\<^sub>V.obj b \ \a. C\<^sub>U.obj a \ D\<^sub>V.equivalent_objects (F\<^sub>U\<^sub>V.map\<^sub>0 a) b" proof - fix b assume b: "D\<^sub>V.obj b" obtain a where a: "C.obj a \ map\<^sub>0 a = b" using b D\<^sub>V.obj_char D\<^sub>V.arr_char V_def by auto have 1: "D.obj b \ b \ V" using a b D\<^sub>V.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': "C\<^sub>U.obj a'" using a' U_def C\<^sub>U.obj_char C\<^sub>U.arr_char by fastforce moreover have "D\<^sub>V.equivalent_objects (F\<^sub>U\<^sub>V.map\<^sub>0 a') b" proof - have "D.equivalent_objects (map\<^sub>0 a') (map\<^sub>0 a)" using a' preserves_equivalent_objects by simp hence 2: "D.equivalent_objects (map\<^sub>0 a') b" using a D.equivalent_objects_transitive by simp obtain e where e: "\e : map\<^sub>0 a' \\<^sub>D b\ \ D.equivalence_map e" using 2 D.equivalent_objects_def by auto have 3: "D.obj (map\<^sub>0 a') \ map\<^sub>0 a' \ V" using a' e U_def V_def by simp moreover have e_in_hhom: "D\<^sub>V.in_hhom e (F\<^sub>U\<^sub>V.map\<^sub>0 a') b" proof show 4: "D\<^sub>V.arr e" using 1 3 a e b D\<^sub>V.arr_char D\<^sub>V.obj_char V_def by (metis (no_types, lifting) D.in_hhomE) show "D\<^sub>V.src e = F\<^sub>U\<^sub>V.map\<^sub>0 a'" using e 4 D\<^sub>V.src_def F\<^sub>U\<^sub>V.map\<^sub>0_def obj_a' map\<^sub>0_def by auto show "D\<^sub>V.trg e = b" using e 4 D\<^sub>V.trg_def by auto qed moreover have "D\<^sub>V.equivalence_map e" proof - obtain d \ \ where d: "equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D e d \ \" using e D.equivalence_map_def by auto interpret e: equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D e d \ \ using d by simp have "equivalence_in_bicategory D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\ \\<^sub>D D\<^sub>V.src D\<^sub>V.trg e d \ \" proof show ide_e: "D\<^sub>V.ide e" using e e_in_hhom D\<^sub>V.ide_char by auto show ide_d: "D\<^sub>V.ide d" using d ide_e e.antipar D\<^sub>V.ide_char D\<^sub>V.arr_char by simp show 4: "D\<^sub>V.in_hom \ (D\<^sub>V.src e) (D\<^sub>V.hcomp d e)" proof - have "D\<^sub>V.hseq d e" using ide_d ide_e e.antipar D\<^sub>V.src_def D\<^sub>V.trg_def by simp thus ?thesis using ide_d ide_e e.unit_in_hom D\<^sub>V.arr_char D\<^sub>V.ide_char D\<^sub>V.dom_char D\<^sub>V.cod_char D\<^sub>V.src_def D\<^sub>V.trg_def by (intro D\<^sub>V.in_homI) auto qed show 5: "D\<^sub>V.in_hom \ (D\<^sub>V.hcomp e d) (D\<^sub>V.src d)" proof - have "D\<^sub>V.hseq e d" using ide_d ide_e e.antipar D\<^sub>V.src_def D\<^sub>V.trg_def by simp thus ?thesis using ide_d ide_e e.antipar e.counit_in_hom D\<^sub>V.arr_char D\<^sub>V.ide_char D\<^sub>V.dom_char D\<^sub>V.cod_char D\<^sub>V.src_def D\<^sub>V.trg_def by (intro D\<^sub>V.in_homI) auto qed show "D\<^sub>V.iso \" using 4 D\<^sub>V.iso_char D\<^sub>V.arr_char by fastforce show "D\<^sub>V.iso \" using 5 D\<^sub>V.iso_char D\<^sub>V.arr_char by fastforce qed thus ?thesis using D\<^sub>V.equivalence_map_def by auto qed ultimately show ?thesis using D\<^sub>V.equivalent_objects_def by auto qed ultimately show "\a. C\<^sub>U.obj a \ D\<^sub>V.equivalent_objects (F\<^sub>U\<^sub>V.map\<^sub>0 a) b" by auto qed qed interpretation F\<^sub>U\<^sub>V: equivalence_pseudofunctor_bij_on_obj C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\ \\<^sub>C C\<^sub>U.src C\<^sub>U.trg D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\ \\<^sub>D D\<^sub>V.src D\<^sub>V.trg F\<^sub>U\<^sub>V.map F\<^sub>U\<^sub>V.cmp proof have "Collect C\<^sub>U.obj = U" using C\<^sub>U.obj_char C\<^sub>U.arr_char U_def by fastforce moreover have "Collect D\<^sub>V.obj = V" using D\<^sub>V.obj_char D\<^sub>V.arr_char V_def D.obj_def' map\<^sub>0_simps(1) by auto moreover have "\a. a \ Collect C\<^sub>U.obj \ F\<^sub>U\<^sub>V.map\<^sub>0 a = map\<^sub>0 a" using F\<^sub>U\<^sub>V.map\<^sub>0_def C\<^sub>U.obj_char C\<^sub>U.arr_char D\<^sub>V.src_def F\<^sub>U.map\<^sub>0_simp F\<^sub>U\<^sub>V.map\<^sub>0_simp by auto ultimately show "bij_betw F\<^sub>U\<^sub>V.map\<^sub>0 (Collect C\<^sub>U.obj) (Collect D\<^sub>V.obj)" using bij_betw_U_V by (simp add: bij_betw_U_V bij_betw_cong) qed interpretation EQ\<^sub>U\<^sub>V: equivalence_of_bicategories (* 25 sec *) D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\ \\<^sub>D D\<^sub>V.src D\<^sub>V.trg C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\ \\<^sub>C C\<^sub>U.src C\<^sub>U.trg F\<^sub>U\<^sub>V.map F\<^sub>U\<^sub>V.cmp F\<^sub>U\<^sub>V.G F\<^sub>U\<^sub>V.\\<^sub>G - F\<^sub>U\<^sub>V.unit\<^sub>0 F\<^sub>U\<^sub>V.unit\<^sub>1 F\<^sub>U\<^sub>V.counit\<^sub>0 F\<^sub>U\<^sub>V.counit\<^sub>1 + F\<^sub>U\<^sub>V.unit\<^sub>0' F\<^sub>U\<^sub>V.unit\<^sub>1' F\<^sub>U\<^sub>V.counit\<^sub>0' F\<^sub>U\<^sub>V.counit\<^sub>1' using F\<^sub>U\<^sub>V.extends_to_equivalence_of_bicategories by blast (* 18 sec, mostly "by" *) text \ Now compose \EQ\<^sub>U\<^sub>V\ with the equivalence from \D\<^sub>V\ to \D\ and the converse of the equivalence from \C\<^sub>U\ to \C\. The result is an equivalence of bicategories from \C\ to \D\. \ interpretation EQ\<^sub>C: equivalence_of_bicategories V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\ \\<^sub>C C\<^sub>U.src C\<^sub>U.trg C\<^sub>U.E C\<^sub>U.\\<^sub>E C\<^sub>U.P C\<^sub>U.\\<^sub>P C\<^sub>U.unit\<^sub>0 C\<^sub>U.unit\<^sub>1 C\<^sub>U.counit\<^sub>0 C\<^sub>U.counit\<^sub>1 using C\<^sub>U.induces_equivalence by simp interpretation EQ\<^sub>C': converse_equivalence_of_bicategories V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\ \\<^sub>C C\<^sub>U.src C\<^sub>U.trg C\<^sub>U.E C\<^sub>U.\\<^sub>E C\<^sub>U.P C\<^sub>U.\\<^sub>P C\<^sub>U.unit\<^sub>0 C\<^sub>U.unit\<^sub>1 C\<^sub>U.counit\<^sub>0 C\<^sub>U.counit\<^sub>1 .. interpretation EQ\<^sub>D: equivalence_of_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\ \\<^sub>D D\<^sub>V.src D\<^sub>V.trg D\<^sub>V.E D\<^sub>V.\\<^sub>E D\<^sub>V.P D\<^sub>V.\\<^sub>P D\<^sub>V.unit\<^sub>0 D\<^sub>V.unit\<^sub>1 D\<^sub>V.counit\<^sub>0 D\<^sub>V.counit\<^sub>1 using D\<^sub>V.induces_equivalence by simp - interpretation EQ\<^sub>U\<^sub>VoEQ\<^sub>C': composite_equivalence_of_bicategories (* 50 sec *) + interpretation EQ\<^sub>U\<^sub>VoEQ\<^sub>C': composite_equivalence_of_bicategories (* 35 sec *) D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\ \\<^sub>D D\<^sub>V.src D\<^sub>V.trg C\<^sub>U.comp C\<^sub>U.hcomp C\<^sub>U.\ \\<^sub>C C\<^sub>U.src C\<^sub>U.trg V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F\<^sub>U\<^sub>V.map F\<^sub>U\<^sub>V.cmp F\<^sub>U\<^sub>V.G F\<^sub>U\<^sub>V.\\<^sub>G C\<^sub>U.P C\<^sub>U.\\<^sub>P C\<^sub>U.E C\<^sub>U.\\<^sub>E - F\<^sub>U\<^sub>V.unit\<^sub>0 F\<^sub>U\<^sub>V.unit\<^sub>1 F\<^sub>U\<^sub>V.counit\<^sub>0 F\<^sub>U\<^sub>V.counit\<^sub>1 + F\<^sub>U\<^sub>V.unit\<^sub>0' F\<^sub>U\<^sub>V.unit\<^sub>1' F\<^sub>U\<^sub>V.counit\<^sub>0' F\<^sub>U\<^sub>V.counit\<^sub>1' EQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>C'.unit\<^sub>1 EQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>C'.counit\<^sub>1 - .. (* 40 sec *) - - interpretation EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C': composite_equivalence_of_bicategories (* 50 sec *) + .. (* 55 sec *) + + interpretation EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C': composite_equivalence_of_bicategories (* 30 sec *) V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D D\<^sub>V.comp D\<^sub>V.hcomp D\<^sub>V.\ \\<^sub>D D\<^sub>V.src D\<^sub>V.trg V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C D\<^sub>V.E D\<^sub>V.\\<^sub>E D\<^sub>V.P D\<^sub>V.\\<^sub>P EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp D\<^sub>V.unit\<^sub>0 D\<^sub>V.unit\<^sub>1 D\<^sub>V.counit\<^sub>0 D\<^sub>V.counit\<^sub>1 EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>1 EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>1 - .. (* 40 sec *) + .. (* 50 sec *) lemma induces_equivalence_of_bicategories: shows "equivalence_of_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>1 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>1" .. lemma left_map_simp: assumes "C.arr \" shows "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map \ = D\<^sub>V.E (F (C\<^sub>U.P \))" using assms by simp lemma right_map_simp: assumes "D.arr \" shows "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map \ = C\<^sub>U.E (F\<^sub>U\<^sub>V.G (D\<^sub>V.P \))" using assms by simp lemma unit\<^sub>0_simp: assumes "C.obj a" shows "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 a = - C\<^sub>U.E (F\<^sub>U\<^sub>V.G (D\<^sub>V.unit\<^sub>0 (D\<^sub>V.src (F (C\<^sub>U.P a))))) \\<^sub>C C\<^sub>U.E (C\<^sub>U.P\<^sub>0 (src\<^sub>C a)) + C\<^sub>U.E (F\<^sub>U\<^sub>V.G (D\<^sub>V.\\<^sub>0 (D\<^sub>V.src (F (C\<^sub>U.P a))))) \\<^sub>C C\<^sub>U.E (C\<^sub>U.P\<^sub>0 (src\<^sub>C a)) \\<^sub>C EQ\<^sub>C'.unit\<^sub>0 a" using assms EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0_simp EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0_simp EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0_def C\<^sub>U.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 \EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map\ and not the originally given equivalence pseudofunctor \F\. However, we can patch things up by showing that \EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.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 \EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map\, rather than \EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map\ and \EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map\. \ abbreviation \\<^sub>0 - where "\\<^sub>0 a \ F (C\<^sub>U.counit\<^sub>0 a)" + where "\\<^sub>0 a \ F (C\<^sub>U.\\<^sub>0 a)" abbreviation \\<^sub>1 - where "\\<^sub>1 f \ D.inv (\ (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)) \\<^sub>D F (C\<^sub>U.counit\<^sub>1 f) \\<^sub>D - \ (f, C\<^sub>U.counit\<^sub>0 (src\<^sub>C f))" + where "\\<^sub>1 f \ D.inv (\ (C\<^sub>U.\\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)) \\<^sub>D F (C\<^sub>U.\\<^sub>1 f) \\<^sub>D \ (f, C\<^sub>U.\\<^sub>0 (src\<^sub>C f))" lemma \\<^sub>0_in_hom [intro]: assumes "C.obj a" shows "\\\<^sub>0 a : map\<^sub>0 (C\<^sub>U.P\<^sub>0 a) \\<^sub>D map\<^sub>0 a\" and "\\\<^sub>0 a : \\<^sub>0 a \\<^sub>D \\<^sub>0 a\" proof - show "\\\<^sub>0 a : map\<^sub>0 (C\<^sub>U.P\<^sub>0 a) \\<^sub>D map\<^sub>0 a\" proof show "D.arr (\\<^sub>0 a)" using assms by simp show "src\<^sub>D (\\<^sub>0 a) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)" using assms map\<^sub>0_def C\<^sub>U.counit.ide_map\<^sub>0_obj C\<^sub>U.equivalence_data_simps\<^sub>B(7) C.ideD(1) preserves_src by presburger show "trg\<^sub>D (\\<^sub>0 a) = map\<^sub>0 a" using assms by auto qed show "\\\<^sub>0 a : \\<^sub>0 a \\<^sub>D \\<^sub>0 a\" using assms by auto qed lemma \\<^sub>0_simps [simp]: assumes "C.obj a" shows "D.ide (\\<^sub>0 a)" and "src\<^sub>D (\\<^sub>0 a) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)" and "trg\<^sub>D (\\<^sub>0 a) = map\<^sub>0 a" using assms \\<^sub>0_in_hom(1) [of a] by auto lemma \\<^sub>1_in_hom [intro]: assumes "C.ide f" shows "\\\<^sub>1 f : map\<^sub>0 (C\<^sub>U.P\<^sub>0 (src\<^sub>C f)) \\<^sub>D map\<^sub>0 (trg\<^sub>C f)\" and "\\\<^sub>1 f : F f \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D F (C\<^sub>U.P f)\" proof - show 1: "\\\<^sub>1 f : F f \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D F (C\<^sub>U.P f)\" proof (intro D.comp_in_homI) - show "\\ (f, C\<^sub>U.d (src\<^sub>C f)) : F f \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D F (f \\<^sub>C C\<^sub>U.counit\<^sub>0 (src\<^sub>C f))\" + show "\\ (f, C\<^sub>U.d (src\<^sub>C f)) : F f \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D F (f \\<^sub>C C\<^sub>U.\\<^sub>0 (src\<^sub>C f))\" using assms by auto - show "\F (C\<^sub>U.counit\<^sub>1 f) : - F (f \\<^sub>C C\<^sub>U.counit\<^sub>0 (src\<^sub>C f)) \\<^sub>D F (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C f) \\<^sub>C C\<^sub>U.P f)\" + show "\F (C\<^sub>U.\\<^sub>1 f) : F (f \\<^sub>C C\<^sub>U.\\<^sub>0 (src\<^sub>C f)) \\<^sub>D F (C\<^sub>U.\\<^sub>0 (trg\<^sub>C f) \\<^sub>C C\<^sub>U.P f)\" using assms C\<^sub>U.counit\<^sub>1_in_hom [of f] C\<^sub>U.P_def by auto - show "\D.inv (\ (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)) : - F (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C f) \\<^sub>C C\<^sub>U.P f) \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D F (C\<^sub>U.P f)\" + show "\D.inv (\ (C\<^sub>U.\\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)) : + F (C\<^sub>U.\\<^sub>0 (trg\<^sub>C f) \\<^sub>C C\<^sub>U.P f) \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D F (C\<^sub>U.P f)\" using assms C.VV.arr_char C.VV.ide_char \.components_are_iso by (metis (no_types, lifting) C\<^sub>U.P_def C\<^sub>U.counit.ide_map\<^sub>0_obj C\<^sub>U.counit\<^sub>1_simps(1,5) C\<^sub>U.equivalence_data_simps\<^sub>B(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 "\\\<^sub>1 f : map\<^sub>0 (C\<^sub>U.P\<^sub>0 (src\<^sub>C f)) \\<^sub>D map\<^sub>0 (trg\<^sub>C f)\" using assms 1 map\<^sub>0_def [of "C\<^sub>U.P\<^sub>0 (src\<^sub>C f)"] C\<^sub>U.emb.map\<^sub>0_simp C\<^sub>U.P\<^sub>0_props apply (intro D.in_hhomI) apply auto using D.src_dom [of "\\<^sub>1 f"] apply (elim D.in_homE) apply auto using D.trg_dom [of "\\<^sub>1 f"] apply (elim D.in_homE) by auto qed lemma \\<^sub>1_simps [simp]: assumes "C.ide f" shows "D.arr (\\<^sub>1 f)" and "src\<^sub>D (\\<^sub>1 f) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 (src\<^sub>C f))" and "trg\<^sub>D (\\<^sub>1 f) = map\<^sub>0 (trg\<^sub>C f)" and "D.dom (\\<^sub>1 f) = F f \\<^sub>D \\<^sub>0 (src\<^sub>C f)" and "D.cod (\\<^sub>1 f) = \\<^sub>0 (trg\<^sub>C f) \\<^sub>D F (C\<^sub>U.P f)" using assms \\<^sub>1_in_hom by blast+ lemma iso_\\<^sub>1: assumes "C.ide f" shows "D.iso (\\<^sub>1 f)" proof - - have "C.VV.ide (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)" - proof - - have "C.VV.arr (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)" + have "C.VV.ide (C\<^sub>U.\\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)" + proof - + have "C.VV.arr (C\<^sub>U.\\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)" using assms C\<^sub>U.equivalence_data_simps\<^sub>B(7) C.ideD(1) by auto - moreover have "C.VxV.ide (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)" + moreover have "C.VxV.ide (C\<^sub>U.\\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)" using assms C\<^sub>U.prj.preserves_ide [of f] C\<^sub>U.ide_char by simp ultimately show ?thesis - using assms C.VV.ide_char [of "(C\<^sub>U.counit\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)"] by blast + using assms C.VV.ide_char [of "(C\<^sub>U.\\<^sub>0 (trg\<^sub>C f), C\<^sub>U.P f)"] by blast qed hence "D.iso (\ (C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f))" by simp thus ?thesis using assms C.VV.ide_char C.VV.arr_char \.components_are_iso by (intro D.isos_compose) auto qed interpretation \: pseudonatural_equivalence V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp F \ \\<^sub>0 \\<^sub>1 proof show "\a. C.obj a \ D.ide (\\<^sub>0 a)" by simp show "\a. C.obj a \ D.equivalence_map (\\<^sub>0 a)" using C\<^sub>U.counit.components_are_equivalences preserves_equivalence_maps by blast show "\a. C.obj a \ \\\<^sub>0 a : src\<^sub>D (EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map a) \\<^sub>D src\<^sub>D (F a)\" proof - fix a assume a: "C.obj a" show "\\\<^sub>0 a : src\<^sub>D (EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map a) \\<^sub>D src\<^sub>D (F a)\" proof show "D.arr (\\<^sub>0 a)" using a by simp show "src\<^sub>D (\\<^sub>0 a) = src\<^sub>D (EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map a)" proof - have "D\<^sub>V.arr (F (C\<^sub>U.P a))" proof - have "src\<^sub>D (F (C\<^sub>U.P a)) \ V" proof - have "src\<^sub>D (F (C\<^sub>U.P a)) = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)" using a by auto moreover have "C.obj (C\<^sub>U.prj.map\<^sub>0 a)" using a C\<^sub>U.obj_char [of "C\<^sub>U.prj.map\<^sub>0 a"] C\<^sub>U.prj.map\<^sub>0_simps(1) by auto ultimately show ?thesis using V_def by simp qed moreover have "trg\<^sub>D (F (C\<^sub>U.P a)) \ V" proof - have "trg\<^sub>D (F (C\<^sub>U.P a)) = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)" using a by auto moreover have "C.obj (C\<^sub>U.prj.map\<^sub>0 a)" using a C\<^sub>U.obj_char [of "C\<^sub>U.prj.map\<^sub>0 a"] C\<^sub>U.prj.map\<^sub>0_simps(1) by auto ultimately show ?thesis using V_def by simp qed ultimately show ?thesis using a D\<^sub>V.arr_char by fastforce qed thus ?thesis using a D\<^sub>V.emb.map\<^sub>0_def D\<^sub>V.emb.map_def D\<^sub>V.src_def C.obj_simps(1-2) C\<^sub>U.P_simps\<^sub>B(2) by (simp add: C\<^sub>U.P\<^sub>0_props(1)) qed show "trg\<^sub>D (F (C\<^sub>U.d a)) = src\<^sub>D (F a)" using a by auto qed qed show "\f. C.ide f \ \\\<^sub>1 f : F f \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f\" proof - fix f assume f: "C.ide f" show "\\\<^sub>1 f : F f \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f\" proof - have "D\<^sub>V.arr (F (C\<^sub>U.P f))" using f F\<^sub>U\<^sub>V.preserves_arr by auto hence "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f = F (C\<^sub>U.P f)" using f EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0_def D\<^sub>V.emb.map_def by simp thus ?thesis using f by auto qed qed show "\f. C.ide f \ D.iso (\\<^sub>1 f)" using iso_\\<^sub>1 by simp show "\\. C.arr \ \ \\<^sub>1 (C.cod \) \\<^sub>D (F \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) = (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map \) \\<^sub>D \\<^sub>1 (C.dom \)" proof - fix \ assume \: "C.arr \" show "\\<^sub>1 (C.cod \) \\<^sub>D (F \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) = (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map \) \\<^sub>D \\<^sub>1 (C.dom \)" proof - have "\\<^sub>1 (C.cod \) \\<^sub>D (F \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) - = D.inv (\ (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.cod \))) \\<^sub>D - F (C\<^sub>U.counit\<^sub>1 (C.cod \)) \\<^sub>D - \ (C.cod \, C\<^sub>U.counit\<^sub>0 (src\<^sub>C \)) \\<^sub>D + = D.inv (\ (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.cod \))) \\<^sub>D + F (C\<^sub>U.\\<^sub>1 (C.cod \)) \\<^sub>D + \ (C.cod \, C\<^sub>U.\\<^sub>0 (src\<^sub>C \)) \\<^sub>D (F \ \\<^sub>D \\<^sub>0 (src\<^sub>C \))" using \ D.comp_assoc by simp - also have "... = D.inv (\ (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.cod \))) \\<^sub>D - (F (C\<^sub>U.counit\<^sub>1 (C.cod \)) \\<^sub>D - F (\ \\<^sub>C C\<^sub>U.counit\<^sub>0 (src\<^sub>C \))) \\<^sub>D - \ (C.dom \, C\<^sub>U.counit\<^sub>0 (src\<^sub>C \))" - using \ \.naturality [of "(\, C\<^sub>U.counit\<^sub>0 (src\<^sub>C \))"] D.comp_assoc + also have "... = D.inv (\ (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.cod \))) \\<^sub>D + (F (C\<^sub>U.\\<^sub>1 (C.cod \)) \\<^sub>D + F (\ \\<^sub>C C\<^sub>U.\\<^sub>0 (src\<^sub>C \))) \\<^sub>D + \ (C.dom \, C\<^sub>U.\\<^sub>0 (src\<^sub>C \))" + using \ \.naturality [of "(\, C\<^sub>U.\\<^sub>0 (src\<^sub>C \))"] D.comp_assoc C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def by simp - also have "... = (D.inv (\ (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.cod \))) \\<^sub>D - F (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \) \\<^sub>C C\<^sub>U.EoP.map \)) \\<^sub>D - F (C\<^sub>U.counit\<^sub>1 (C.dom \)) \\<^sub>D - \ (C.dom \, C\<^sub>U.counit\<^sub>0 (src\<^sub>C \))" + also have "... = (D.inv (\ (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.cod \))) \\<^sub>D + F (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \) \\<^sub>C C\<^sub>U.EoP.map \)) \\<^sub>D + F (C\<^sub>U.\\<^sub>1 (C.dom \)) \\<^sub>D + \ (C.dom \, C\<^sub>U.\\<^sub>0 (src\<^sub>C \))" proof - - have "F (C\<^sub>U.counit\<^sub>1 (C.cod \)) \\<^sub>D F (\ \\<^sub>C C\<^sub>U.counit\<^sub>0 (src\<^sub>C \)) - = F (C\<^sub>U.counit\<^sub>1 (C.cod \) \\<^sub>C (\ \\<^sub>C C\<^sub>U.counit\<^sub>0 (src\<^sub>C \)))" + have "F (C\<^sub>U.\\<^sub>1 (C.cod \)) \\<^sub>D F (\ \\<^sub>C C\<^sub>U.\\<^sub>0 (src\<^sub>C \)) + = F (C\<^sub>U.\\<^sub>1 (C.cod \) \\<^sub>C (\ \\<^sub>C C\<^sub>U.\\<^sub>0 (src\<^sub>C \)))" using \ by simp - also have "... = F ((C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \) \\<^sub>C C\<^sub>U.EoP.map \) \\<^sub>C - C\<^sub>U.counit\<^sub>1 (C.dom \))" + also have "... = F ((C\<^sub>U.\\<^sub>0 (trg\<^sub>C \) \\<^sub>C C\<^sub>U.EoP.map \) \\<^sub>C C\<^sub>U.\\<^sub>1 (C.dom \))" using \ C\<^sub>U.counit.naturality [of \] by simp - also have "... = F (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \) \\<^sub>C C\<^sub>U.EoP.map \) \\<^sub>D - F (C\<^sub>U.counit\<^sub>1 (C.dom \))" + also have "... = F (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \) \\<^sub>C C\<^sub>U.EoP.map \) \\<^sub>D + F (C\<^sub>U.\\<^sub>1 (C.dom \))" using \ by simp - finally have "F (C\<^sub>U.counit\<^sub>1 (C.cod \)) \\<^sub>D F (\ \\<^sub>C C\<^sub>U.counit\<^sub>0 (src\<^sub>C \)) - = F (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \) \\<^sub>C C\<^sub>U.EoP.map \) \\<^sub>D - F (C\<^sub>U.counit\<^sub>1 (C.dom \))" + finally have "F (C\<^sub>U.\\<^sub>1 (C.cod \)) \\<^sub>D F (\ \\<^sub>C C\<^sub>U.\\<^sub>0 (src\<^sub>C \)) + = F (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \) \\<^sub>C C\<^sub>U.EoP.map \) \\<^sub>D F (C\<^sub>U.\\<^sub>1 (C.dom \))" by blast thus ?thesis using D.comp_assoc by simp qed - also have "... = (F (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \)) \\<^sub>D F (C\<^sub>U.EoP.map \)) \\<^sub>D - D.inv (\ (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.dom \))) \\<^sub>D - F (C\<^sub>U.counit\<^sub>1 (C.dom \)) \\<^sub>D - \ (C.dom \, C\<^sub>U.counit\<^sub>0 (src\<^sub>C \))" + also have "... = (F (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \)) \\<^sub>D F (C\<^sub>U.EoP.map \)) \\<^sub>D + D.inv (\ (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.dom \))) \\<^sub>D + F (C\<^sub>U.\\<^sub>1 (C.dom \)) \\<^sub>D + \ (C.dom \, C\<^sub>U.\\<^sub>0 (src\<^sub>C \))" proof - - have "D.inv (\ (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.cod \))) \\<^sub>D - F (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \) \\<^sub>C C\<^sub>U.EoP.map \) - = (F (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \)) \\<^sub>D F (C\<^sub>U.EoP.map \)) \\<^sub>D - D.inv (\ (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.dom \)))" + have "D.inv (\ (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.cod \))) \\<^sub>D + F (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \) \\<^sub>C C\<^sub>U.EoP.map \) + = (F (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \)) \\<^sub>D F (C\<^sub>U.EoP.map \)) \\<^sub>D + D.inv (\ (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P (C.dom \)))" proof - - have "C.VV.arr (C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P \)" + have "C.VV.arr (C\<^sub>U.\\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P \)" proof (unfold C.VV.arr_char, intro conjI) show "C.arr (fst (C\<^sub>U.d (trg\<^sub>C \), C\<^sub>U.P \))" using \ by simp show "C.arr (snd (C\<^sub>U.d (trg\<^sub>C \), C\<^sub>U.P \))" using \ by simp show "src\<^sub>C (fst (C\<^sub>U.d (trg\<^sub>C \), C\<^sub>U.P \)) = trg\<^sub>C (snd (C\<^sub>U.d (trg\<^sub>C \), C\<^sub>U.P \))" using \ C\<^sub>U.emb.map\<^sub>0_def C\<^sub>U.emb.map_def apply simp using C\<^sub>U.P\<^sub>0_props(1) C\<^sub>U.P_simps\<^sub>B(3) C\<^sub>U.src_def by fastforce qed moreover have "C\<^sub>U.EoP.map \ = C\<^sub>U.P \" using \ by (simp add: C\<^sub>U.emb.map_def) moreover have "C\<^sub>U.emb.map\<^sub>0 (C\<^sub>U.P\<^sub>0 (trg\<^sub>C \)) = C\<^sub>U.P\<^sub>0 (trg\<^sub>C \)" using \ C\<^sub>U.P\<^sub>0_props(1) C\<^sub>U.emb.map\<^sub>0_simp by blast ultimately show ?thesis using \ C.VV.arr_char C.VV.dom_char C.VV.cod_char C.VV.ide_char FF_def \.inv_naturality [of "(C\<^sub>U.counit\<^sub>0 (trg\<^sub>C \), C\<^sub>U.P \)"] by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map \) \\<^sub>D \\<^sub>1 (C.dom \)" using \ F\<^sub>U\<^sub>V.preserves_arr D.comp_assoc D\<^sub>V.emb.map_def C\<^sub>U.emb.map_def by simp finally show "\\<^sub>1 (C.cod \) \\<^sub>D (F \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) = (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map \) \\<^sub>D \\<^sub>1 (C.dom \)" by blast qed qed show "\a. C.obj a \ (\\<^sub>0 a \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] = \\<^sub>1 a \\<^sub>D (unit a \\<^sub>D \\<^sub>0 a)" proof - fix a assume a: "C.obj a" have 1: "C\<^sub>U.obj (C\<^sub>U.prj.map\<^sub>0 a)" using a C\<^sub>U.prj.map\<^sub>0_simps(1) by simp have 2: "\C\<^sub>U.prj.unit a : C\<^sub>U.prj.map\<^sub>0 a \\<^sub>C C\<^sub>U.P a\" using a C\<^sub>U.in_hom_char by blast have 3: "\C\<^sub>U.prj.unit a : C\<^sub>U.prj.map\<^sub>0 a \\<^sub>C C\<^sub>U.prj.map\<^sub>0 a\" proof (intro C.in_hhomI) show 4: "C.arr (C\<^sub>U.prj.unit a)" using 2 by auto show "src\<^sub>C (C\<^sub>U.prj.unit a) = C\<^sub>U.prj.map\<^sub>0 a" proof - have "src\<^sub>C (C\<^sub>U.prj.unit a) = src\<^sub>C (C\<^sub>U.P a)" using 2 4 C.src_cod [of "C\<^sub>U.prj.unit a"] C.vconn_implies_hpar(1,3) by auto also have "... = C\<^sub>U.prj.map\<^sub>0 a" using a C.obj_simps(1) C\<^sub>U.prj.map\<^sub>0_def [of a] C\<^sub>U.src_def [of "C\<^sub>U.P a"] by simp finally show ?thesis by blast qed show "trg\<^sub>C (C\<^sub>U.prj.unit a) = C\<^sub>U.prj.map\<^sub>0 a" proof - have "trg\<^sub>C (C\<^sub>U.prj.unit a) = trg\<^sub>C (C\<^sub>U.P a)" using 2 4 C.trg_cod [of "C\<^sub>U.prj.unit a"] C.vconn_implies_hpar(2,4) by auto also have "... = C\<^sub>U.prj.map\<^sub>0 a" using a C.obj_simps(1,3) C\<^sub>U.prj.map\<^sub>0_def [of a] C\<^sub>U.trg_def [of "C\<^sub>U.P a"] by simp finally show ?thesis by blast qed qed have [simp]: "C.dom (C\<^sub>U.prj.unit a) = C\<^sub>U.prj.map\<^sub>0 a" using 2 by auto have [simp]: "C.arr (C\<^sub>U.prj.unit a)" using 2 by auto have [simp]: "C.cod (C\<^sub>U.prj.unit a) = C\<^sub>U.P a" using 2 by auto have [simp]: "src\<^sub>C (C\<^sub>U.prj.unit a) = C\<^sub>U.prj.map\<^sub>0 a" using 3 by auto have [simp]: "trg\<^sub>C (C\<^sub>U.prj.unit a) = C\<^sub>U.prj.map\<^sub>0 a" using 3 by auto have [simp]: "C.arr (C\<^sub>U.d a)" using a by simp have [simp]: "C.ide (C\<^sub>U.d a)" using a by simp have [simp]: "C.dom (C\<^sub>U.d a) = C\<^sub>U.d a" using a by auto have [simp]: "C.cod (C\<^sub>U.d a) = C\<^sub>U.d a" using a by auto have [simp]: "src\<^sub>C (C\<^sub>U.d a) = C\<^sub>U.prj.map\<^sub>0 a" using a C\<^sub>U.equivalence_data_simps\<^sub>B(7) by auto have [simp]: "trg\<^sub>C (C\<^sub>U.d a) = a" using a C\<^sub>U.equivalence_data_simps\<^sub>B(7) by auto have seq: "D.seq (F (C\<^sub>U.prj.unit a)) (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))" using a apply (intro D.seqI) using C\<^sub>U.prj.map\<^sub>0_simps(1) D\<^sub>V.arr_char F\<^sub>U\<^sub>V.unit_simps(1) apply blast using C\<^sub>U.prj.unit_simps(1) C\<^sub>U.arr_char apply blast proof - have "D.dom (F (C\<^sub>U.prj.unit a)) = F (C.dom (C\<^sub>U.prj.unit a))" using a C\<^sub>U.prj.unit_simps(1) C\<^sub>U.arr_char by simp also have "... = F (C\<^sub>U.prj.map\<^sub>0 a)" using a C\<^sub>U.prj.unit_simps [of a] C\<^sub>U.dom_char by simp also have "... = D.cod (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))" using a 1 C\<^sub>U.prj.unit_simps [of a] F\<^sub>U\<^sub>V.unit_simps(5) [of "C\<^sub>U.prj.map\<^sub>0 a"] D\<^sub>V.cod_char by simp finally show "D.dom (F (C\<^sub>U.prj.unit a)) = D.cod (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))" by blast qed have 4: "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit a = F (C\<^sub>U.prj.unit a) \\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)" proof - have "EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0 a = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)" using a EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0_def D\<^sub>V.src_def C.obj_def F\<^sub>U\<^sub>V.preserves_arr by auto have seq': "D\<^sub>V.arr (F (C\<^sub>U.prj.unit a) \\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))" proof - have 5: "D.arr (F (C\<^sub>U.prj.unit a) \\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))" using seq by simp moreover have "src\<^sub>D (F (C\<^sub>U.prj.unit a) \\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)) \ V" proof - have "src\<^sub>D (F (C\<^sub>U.prj.unit a) \\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)) = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)" proof - have "src\<^sub>D (F (C\<^sub>U.prj.unit a) \\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)) = src\<^sub>D (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))" using 5 D.src_vcomp D.vseq_implies_hpar by simp also have "... = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)" using a F\<^sub>U\<^sub>V.unit_simps(1-2) [of "C\<^sub>U.prj.map\<^sub>0 a"] D\<^sub>V.src_def F\<^sub>U\<^sub>V.map\<^sub>0_def apply simp using a C\<^sub>U.prj.map\<^sub>0_simps(1) F\<^sub>U\<^sub>V.preserves_arr map\<^sub>0_def by force finally show ?thesis by blast qed moreover have "C.obj (C\<^sub>U.prj.map\<^sub>0 a)" using a C\<^sub>U.prj.map\<^sub>0_simps(1) C\<^sub>U.obj_char by blast ultimately show ?thesis using V_def by blast qed moreover have "trg\<^sub>D (F (C\<^sub>U.prj.unit a) \\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)) \ V" proof - have "trg\<^sub>D (F (C\<^sub>U.prj.unit a) \\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)) = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)" proof - have "trg\<^sub>D (F (C\<^sub>U.prj.unit a) \\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)) = trg\<^sub>D (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))" using 5 D.trg_vcomp D.vseq_implies_hpar by simp also have "... = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)" using a F\<^sub>U\<^sub>V.unit_simps(1,3) [of "C\<^sub>U.prj.map\<^sub>0 a"] D\<^sub>V.src_def D\<^sub>V.trg_def F\<^sub>U\<^sub>V.map\<^sub>0_def apply simp using a C\<^sub>U.prj.map\<^sub>0_simps(1) F\<^sub>U\<^sub>V.preserves_arr map\<^sub>0_def by force finally show ?thesis by blast qed moreover have "C.obj (C\<^sub>U.prj.map\<^sub>0 a)" using a C\<^sub>U.prj.map\<^sub>0_simps(1) C\<^sub>U.obj_char by blast ultimately show ?thesis using V_def by blast qed ultimately show ?thesis using a D\<^sub>V.arr_char by simp qed have "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit a = F (C\<^sub>U.prj.unit a) \\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a) \\<^sub>D map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)" proof - have 5: "D\<^sub>V.obj (src\<^sub>D (F (C\<^sub>U.P\<^sub>0 a)))" using a C\<^sub>U.P\<^sub>0_props [of a] by (metis (no_types, lifting) EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0_simps(1) \EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0 a = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)\ map\<^sub>0_def) have 6: "D\<^sub>V.emb.unit (map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)" using a 5 D\<^sub>V.emb.unit_char' EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0_simps(1) \EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0 a = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)\ map\<^sub>0_def by simp moreover have "D\<^sub>V.emb.unit (map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)) = map\<^sub>0 (C\<^sub>U.prj.map\<^sub>0 a)" using 6 a C.obj_simps C\<^sub>U.equivalence_data_simps\<^sub>B(7) by simp moreover have "D\<^sub>V.arr (F (C\<^sub>U.prj.unit a) \\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))" using a C.obj_simps seq' by blast moreover have "D\<^sub>V.arr (F\<^sub>U\<^sub>V.unit (C\<^sub>U.P\<^sub>0 a))" using a 1 by simp moreover have "D\<^sub>V.arr (F (C\<^sub>U.prj.unit a))" using a F\<^sub>U\<^sub>V.preserves_arr by auto moreover have "D\<^sub>V.obj (F\<^sub>U\<^sub>V.map\<^sub>0 (C\<^sub>U.P\<^sub>0 a))" using a 1 F\<^sub>U\<^sub>V.map\<^sub>0_simps(1) by auto moreover have "F\<^sub>U\<^sub>V.map\<^sub>0 (C\<^sub>U.P\<^sub>0 a) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)" using a \EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.map\<^sub>0 a = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)\ map\<^sub>0_def F\<^sub>U\<^sub>V.map\<^sub>0_simps by auto ultimately show ?thesis using a seq EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit_char' [of a] EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit_char' [of a] D\<^sub>V.emb.map_def D\<^sub>V.seq_char D\<^sub>V.comp_char D.comp_assoc by simp qed also have "... = F (C\<^sub>U.prj.unit a) \\<^sub>D F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)" proof - have "D.dom (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a)) = map\<^sub>0 (C\<^sub>U.P\<^sub>0 a)" using a 1 F\<^sub>U\<^sub>V.unit_simps(4) [of "C\<^sub>U.prj.map\<^sub>0 a"] D\<^sub>V.dom_char F\<^sub>U\<^sub>V.map\<^sub>0_def D\<^sub>V.src_def map\<^sub>0_def C.obj_def F\<^sub>U\<^sub>V.preserves_arr by simp moreover have "D.arr (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))" using a \D.seq (F (C\<^sub>U.prj.unit a)) (F\<^sub>U\<^sub>V.unit (C\<^sub>U.prj.map\<^sub>0 a))\ by blast ultimately show ?thesis using a D.comp_arr_dom by auto qed also have "... = F (C\<^sub>U.prj.unit a) \\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)" using a 1 F\<^sub>U\<^sub>V.unit_char' F\<^sub>U.unit_char' by simp finally show ?thesis by blast qed have "(\\<^sub>0 a \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] = (\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a) \\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a]" using 4 by simp also have "(\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a) \\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] = (\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a) \\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)) \\<^sub>D (\\<^sub>0 a \\<^sub>D D.inv (unit (src\<^sub>C (C\<^sub>U.d a)))) \\<^sub>D D.inv (\ (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \\<^sub>D F \\<^sub>C[C\<^sub>U.d a] \\<^sub>D \ (trg\<^sub>C (C\<^sub>U.d a), C\<^sub>U.d a) \\<^sub>D (unit (trg\<^sub>C (C\<^sub>U.d a)) \\<^sub>D \\<^sub>0 a)" proof - have "\\<^sub>D[\\<^sub>0 a] = F \\<^sub>C[C\<^sub>U.d a] \\<^sub>D \ (trg\<^sub>C (C\<^sub>U.d a), C\<^sub>U.d a) \\<^sub>D (unit (trg\<^sub>C (C\<^sub>U.d a)) \\<^sub>D \\<^sub>0 a)" - using a preserves_lunit [of "C\<^sub>U.counit\<^sub>0 a"] C\<^sub>U.counit.ide_map\<^sub>0_obj lunit_coherence + using a preserves_lunit [of "C\<^sub>U.\\<^sub>0 a"] C\<^sub>U.counit.ide_map\<^sub>0_obj lunit_coherence by blast moreover have "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] = (\\<^sub>0 a \\<^sub>D D.inv (unit (src\<^sub>C (C\<^sub>U.d a)))) \\<^sub>D D.inv (\ (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a]" proof - have "F \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] = \ (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a)) \\<^sub>D (\\<^sub>0 a \\<^sub>D unit (src\<^sub>C (C\<^sub>U.d a))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]" - using preserves_runit(2) [of "C\<^sub>U.counit\<^sub>0 a"] by simp + using preserves_runit(2) [of "C\<^sub>U.\\<^sub>0 a"] by simp moreover have 1: "D.iso (\ (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a)))" using a C.VV.ide_char C.VV.arr_char C.obj_src [of "C\<^sub>U.d a"] \.components_are_iso [of "(C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))"] by auto ultimately have "D.inv (\ (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] = (\\<^sub>0 a \\<^sub>D unit (src\<^sub>C (C\<^sub>U.d a))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]" using D.invert_side_of_triangle(1) [of "F \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a]" "\ (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))" "(\\<^sub>0 a \\<^sub>D unit (src\<^sub>C (C\<^sub>U.d a))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]"] by fastforce thus ?thesis using D.invert_side_of_triangle(1) [of "D.inv (\ (C\<^sub>U.d a, src\<^sub>C (C\<^sub>U.d a))) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a]" "\\<^sub>0 a \\<^sub>D unit (src\<^sub>C (C\<^sub>U.d a))" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]"] using C.obj_src [of "C\<^sub>U.d a"] unit_char(2) by auto qed ultimately show ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a)) \\<^sub>D ((\\<^sub>0 a \\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)) \\<^sub>D (\\<^sub>0 a \\<^sub>D D.inv (unit (C\<^sub>U.prj.map\<^sub>0 a))))) \\<^sub>D D.inv (\ (C\<^sub>U.d a, C\<^sub>U.prj.map\<^sub>0 a)) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \\<^sub>D F \\<^sub>C[C\<^sub>U.d a] \\<^sub>D \ (a, C\<^sub>U.d a) \\<^sub>D (unit a \\<^sub>D \\<^sub>0 a)" proof - have "\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a) \\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a) = (\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a)) \\<^sub>D (\\<^sub>0 a \\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a))" using 1 seq D.whisker_left [of "\\<^sub>0 a"] by (simp add: F\<^sub>U.unit_char' F\<^sub>U\<^sub>V.unit_char') thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a)) \\<^sub>D D.inv (\ (C\<^sub>U.d a, C\<^sub>U.prj.map\<^sub>0 a))) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \\<^sub>D F \\<^sub>C[C\<^sub>U.d a] \\<^sub>D \ (a, C\<^sub>U.d a) \\<^sub>D (unit a \\<^sub>D \\<^sub>0 a)" proof - have "(\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a)) \\<^sub>D ((\\<^sub>0 a \\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)) \\<^sub>D (\\<^sub>0 a \\<^sub>D D.inv (unit (C\<^sub>U.prj.map\<^sub>0 a)))) = (\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a)) \\<^sub>D (\\<^sub>0 a \\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a) \\<^sub>D D.inv (unit (C\<^sub>U.prj.map\<^sub>0 a)))" proof - have "D.seq (unit (C\<^sub>U.prj.map\<^sub>0 a)) (D.inv (unit (C\<^sub>U.prj.map\<^sub>0 a)))" using a 1 unit_char(1-2) C\<^sub>U.obj_char by (intro D.seqI) auto thus ?thesis using a D.whisker_left [of "\\<^sub>0 a"] by simp qed also have "... = (\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a)) \\<^sub>D (\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.map\<^sub>0 a))" using a 1 unit_char(1-2) C\<^sub>U.obj_char unit_simps [of "C\<^sub>U.prj.map\<^sub>0 a"] by (simp add: D.comp_arr_inv') also have "... = \\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a) \\<^sub>D F (C\<^sub>U.prj.map\<^sub>0 a)" using seq D.whisker_left by auto also have "... = \\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a)" using D.comp_arr_dom by simp finally have "(\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a)) \\<^sub>D ((\\<^sub>0 a \\<^sub>D unit (C\<^sub>U.prj.map\<^sub>0 a)) \\<^sub>D (\\<^sub>0 a \\<^sub>D D.inv (unit (C\<^sub>U.prj.map\<^sub>0 a)))) = \\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = D.inv (\ (C\<^sub>U.d a, C\<^sub>U.P a)) \\<^sub>D (F (C\<^sub>U.d a \\<^sub>C C\<^sub>U.prj.unit a) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \\<^sub>D F \\<^sub>C[C\<^sub>U.d a]) \\<^sub>D \ (a, C\<^sub>U.d a) \\<^sub>D (unit a \\<^sub>D \\<^sub>0 a)" proof - have "(\\<^sub>0 a \\<^sub>D F (C\<^sub>U.prj.unit a)) \\<^sub>D D.inv (\ (C\<^sub>U.d a, C\<^sub>U.prj.map\<^sub>0 a)) = D.inv (\ (C\<^sub>U.d a, C\<^sub>U.P a)) \\<^sub>D F (C\<^sub>U.d a \\<^sub>C C\<^sub>U.prj.unit a)" proof - have "C.VV.arr (C\<^sub>U.d a, C\<^sub>U.prj.unit a)" using a C.VV.arr_char by simp thus ?thesis using a \.inv_naturality [of "(C\<^sub>U.d a, C\<^sub>U.prj.unit a)"] C\<^sub>U.prj.unit_simps [of a] C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = (D.inv (\ (C\<^sub>U.d a, C\<^sub>U.P a)) \\<^sub>D F (C\<^sub>U.counit\<^sub>1 a \\<^sub>C (a \\<^sub>C C\<^sub>U.d a)) \\<^sub>D \ (a, C\<^sub>U.d a)) \\<^sub>D (unit a \\<^sub>D \\<^sub>0 a)" proof - have "F (C\<^sub>U.d a \\<^sub>C C\<^sub>U.prj.unit a) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \\<^sub>D F \\<^sub>C[C\<^sub>U.d a] = F ((C\<^sub>U.d a \\<^sub>C C\<^sub>U.prj.unit a) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \\<^sub>C \\<^sub>C[C\<^sub>U.d a])" using a by simp also have "... = F (C\<^sub>U.counit\<^sub>1 a \\<^sub>C (a \\<^sub>C C\<^sub>U.d a))" proof - have "C\<^sub>U.EoP.unit a = C\<^sub>U.prj.unit a" using a 1 C\<^sub>U.emb.map_def C\<^sub>U.EoP.unit_char' C\<^sub>U.emb.unit_char' C.comp_arr_dom by simp thus ?thesis using a C\<^sub>U.counit.respects_unit [of a] C\<^sub>U.I\<^sub>C.unit_char' [of a] by simp qed finally have "F (C\<^sub>U.d a \\<^sub>C C\<^sub>U.prj.unit a) \\<^sub>D F \\<^sub>C\<^sup>-\<^sup>1[C\<^sub>U.d a] \\<^sub>D F \\<^sub>C[C\<^sub>U.d a] = F (C\<^sub>U.counit\<^sub>1 a \\<^sub>C (a \\<^sub>C C\<^sub>U.d a))" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>1 a \\<^sub>D (unit a \\<^sub>D \\<^sub>0 a)" proof - - have "D.inv (\ (C\<^sub>U.d a, C\<^sub>U.P a)) \\<^sub>D F (C\<^sub>U.counit\<^sub>1 a \\<^sub>C (a \\<^sub>C C\<^sub>U.d a)) \\<^sub>D - \ (a, C\<^sub>U.d a) - = D.inv (\ (C\<^sub>U.d a, C\<^sub>U.P a)) \\<^sub>D F (C\<^sub>U.counit\<^sub>1 a) \\<^sub>D \ (a, C\<^sub>U.d a)" + have "D.inv (\ (C\<^sub>U.d a, C\<^sub>U.P a)) \\<^sub>D F (C\<^sub>U.\\<^sub>1 a \\<^sub>C (a \\<^sub>C C\<^sub>U.d a)) \\<^sub>D \ (a, C\<^sub>U.d a) + = D.inv (\ (C\<^sub>U.d a, C\<^sub>U.P a)) \\<^sub>D F (C\<^sub>U.\\<^sub>1 a) \\<^sub>D \ (a, C\<^sub>U.d a)" using a C.obj_def' C.comp_arr_dom by (metis C\<^sub>U.counit\<^sub>1_simps(1) C\<^sub>U.counit\<^sub>1_simps(4) C.objE) also have "... = \\<^sub>1 a" using a by auto - finally have "D.inv (\ (C\<^sub>U.d a, C\<^sub>U.P a)) \\<^sub>D F (C\<^sub>U.counit\<^sub>1 a \\<^sub>C (a \\<^sub>C C\<^sub>U.d a)) \\<^sub>D - \ (a, C\<^sub>U.d a) + finally have "D.inv (\ (C\<^sub>U.d a, C\<^sub>U.P a)) \\<^sub>D F (C\<^sub>U.\\<^sub>1 a \\<^sub>C (a \\<^sub>C C\<^sub>U.d a)) \\<^sub>D \ (a, C\<^sub>U.d a) = \\<^sub>1 a" by blast thus ?thesis using D.comp_assoc by simp qed finally show "(\\<^sub>0 a \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] = \\<^sub>1 a \\<^sub>D (unit a \\<^sub>D \\<^sub>0 a)" by blast qed show "\f g. \C.ide f; C.ide g; src\<^sub>C g = trg\<^sub>C f\ \ (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map g, EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \\<^sub>D (\\<^sub>1 g \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0 (src\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \\<^sub>D (F g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)] = \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\ (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" proof - fix f g assume f: "C.ide f" and g: "C.ide g" and fg: "src\<^sub>C g = trg\<^sub>C f" have 1: "C.ide f \ C.ide g \ C.ide (C\<^sub>U.P f) \ C.ide (C\<^sub>U.P g) \ C.ide (C\<^sub>U.d (trg\<^sub>C f)) \ C.ide (C\<^sub>U.d (trg\<^sub>C g)) \ src\<^sub>C (C\<^sub>U.d (trg\<^sub>C f)) = trg\<^sub>C (C\<^sub>U.P f) \ src\<^sub>C (C\<^sub>U.d (trg\<^sub>C g)) = trg\<^sub>C (C\<^sub>U.P g) \ trg\<^sub>C (C\<^sub>U.d (trg\<^sub>C f)) = src\<^sub>C g" using f g fg C\<^sub>U.emb.map\<^sub>0_def C\<^sub>U.emb.map_def C\<^sub>U.obj_char C\<^sub>U.P\<^sub>0_props(1) C.obj_simps(2) C\<^sub>U.prj.preserves_ide C\<^sub>U.ide_char by auto have "\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\ (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = (D.inv (\ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f))) \\<^sub>D F ((C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.EoP.cmp (g, f)) \\<^sub>C \\<^sub>C[C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g, C\<^sub>U.P f] \\<^sub>C - (C\<^sub>U.counit\<^sub>1 g \\<^sub>C C\<^sub>U.P f) \\<^sub>C + (C\<^sub>U.\\<^sub>1 g \\<^sub>C C\<^sub>U.P f) \\<^sub>C C.inv \\<^sub>C[g, C\<^sub>U.d (src\<^sub>C g), C\<^sub>U.P f] \\<^sub>C - (g \\<^sub>C C\<^sub>U.counit\<^sub>1 f) \\<^sub>C + (g \\<^sub>C C\<^sub>U.\\<^sub>1 f) \\<^sub>C \\<^sub>C[g, f, C\<^sub>U.d (src\<^sub>C f)]) \\<^sub>D \ (g \\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \\<^sub>D (\ (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg C\<^sub>U.emb.map_def C\<^sub>U.counit.respects_hcomp [of f g] D.comp_arr_dom by simp also have "... = ((D.inv (\ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f))) \\<^sub>D \ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f)))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \\<^sub>D (\\<^sub>1 g \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \\<^sub>D (F g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (D.inv (\ (g \\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \\<^sub>D \ (g \\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \\<^sub>D (\ (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" proof - have "F ((C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.EoP.cmp (g, f)) \\<^sub>C \\<^sub>C[C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g, C\<^sub>U.P f] \\<^sub>C - (C\<^sub>U.counit\<^sub>1 g \\<^sub>C C\<^sub>U.P f) \\<^sub>C + (C\<^sub>U.\\<^sub>1 g \\<^sub>C C\<^sub>U.P f) \\<^sub>C C.inv \\<^sub>C[g, C\<^sub>U.d (src\<^sub>C g), C\<^sub>U.P f] \\<^sub>C - (g \\<^sub>C C\<^sub>U.counit\<^sub>1 f) \\<^sub>C + (g \\<^sub>C C\<^sub>U.\\<^sub>1 f) \\<^sub>C \\<^sub>C[g, f, C\<^sub>U.d (src\<^sub>C f)]) = F (C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.EoP.cmp (g, f)) \\<^sub>D F \\<^sub>C[C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g, C\<^sub>U.P f] \\<^sub>D - F (C\<^sub>U.counit\<^sub>1 g \\<^sub>C C\<^sub>U.P f) \\<^sub>D + F (C\<^sub>U.\\<^sub>1 g \\<^sub>C C\<^sub>U.P f) \\<^sub>D F (C.inv \\<^sub>C[g, C\<^sub>U.d (src\<^sub>C g), C\<^sub>U.P f]) \\<^sub>D - F (g \\<^sub>C C\<^sub>U.counit\<^sub>1 f) \\<^sub>D + F (g \\<^sub>C C\<^sub>U.\\<^sub>1 f) \\<^sub>D F \\<^sub>C[g, f, C\<^sub>U.d (src\<^sub>C f)]" proof - have "C.arr ((C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.EoP.cmp (g, f)) \\<^sub>C \\<^sub>C[C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g, C\<^sub>U.P f] \\<^sub>C - (C\<^sub>U.counit\<^sub>1 g \\<^sub>C C\<^sub>U.P f) \\<^sub>C + (C\<^sub>U.\\<^sub>1 g \\<^sub>C C\<^sub>U.P f) \\<^sub>C C.inv \\<^sub>C[g, C\<^sub>U.d (src\<^sub>C g), C\<^sub>U.P f] \\<^sub>C - (g \\<^sub>C C\<^sub>U.counit\<^sub>1 f) \\<^sub>C + (g \\<^sub>C C\<^sub>U.\\<^sub>1 f) \\<^sub>C \\<^sub>C[g, f, C\<^sub>U.d (src\<^sub>C f)])" using f g fg 1 C\<^sub>U.emb.map_def C.VV.arr_char C.VV.dom_char C\<^sub>U.EoP.FF_def by (intro C.seqI C.hseqI') auto thus ?thesis using f g fg by fastforce qed also have "... = \ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \\<^sub>D ((D.inv (\ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g \\<^sub>C C\<^sub>U.P f)) \\<^sub>D \ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g \\<^sub>C C\<^sub>U.P f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f))) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \\<^sub>D (D.inv (\ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g)) \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D ((D.inv (\ (C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D \ (C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D - (F (C\<^sub>U.counit\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f))) \\<^sub>D + (F (C\<^sub>U.\\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f))) \\<^sub>D ((D.inv (\ (g \\<^sub>C C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f)) \\<^sub>D \ (g \\<^sub>C C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f)) \\<^sub>D (\ (g, C\<^sub>U.d (trg\<^sub>C f)) \\<^sub>D F (C\<^sub>U.P f))) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \\<^sub>D (F g \\<^sub>D D.inv (\ (C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f))) \\<^sub>D ((D.inv (\ (g, C\<^sub>U.d (trg\<^sub>C f) \\<^sub>C C\<^sub>U.P f)) \\<^sub>D \ (g, C\<^sub>U.d (trg\<^sub>C f) \\<^sub>C C\<^sub>U.P f)) \\<^sub>D - (F g \\<^sub>D F (C\<^sub>U.counit\<^sub>1 f))) \\<^sub>D + (F g \\<^sub>D F (C\<^sub>U.\\<^sub>1 f))) \\<^sub>D ((D.inv (\ (g, f \\<^sub>C C\<^sub>U.d (src\<^sub>C f))) \\<^sub>D \ (g, f \\<^sub>C C\<^sub>U.d (src\<^sub>C f))) \\<^sub>D (F g \\<^sub>D \ (f, C\<^sub>U.d (src\<^sub>C f)))) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D D.inv (\ (g \\<^sub>C f, C\<^sub>U.d (src\<^sub>C f)))" using 1 f g fg preserves_hcomp preserves_assoc C\<^sub>U.emb.map_def C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def D.comp_assoc by simp also have "... = \ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \\<^sub>D ((D.inv (\ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g)) \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D - (F (C\<^sub>U.counit\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D + (F (C\<^sub>U.\\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D (\ (g, C\<^sub>U.d (trg\<^sub>C f)) \\<^sub>D F (C\<^sub>U.P f))) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \\<^sub>D ((F g \\<^sub>D D.inv (\ (C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f))) \\<^sub>D - (F g \\<^sub>D F (C\<^sub>U.counit\<^sub>1 f)) \\<^sub>D + (F g \\<^sub>D F (C\<^sub>U.\\<^sub>1 f)) \\<^sub>D (F g \\<^sub>D \ (f, C\<^sub>U.d (src\<^sub>C f)))) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D D.inv (\ (g \\<^sub>C f, C\<^sub>U.d (src\<^sub>C f)))" proof - have "(D.inv (\ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g \\<^sub>C C\<^sub>U.P f)) \\<^sub>D \ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g \\<^sub>C C\<^sub>U.P f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)) = \\<^sub>0 (trg\<^sub>C g) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)" using f g fg 1 C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr \.components_are_iso C.VV.ide_char FF_def by simp moreover have "(D.inv (\ (C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D \ (C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D - (F (C\<^sub>U.counit\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f)) - = F (C\<^sub>U.counit\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f)" + (F (C\<^sub>U.\\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f)) + = F (C\<^sub>U.\\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f)" proof - have "(D.inv (\ (C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D \ (C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D - (F (C\<^sub>U.counit\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f)) + (F (C\<^sub>U.\\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f)) = (F (C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.P g) \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D - (F (C\<^sub>U.counit\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f))" + (F (C\<^sub>U.\\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f))" using f g fg 1 C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr \.components_are_iso C.VV.ide_char FF_def by simp also have "... = F (C\<^sub>U.counit\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f)" using g D.comp_cod_arr f g fg C\<^sub>U.P\<^sub>0_props(1) C\<^sub>U.emb.map_def by simp finally show ?thesis by blast qed moreover have "(D.inv (\ (g \\<^sub>C C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f)) \\<^sub>D (\ (g \\<^sub>C C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f))) \\<^sub>D (\ (g, C\<^sub>U.d (trg\<^sub>C f)) \\<^sub>D F (C\<^sub>U.P f)) = (\ (g, C\<^sub>U.d (trg\<^sub>C f)) \\<^sub>D F (C\<^sub>U.P f))" using f g fg 1 C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr \.components_are_iso C.VV.ide_char FF_def by simp moreover have "(D.inv (\ (g, C\<^sub>U.d (trg\<^sub>C f) \\<^sub>C C\<^sub>U.P f)) \\<^sub>D \ (g, C\<^sub>U.d (trg\<^sub>C f) \\<^sub>C C\<^sub>U.P f)) \\<^sub>D (F g \\<^sub>D F (C\<^sub>U.counit\<^sub>1 f)) = F g \\<^sub>D F (C\<^sub>U.counit\<^sub>1 f)" using f g fg 1 C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr \.components_are_iso C.VV.ide_char FF_def by simp moreover have "(D.inv (\ (g, f \\<^sub>C C\<^sub>U.d (src\<^sub>C f))) \\<^sub>D \ (g, f \\<^sub>C C\<^sub>U.d (src\<^sub>C f))) \\<^sub>D (F g \\<^sub>D \ (f, C\<^sub>U.d (src\<^sub>C f))) = F g \\<^sub>D \ (f, C\<^sub>U.d (src\<^sub>C f))" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr \.components_are_iso C.VV.ide_char FF_def by simp ultimately show ?thesis using D.comp_assoc by simp qed also have "... = \ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \\<^sub>D (\\<^sub>1 g \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \\<^sub>D (F g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D D.inv (\ (g \\<^sub>C f, C\<^sub>U.d (src\<^sub>C f)))" proof - have "(F g \\<^sub>D D.inv (\ (C\<^sub>U.d (trg\<^sub>C f), C\<^sub>U.P f))) \\<^sub>D (F g \\<^sub>D F (C\<^sub>U.counit\<^sub>1 f)) \\<^sub>D (F g \\<^sub>D \ (f, C\<^sub>U.d (src\<^sub>C f))) = F g \\<^sub>D \\<^sub>1 f" using f g fg 1 D.whisker_left C.VV.arr_char C.VV.dom_char C.VV.cod_char by simp moreover have "(D.inv (\ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g)) \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D (F (C\<^sub>U.counit\<^sub>1 g) \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D (\ (g, C\<^sub>U.d (trg\<^sub>C f)) \\<^sub>D F (C\<^sub>U.P f)) = \\<^sub>1 g \\<^sub>D F (C\<^sub>U.P f)" using f g fg 1 D.whisker_right C.VV.arr_char C.VV.dom_char C.VV.cod_char C.VV.ide_char \.components_are_iso C\<^sub>U.emb.map_def by simp ultimately show ?thesis using D.comp_assoc by simp qed finally have "F ((C\<^sub>U.d (trg\<^sub>C g) \\<^sub>C C\<^sub>U.EoP.cmp (g, f)) \\<^sub>C \\<^sub>C[C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P g, C\<^sub>U.P f] \\<^sub>C - (C\<^sub>U.counit\<^sub>1 g \\<^sub>C C\<^sub>U.P f) \\<^sub>C + (C\<^sub>U.\\<^sub>1 g \\<^sub>C C\<^sub>U.P f) \\<^sub>C C.inv \\<^sub>C[g, C\<^sub>U.d (src\<^sub>C g), C\<^sub>U.P f] \\<^sub>C - (g \\<^sub>C C\<^sub>U.counit\<^sub>1 f) \\<^sub>C + (g \\<^sub>C C\<^sub>U.\\<^sub>1 f) \\<^sub>C \\<^sub>C[g, f, C\<^sub>U.d (src\<^sub>C f)]) = \ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \\<^sub>D (\\<^sub>1 g \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \\<^sub>D (F g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D D.inv (\ (g \\<^sub>C f, C\<^sub>U.d (src\<^sub>C f)))" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f))) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \\<^sub>D (\\<^sub>1 g \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \\<^sub>D (F g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)]" proof - have "(D.inv (\ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f))) \\<^sub>D \ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) = \\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))" proof - have "(D.inv (\ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f))) \\<^sub>D \ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.P (g \\<^sub>C f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f)))" proof - have "D.iso (\ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f)))" using f g fg \.components_are_iso C.VV.ide_char C.VV.arr_char by (metis (no_types, lifting) C\<^sub>U.prj.preserves_ide C\<^sub>U.P_simps\<^sub>B(3) C\<^sub>U.ide_char C\<^sub>U.counit.ide_map\<^sub>0_obj C\<^sub>U.equivalence_data_simps\<^sub>B(7) C.hcomp_simps(2) C.ideD(1) C.ide_hcomp C.obj_trg cmp_components_are_iso) moreover have "D.dom (\ (C\<^sub>U.d (trg\<^sub>C g), C\<^sub>U.P (g \\<^sub>C f))) = \\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.P (g \\<^sub>C f))" using f g fg C\<^sub>U.ide_char C\<^sub>U.equivalence_data_simps\<^sub>B(7) C\<^sub>U.prj.preserves_ide cmp_simps(4) by simp ultimately show ?thesis using D.comp_inv_arr' by auto qed also have "... = \\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.emb.map_def D.whisker_left [of "\\<^sub>0 (trg\<^sub>C g)"] by simp also have "... = \\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.emb.map_def D.comp_cod_arr by simp finally show ?thesis by blast qed moreover have "\\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (D.inv (\ (g \\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \\<^sub>D \ (g \\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \\<^sub>D (\ (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)]" proof - have "D.inv (\ (g \\<^sub>C f, C\<^sub>U.d (src\<^sub>C f))) \\<^sub>D \ (g \\<^sub>C f, C\<^sub>U.d (src\<^sub>C f)) = F (g \\<^sub>C f) \\<^sub>D F (C\<^sub>U.d (src\<^sub>C f))" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr \.components_are_iso C.VV.ide_char FF_def by simp moreover have "(F (g \\<^sub>C f) \\<^sub>D F (C\<^sub>U.d (src\<^sub>C f))) \\<^sub>D (\ (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = \ (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr \.components_are_iso C.VV.ide_char FF_def by simp moreover have "(D.inv (\ (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\ (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = (F g \\<^sub>D F f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.emb.map_def D.comp_inv_arr' D.comp_arr_dom D.comp_cod_arr \.components_are_iso C.VV.ide_char FF_def D.whisker_right [of "\\<^sub>0 (src\<^sub>C f)" "D.inv (\ (g, f))" "\ (g, f)"] by simp moreover have "\\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D ((F g \\<^sub>D F f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)]" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.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 "... = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F (C\<^sub>U.P g), F (C\<^sub>U.P f)] \\<^sub>D (\\<^sub>1 g \\<^sub>D F (C\<^sub>U.P f)) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0 (trg\<^sub>C f), F (C\<^sub>U.P f)] \\<^sub>D (F g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)]" proof - have "(\\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)) = \\<^sub>0 (trg\<^sub>C g) \\<^sub>D F (C\<^sub>U.EoP.cmp (g, f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)" proof - have "D.arr (F (C\<^sub>U.EoP.cmp (g, f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f))" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char C\<^sub>U.EoP.FF_def C\<^sub>U.emb.map_def by (intro D.seqI) auto thus ?thesis using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char D.whisker_left [of "\\<^sub>0 (trg\<^sub>C g)"] by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map g, EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \\<^sub>D (\\<^sub>1 g \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0 (src\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \\<^sub>D (F g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)]" proof - have "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp (g, f) = F (C\<^sub>U.EoP.cmp (g, f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)" proof - have "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp (g, f) = F (C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D ((F (C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D F (C\<^sub>U.\\<^sub>P (g, f))) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D D\<^sub>V.\\<^sub>E (F (C\<^sub>U.P g), F (C\<^sub>U.P f))" proof - have 3: "D.arr (F (C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D F (C\<^sub>U.\\<^sub>P (g, f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f))" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def C\<^sub>U.prj.FF_def C\<^sub>U.\\<^sub>P_in_hom [of g f] C\<^sub>U.hcomp_def C\<^sub>U.in_hom_char by (intro D.seqI) auto have 4: "D\<^sub>V.arr (\ (C\<^sub>U.P g, C\<^sub>U.P f))" using f g fg 1 cmp_in_hom(1) [of "C\<^sub>U.P g" "C\<^sub>U.P f"] C\<^sub>U.prj.map\<^sub>0_simps(1) C\<^sub>U.obj_char D\<^sub>V.arr_char V_def by auto moreover have 5: "D\<^sub>V.arr (F (C\<^sub>U.\\<^sub>P (g, f)))" using f g fg 1 C\<^sub>U.\\<^sub>P_simps [of g f] F\<^sub>U\<^sub>V.preserves_arr by presburger moreover have 6: "D\<^sub>V.arr (F (C\<^sub>U.P (g \\<^sub>C f)))" using f g fg C\<^sub>U.P_simps(1) [of "g \\<^sub>C f"] F\<^sub>U\<^sub>V.preserves_arr by simp moreover have 7: "D\<^sub>V.arr (F (C\<^sub>U.\\<^sub>P (g, f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f))" using 3 4 5 D\<^sub>V.seq_char [of "F (C\<^sub>U.\\<^sub>P (g, f))" "\ (C\<^sub>U.P g, C\<^sub>U.P f)"] D\<^sub>V.comp_char [of "F (C\<^sub>U.\\<^sub>P (g, f))" "\ (C\<^sub>U.P g, C\<^sub>U.P f)"] by auto moreover have "D.seq (F (C\<^sub>U.\\<^sub>P (g, f))) (\ (C\<^sub>U.P g, C\<^sub>U.P f))" using 3 by blast moreover have "D.seq (F (C\<^sub>U.P (g \\<^sub>C f))) (F (C\<^sub>U.\\<^sub>P (g, f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f))" using 3 by blast moreover have "D\<^sub>V.arr (F (C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D F (C\<^sub>U.\\<^sub>P (g, f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f))" using f g fg 3 6 7 D.vseq_implies_hpar D\<^sub>V.arr_char V_def by (metis (no_types, lifting) D\<^sub>V.comp_char D\<^sub>V.seq_char) ultimately show ?thesis using f g fg EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.cmp_def EQ\<^sub>U\<^sub>VoEQ\<^sub>C'.FH.cmp_def C.VV.arr_char C.VV.dom_char C\<^sub>U.VV.arr_char D\<^sub>V.comp_char D\<^sub>V.emb.map_def D.comp_assoc by simp qed also have "... = F (C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D F (C\<^sub>U.\\<^sub>P (g, f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f) \\<^sub>D D\<^sub>V.\\<^sub>E (F (C\<^sub>U.P g), F (C\<^sub>U.P f))" proof - have "F (C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D F (C\<^sub>U.\\<^sub>P (g, f)) = F (C\<^sub>U.\\<^sub>P (g, f))" proof - have 1: "C.arr (C\<^sub>U.\\<^sub>P (g, f))" using f g fg C.VV.arr_char by (metis (no_types, lifting) C\<^sub>U.prj.preserves_hcomp C\<^sub>U.prj.preserves_ide C\<^sub>U.ide_char C\<^sub>U.seq_char C.ideD(1) C.ideD(3) C.ide_hcomp C.seqE) moreover have "C.cod (C\<^sub>U.\\<^sub>P (g, f)) = C\<^sub>U.P (g \\<^sub>C f)" using f g fg C.VV.arr_char C\<^sub>U.\\<^sub>P_in_hom(2) [of g f] C\<^sub>U.hcomp_def C\<^sub>U.in_hom_char by auto ultimately have "D.cod (F (C\<^sub>U.\\<^sub>P (g, f))) = F (C\<^sub>U.P (g \\<^sub>C f))" using f g fg C.VV.arr_char preserves_cod [of "C\<^sub>U.\\<^sub>P (g, f)"] by simp thus ?thesis using 1 f g fg D.comp_cod_arr [of "F (C\<^sub>U.\\<^sub>P (g, f))" "F (C\<^sub>U.P (g \\<^sub>C f))"] by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = F (C\<^sub>U.P (g \\<^sub>C f)) \\<^sub>D F (C\<^sub>U.\\<^sub>P (g, f)) \\<^sub>D F (C\<^sub>U.\\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)" proof - have "D\<^sub>V.\\<^sub>E (F (C\<^sub>U.P g), F (C\<^sub>U.P f)) = F (C\<^sub>U.P g) \\<^sub>D F (C\<^sub>U.P f)" using f g fg D\<^sub>V.emb.cmp_def D\<^sub>V.VV.arr_char D\<^sub>V.src_def D\<^sub>V.trg_def F\<^sub>U\<^sub>V.preserves_arr by auto moreover have "C\<^sub>U.\\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f) = C\<^sub>U.P g \\<^sub>C C\<^sub>U.P f" using f g fg C\<^sub>U.VV.arr_char C\<^sub>U.emb.cmp_def by simp ultimately have "\ (C\<^sub>U.P g, C\<^sub>U.P f) \\<^sub>D D\<^sub>V.\\<^sub>E (F (C\<^sub>U.P g), F (C\<^sub>U.P f)) = F (C\<^sub>U.\\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def \.naturality [of "(C\<^sub>U.P g, C\<^sub>U.P f)"] by simp thus ?thesis using D.comp_assoc by simp qed also have "... = F (C\<^sub>U.P (g \\<^sub>C f) \\<^sub>C C\<^sub>U.\\<^sub>P (g, f) \\<^sub>C C\<^sub>U.\\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)" proof - have "C.arr (C\<^sub>U.P (g \\<^sub>C f) \\<^sub>C C\<^sub>U.\\<^sub>P (g, f) \\<^sub>C C\<^sub>U.\\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))" proof (intro C.seqI) show "C.arr (C\<^sub>U.\\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))" using f g fg by auto show 1: "C.arr (C\<^sub>U.\\<^sub>P (g, f))" using f g fg C\<^sub>U.arr_char C\<^sub>U.\\<^sub>P_simps(1) [of g f] by auto show "C.dom (C\<^sub>U.\\<^sub>P (g, f)) = C.cod (C\<^sub>U.\\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))" proof - have "C.dom (C\<^sub>U.\\<^sub>P (g, f)) = C\<^sub>U.hcomp (C\<^sub>U.P g) (C\<^sub>U.P f)" using f g fg C\<^sub>U.VV.arr_char C\<^sub>U.\\<^sub>P_simps(4) [of g f] C\<^sub>U.dom_char [of "C\<^sub>U.\\<^sub>P (g, f)"] by fastforce also have "... = C\<^sub>U.P g \\<^sub>C (C\<^sub>U.P f)" using f g fg by auto also have "... = C.cod (C\<^sub>U.\\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))" using f g fg C\<^sub>U.emb.cmp_def [of "(C\<^sub>U.P g, C\<^sub>U.P f)"] by auto finally show ?thesis by blast qed show "C.arr (C\<^sub>U.P (g \\<^sub>C f))" using f g fg by simp show "C.dom (C\<^sub>U.P (g \\<^sub>C f)) = C.cod (C\<^sub>U.\\<^sub>P (g, f) \\<^sub>C C\<^sub>U.\\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))" proof - have "C.dom (C\<^sub>U.P (g \\<^sub>C f)) = C\<^sub>U.P (g \\<^sub>C f)" using f g fg by simp also have "... = C.cod (C\<^sub>U.\\<^sub>P (g, f))" using f g fg C\<^sub>U.\\<^sub>P_simps(5) [of g f] C\<^sub>U.cod_char [of "C\<^sub>U.\\<^sub>P (g, f)"] by fastforce also have "... = C.cod (C\<^sub>U.\\<^sub>P (g, f) \\<^sub>C C\<^sub>U.\\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))" proof - have 2: "C\<^sub>U.hcomp (C\<^sub>U.P g) (C\<^sub>U.P f) = C\<^sub>U.P g \\<^sub>C C\<^sub>U.P f" using f g fg by auto have "C\<^sub>U.arr (C\<^sub>U.P g \\<^sub>C C\<^sub>U.P f)" proof - have "C\<^sub>U.arr (C\<^sub>U.hcomp (C\<^sub>U.P g) (C\<^sub>U.P f))" using f g fg by simp thus ?thesis using 2 by simp qed moreover have "C.dom (C\<^sub>U.\\<^sub>P (g, f)) = C\<^sub>U.P g \\<^sub>C C\<^sub>U.P f" using f g fg 2 C\<^sub>U.VV.arr_char C\<^sub>U.\\<^sub>P_simps(4) [of g f] C\<^sub>U.dom_char [of "C\<^sub>U.\\<^sub>P (g, f)"] by fastforce ultimately have "C.arr (C\<^sub>U.\\<^sub>P (g, f) \\<^sub>C C\<^sub>U.\\<^sub>E (C\<^sub>U.P g, C\<^sub>U.P f))" using f g fg 1 C\<^sub>U.VV.arr_char C\<^sub>U.VV.cod_char C\<^sub>U.hcomp_char C\<^sub>U.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 (C\<^sub>U.EoP.cmp (g, f)) \\<^sub>D \ (C\<^sub>U.P g, C\<^sub>U.P f)" using f g fg C\<^sub>U.EoP.cmp_def C.VV.arr_char C.VV.dom_char C\<^sub>U.emb.map_def by simp finally show ?thesis by blast qed moreover have "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map g = F (C\<^sub>U.P g)" proof - have "D\<^sub>V.arr (F (C\<^sub>U.P g))" using g D\<^sub>V.arr_char C\<^sub>U.P_simps(1) C.ideD(1) F\<^sub>U\<^sub>V.preserves_arr by presburger thus ?thesis using g D\<^sub>V.emb.map_def by simp qed moreover have "EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f = F (C\<^sub>U.P f)" proof - have "D\<^sub>V.arr (F (C\<^sub>U.P f))" using f D\<^sub>V.arr_char C\<^sub>U.P_simps(1) C.ideD(1) F\<^sub>U\<^sub>V.preserves_arr by presburger thus ?thesis using f D\<^sub>V.emb.map_def by simp qed ultimately show ?thesis using fg by argo qed finally show "(\\<^sub>0 (trg\<^sub>C g) \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map g, EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \\<^sub>D (\\<^sub>1 g \\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0 (src\<^sub>C g), EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map f] \\<^sub>D (F g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0 (src\<^sub>C f)] = \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\ (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" by argo qed qed - interpretation EQ: equivalence_of_bicategories_and_pseudonatural_equivalence_left (* 25 sec *) + interpretation EQ: equivalence_of_bicategories_and_pseudonatural_equivalence_left (* 17 sec *) V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>1 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>1 F \ \\<^sub>0 \\<^sub>1 proof - interpret E: equivalence_of_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>1 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>1 using induces_equivalence_of_bicategories by blast interpret \: pseudonatural_equivalence V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp F \ \\<^sub>0 \\<^sub>1 .. show "equivalence_of_bicategories_and_pseudonatural_equivalence_left V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.left_cmp EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.unit\<^sub>1 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>0 EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.counit\<^sub>1 F \ \\<^sub>0 \\<^sub>1" .. qed definition right_map where "right_map \ EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_map" definition right_cmp where "right_cmp \ EQ\<^sub>DoEQ\<^sub>U\<^sub>VoEQ\<^sub>C'.right_cmp" definition unit\<^sub>0 where "unit\<^sub>0 \ EQ.unit.map\<^sub>0" definition unit\<^sub>1 where "unit\<^sub>1 \ EQ.unit.map\<^sub>1" definition counit\<^sub>0 where "counit\<^sub>0 \ EQ.counit.map\<^sub>0" definition counit\<^sub>1 where "counit\<^sub>1 \ EQ.counit.map\<^sub>1" theorem extends_to_equivalence_of_bicategories: shows "equivalence_of_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \ right_map right_cmp unit\<^sub>0 unit\<^sub>1 counit\<^sub>0 counit\<^sub>1" unfolding right_map_def right_cmp_def unit\<^sub>0_def unit\<^sub>1_def counit\<^sub>0_def counit\<^sub>1_def .. end locale converse_equivalence_pseudofunctor = (* 10 sec *) C: bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + D: bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + F: equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \\<^sub>F for V\<^sub>C :: "'c comp" (infixr "\\<^sub>C" 55) and H\<^sub>C :: "'c comp" (infixr "\\<^sub>C" 53) and \\<^sub>C :: "'c \ 'c \ 'c \ 'c" ("\\<^sub>C[_, _, _]") and \\<^sub>C :: "'c \ 'c" ("\\<^sub>C[_]") and src\<^sub>C :: "'c \ 'c" and trg\<^sub>C :: "'c \ 'c" and V\<^sub>D :: "'d comp" (infixr "\\<^sub>D" 55) and H\<^sub>D :: "'d comp" (infixr "\\<^sub>D" 53) and \\<^sub>D :: "'d \ 'd \ 'd \ 'd" ("\\<^sub>D[_, _, _]") and \\<^sub>D :: "'d \ 'd" ("\\<^sub>D[_]") and src\<^sub>D :: "'d \ 'd" and trg\<^sub>D :: "'d \ 'd" and F :: "'c \ 'd" and \\<^sub>F :: "'c * 'c \ 'd" begin interpretation E: equivalence_of_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \\<^sub>F F.right_map F.right_cmp F.unit\<^sub>0 F.unit\<^sub>1 F.counit\<^sub>0 F.counit\<^sub>1 using F.extends_to_equivalence_of_bicategories by simp interpretation E': converse_equivalence_of_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \\<^sub>F F.right_map F.right_cmp F.unit\<^sub>0 F.unit\<^sub>1 F.counit\<^sub>0 F.counit\<^sub>1 .. sublocale equivalence_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F.right_map F.right_cmp using E'.equivalence_pseudofunctor_left by simp lemma is_equivalence_pseudofunctor: shows "equivalence_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C 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 V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C \ \F \. equivalence_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \" lemma equivalent_bicategories_reflexive: assumes "bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C" shows "equivalent_bicategories V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C" proof - interpret bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C using assms by simp interpret I: identity_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C .. show ?thesis using I.equivalence_pseudofunctor_axioms equivalent_bicategories_def by blast qed lemma equivalent_bicategories_symmetric: assumes "equivalent_bicategories V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D" shows "equivalent_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C" proof - obtain F \\<^sub>F where F: "equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \\<^sub>F" using assms equivalent_bicategories_def by blast interpret F: equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \\<^sub>F using F by simp interpret G: converse_equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \\<^sub>F .. show ?thesis using G.is_equivalence_pseudofunctor equivalent_bicategories_def by blast qed lemma equivalent_bicategories_transitive: assumes "equivalent_bicategories V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C" and "equivalent_bicategories V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D" shows "equivalent_bicategories V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D" proof - obtain F \\<^sub>F where F: "equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \\<^sub>F" using assms(1) equivalent_bicategories_def by blast obtain G \\<^sub>G where G: "equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D G \\<^sub>G" using assms(2) equivalent_bicategories_def by blast interpret F: equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \\<^sub>F using F by simp interpret G: equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D G \\<^sub>G using G by simp interpret GF: composite_equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \\<^sub>F G \\<^sub>G .. show "equivalent_bicategories V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D" using equivalent_bicategories_def GF.equivalence_pseudofunctor_axioms by blast qed end