diff --git a/thys/Bicategory/PseudonaturalTransformation.thy b/thys/Bicategory/PseudonaturalTransformation.thy --- a/thys/Bicategory/PseudonaturalTransformation.thy +++ b/thys/Bicategory/PseudonaturalTransformation.thy @@ -1,8201 +1,8201 @@ (* Title: PseudonaturalTransformation Author: Eugene W. Stark , 2020 Maintainer: Eugene W. Stark *) section "Pseudonatural Transformations" theory PseudonaturalTransformation -imports Bicategory.InternalEquivalence Bicategory.InternalAdjunction Bicategory.Pseudofunctor +imports InternalEquivalence InternalAdjunction Pseudofunctor begin subsection "Definition of Pseudonatural Transformation" text \ Pseudonatural transformations are a generalization of natural transformations that is appropriate for pseudofunctors. The ``components'' of a pseudonatural transformation \\\ from a pseudofunctor \F\ to a pseudofunctor \G\ (both from bicategory \C\ to \D\), are 1-cells \\\\<^sub>0 a : F\<^sub>0 a \\<^sub>D G\<^sub>0 a\\ associated with the objects of \C\. Instead of ``naturality squares'' that commute on-the-nose, a pseudonatural transformation associates an invertible 2-cell \\\\<^sub>1 f : G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0 a' \\<^sub>D F f\\ with each 1-cell \\f : a \\<^sub>C a'\\ of \C\. The 1-cells \\\<^sub>0 a\ and 2-cells \\\<^sub>1 f\ are subject to coherence conditions which express that they transform naturally across 2-cells of \C\ and behave sensibly with respect to objects and horizontal composition. In formalizing ordinary natural transformations, we found it convenient to treat them similarly to functors in that a natural transformation \\ : C \ D\ maps arrows of \C\ to arrows of \D\; the components \\ a\ at objects \a\ being merely special cases. However, it is not possible to take the same approach for pseudofunctors, because it is not possible to treat the components \\\<^sub>0 a\ at objects \a\ as a special case of the components \\\<^sub>1 f\ at 1-cells \f\. So we have to regard a pseudonatural transformation \\\ as consisting of two separate mappings: a mapping \\\<^sub>0\ from objects of \C\ to 1-cells of \D\ and a mapping \\\<^sub>1\ from 1-cells of \C\ to invertible 2-cells of \D\. Pseudonatural transformations are themselves a special case of the more general notion of ``lax natural transformations'' between pseudofunctors. For a lax natural transformation \\\, the 2-cells \\\<^sub>1 f\ are not required to be invertible. This means that there is a distinction between a lax natural transformation with \\\\<^sub>1 f : G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0 a' \\<^sub>D F f\\ and an ``op-lax'' natural transformation with \\\\<^sub>1 f : \\<^sub>0 a' \\<^sub>D F f \\<^sub>D G f \\<^sub>D \\<^sub>0 a\\. There is some variation in the literature on which direction is considered ``lax'' and which is ``op-lax'' and this variation extends as well to the special case of pseudofunctors, though in that case it does not result in any essential difference in the notion, due to the assumed invertibility. We have chosen the direction that agrees with Leinster \cite{leinster-basic-bicategories}, and with the ``nLab'' article \cite{nlab-lax-natural-transformation} on lax natural transformations, but note that the ``nLab'' article \cite{nlab-pseudonatural-transformation} on pseudonatural transformations seems to make the opposite choice. \ locale pseudonatural_transformation = 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>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: 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 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" and G :: "'c \ 'd" and \\<^sub>G :: "'c * 'c \ 'd" and \\<^sub>0 :: "'c \ 'd" and \\<^sub>1 :: "'c \ 'd" + assumes map\<^sub>0_in_hhom: "C.obj a \ \\\<^sub>0 a : src\<^sub>D (F a) \\<^sub>D src\<^sub>D (G a)\" and map\<^sub>1_in_vhom: "C.ide f \ \\\<^sub>1 f : G f \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D F f\" and ide_map\<^sub>0_obj: "C.obj a \ D.ide (\\<^sub>0 a)" and iso_map\<^sub>1_ide: "C.ide f \ D.iso (\\<^sub>1 f)" and naturality: "C.arr \ \ \\<^sub>1 (C.cod \) \\<^sub>D (G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) = (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D F \) \\<^sub>D \\<^sub>1 (C.dom \)" and respects_unit: "C.obj a \ (\\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] = \\<^sub>1 a \\<^sub>D (G.unit a \\<^sub>D \\<^sub>0 a)" and respects_hcomp: "\ C.ide f; C.ide g; src\<^sub>C g = trg\<^sub>C f \ \ (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>1 g \\<^sub>D F f) \\<^sub>D D.inv \\<^sub>D[G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 (src\<^sub>C f)] = \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" begin lemma map\<^sub>0_in_hom [intro]: assumes "C.obj a" shows "\\\<^sub>0 a : F.map\<^sub>0 a \\<^sub>D G.map\<^sub>0 a\" and "\\\<^sub>0 a : \\<^sub>0 a \\<^sub>D \\<^sub>0 a\" using assms map\<^sub>0_in_hhom [of a] apply fastforce using assms ide_map\<^sub>0_obj [of a] by fast lemma map\<^sub>0_simps [simp]: assumes "C.obj a" shows "D.ide (\\<^sub>0 a)" and "src\<^sub>D (\\<^sub>0 a) = F.map\<^sub>0 a" and "trg\<^sub>D (\\<^sub>0 a) = G.map\<^sub>0 a" using assms map\<^sub>0_in_hom iso_map\<^sub>1_ide ide_map\<^sub>0_obj by auto lemma map\<^sub>1_in_hom [intro]: assumes "C.ide f" shows "\\\<^sub>1 f : F.map\<^sub>0 (src\<^sub>C f) \\<^sub>D G.map\<^sub>0 (trg\<^sub>C f)\" and "\\\<^sub>1 f : G f \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D F f\" using assms map\<^sub>1_in_vhom [of f] D.vconn_implies_hpar(1-2) by auto lemma map\<^sub>1_simps [simp]: assumes "C.ide f" shows "D.arr (\\<^sub>1 f)" and "src\<^sub>D (\\<^sub>1 f) = F.map\<^sub>0 (src\<^sub>C f)" and "trg\<^sub>D (\\<^sub>1 f) = G.map\<^sub>0 (trg\<^sub>C f)" and "D.dom (\\<^sub>1 f) = G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)" and "D.cod (\\<^sub>1 f) = \\<^sub>0 (trg\<^sub>C f) \\<^sub>D F f" using assms map\<^sub>1_in_hom iso_map\<^sub>1_ide by auto lemma inv_naturality: assumes "C.arr \" shows "(G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D D.inv (\\<^sub>1 (C.dom \)) = D.inv (\\<^sub>1 (C.cod \)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D F \)" using assms naturality iso_map\<^sub>1_ide D.invert_opposite_sides_of_square by simp end subsection "Identity Pseudonatural Transformation" locale identity_pseudonatural_transformation = 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>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 abbreviation (input) map\<^sub>0 where "map\<^sub>0 a \ F.map\<^sub>0 a" abbreviation (input) map\<^sub>1 where "map\<^sub>1 f \ \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>D[F f]" sublocale pseudonatural_transformation 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 map\<^sub>0 map\<^sub>1 proof show 1: "\a. C.obj a \ D.ide (F.map\<^sub>0 a)" using F.map\<^sub>0_simps(1) by blast show "\f. C.ide f \ D.iso (map\<^sub>1 f)" by auto show "\a. C.obj a \ \F.map\<^sub>0 a : src\<^sub>D (F a) \\<^sub>D src\<^sub>D (F a)\" by (metis C.obj_def' D.in_hhomI D.src.preserves_arr F.map\<^sub>0_def F.map\<^sub>0_simps(2) F.map\<^sub>0_simps(3) F.preserves_reflects_arr) show "\f. C.ide f \ \map\<^sub>1 f : F f \\<^sub>D F.map\<^sub>0 (src\<^sub>C f) \\<^sub>D F.map\<^sub>0 (trg\<^sub>C f) \\<^sub>D F f\" by simp show "\\. C.arr \ \ (map\<^sub>1 (C.cod \)) \\<^sub>D (F \ \\<^sub>D F.map\<^sub>0 (src\<^sub>C \)) = (F.map\<^sub>0 (trg\<^sub>C \) \\<^sub>D F \) \\<^sub>D map\<^sub>1 (C.dom \)" by (metis D.comp_assoc D.lunit'_naturality D.runit_naturality F.preserves_arr F.preserves_cod F.preserves_dom F.preserves_src F.preserves_trg) show "\f g. \C.ide f; C.ide g; src\<^sub>C g = trg\<^sub>C f\ \ (F.map\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[F.map\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D[F g] \\<^sub>D F f) \\<^sub>D D.inv \\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>D[F f]) \\<^sub>D \\<^sub>D[F g, F f, F.map\<^sub>0 (src\<^sub>C f)] = (\\<^sub>D\<^sup>-\<^sup>1[F (g \\<^sub>C f)] \\<^sub>D \\<^sub>D[F (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>F (g, f) \\<^sub>D F.map\<^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 "(F.map\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[F.map\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D[F g] \\<^sub>D F f) \\<^sub>D D.inv \\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>D[F f]) \\<^sub>D \\<^sub>D[F g, F f, F.map\<^sub>0 (src\<^sub>C f)] = (F.map\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D (\\<^sub>D[F.map\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f)) \\<^sub>D (\\<^sub>D[F g] \\<^sub>D F f) \\<^sub>D (D.inv \\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f])) \\<^sub>D (F g \\<^sub>D \\<^sub>D[F f]) \\<^sub>D \\<^sub>D[F g, F f, F.map\<^sub>0 (src\<^sub>C f)]" using f g fg D.whisker_right [of "F f"] D.whisker_left [of "F g"] D.comp_assoc by simp also have "... = (F.map\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D ((\\<^sub>D[F g] \\<^sub>D F f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f)) \\<^sub>D \\<^sub>D[F g \\<^sub>D F f]" proof - have "D.inv \\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]) = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f" proof - have "D.inv \\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]) = D.inv ((F g \\<^sub>D \\<^sub>D[F f]) \\<^sub>D \\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f])" using f g fg 1 D.inv_comp by auto also have "... = D.inv (\\<^sub>D[F g] \\<^sub>D F f)" using f g fg D.triangle [of "F f" "F g"] by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f" using f g fg by simp finally show ?thesis by blast qed moreover have "(F g \\<^sub>D \\<^sub>D[F f]) \\<^sub>D \\<^sub>D[F g, F f, F.map\<^sub>0 (src\<^sub>C f)] = \\<^sub>D[F g \\<^sub>D F f]" using f g fg D.runit_hcomp(1) by simp moreover have "\\<^sub>D[F.map\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f) = \\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f]" using f g fg D.lunit_hcomp(4) by simp ultimately show ?thesis using D.comp_assoc by simp qed also have "... = ((F.map\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f]) \\<^sub>D \\<^sub>D[F g \\<^sub>D F f]" proof - have "((\\<^sub>D[F g] \\<^sub>D F f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f)) \\<^sub>D \\<^sub>D[F g \\<^sub>D F f] = \\<^sub>D[F g \\<^sub>D F f]" proof - have "((\\<^sub>D[F g] \\<^sub>D F f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f)) \\<^sub>D \\<^sub>D[F g \\<^sub>D F f] = (\\<^sub>D[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f) \\<^sub>D \\<^sub>D[F g \\<^sub>D F f]" using f g fg D.whisker_right by simp also have "... = (F g \\<^sub>D F f) \\<^sub>D \\<^sub>D[F g \\<^sub>D F f]" using f g fg D.comp_arr_inv' by simp also have "... = \\<^sub>D[F g \\<^sub>D F f]" using f g fg D.comp_cod_arr by simp finally show ?thesis by blast qed thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[F (g \\<^sub>C f)] \\<^sub>D \\<^sub>F (g, f) \\<^sub>D \\<^sub>D[F g \\<^sub>D F f]" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def D.lunit'_naturality [of "\\<^sub>F (g, f)"] D.comp_assoc by simp also have "... = (\\<^sub>D\<^sup>-\<^sup>1[F (g \\<^sub>C f)] \\<^sub>D \\<^sub>D[F (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>F (g, f) \\<^sub>D F.map\<^sub>0 (src\<^sub>C f))" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def D.runit_naturality [of "\\<^sub>F (g, f)"] D.comp_assoc by simp finally show "(F.map\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[F.map\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D[F g] \\<^sub>D F f) \\<^sub>D D.inv \\<^sub>D[F g, F.map\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>D[F f]) \\<^sub>D \\<^sub>D[F g, F f, F.map\<^sub>0 (src\<^sub>C f)] = (\\<^sub>D\<^sup>-\<^sup>1[F (g \\<^sub>C f)] \\<^sub>D \\<^sub>D[F (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>F (g, f) \\<^sub>D F.map\<^sub>0 (src\<^sub>C f))" by simp qed show "\a. C.obj a \ (F.map\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F.map\<^sub>0 a] \\<^sub>D \\<^sub>D[F.map\<^sub>0 a] = map\<^sub>1 a \\<^sub>D (F.unit a \\<^sub>D F.map\<^sub>0 a)" proof - fix a assume a: "C.obj a" have "(F.map\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F.map\<^sub>0 a] \\<^sub>D \\<^sub>D[F.map\<^sub>0 a] = (F.map\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F.map\<^sub>0 a] \\<^sub>D \\<^sub>D[F.map\<^sub>0 a]" using a 1 D.unitor_coincidence by simp also have "... = ((F.map\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F.map\<^sub>0 a]) \\<^sub>D \\<^sub>D[F.map\<^sub>0 a]" using D.comp_assoc by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[F a] \\<^sub>D F.unit a \\<^sub>D \\<^sub>D[F.map\<^sub>0 a]" using a 1 D.lunit'_naturality [of "F.unit a"] D.comp_assoc by simp also have "... = map\<^sub>1 a \\<^sub>D (F.unit a \\<^sub>D F.map\<^sub>0 a)" using a 1 D.runit_naturality [of "F.unit a"] D.comp_assoc by simp finally show "(F.map\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F.map\<^sub>0 a] \\<^sub>D \\<^sub>D[F.map\<^sub>0 a] = map\<^sub>1 a \\<^sub>D (F.unit a \\<^sub>D F.map\<^sub>0 a)" by blast qed qed lemma is_pseudonatural_transformation: shows "pseudonatural_transformation 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 map\<^sub>0 map\<^sub>1" .. end subsection "Composite Pseudonatural Transformation" text \ A pseudonatural transformation \\\ from \F\ to \G\ and a pseudonatural transformation \\\ from \G\ to \H\ can be composed to obtain a pseudonatural transformation \\\ from \F\ to \H\. The components at objects are composed via horizontal composition. Composing the components at 1-cells is straightforward, but is formally complicated by the need for associativities. The additional complexity leads to somewhat lengthy proofs of the coherence conditions. This issue only gets worse as we consider additional constructions on pseudonatural transformations. \ locale composite_pseudonatural_transformation = 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>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: 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 + H: 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 H \\<^sub>H + \: pseudonatural_transformation 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 \\<^sub>0 \\<^sub>1 + \: pseudonatural_transformation 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 H \\<^sub>H \\<^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 :: "'c \ 'd" and \\<^sub>F :: "'c * 'c \ 'd" and G :: "'c \ 'd" and \\<^sub>G :: "'c * 'c \ 'd" and H :: "'c \ 'd" and \\<^sub>H :: "'c * 'c \ 'd" and \\<^sub>0 :: "'c \ 'd" and \\<^sub>1 :: "'c \ 'd" and \\<^sub>0 :: "'c \ 'd" and \\<^sub>1 :: "'c \ 'd" begin definition map\<^sub>0 where "map\<^sub>0 a = \\<^sub>0 a \\<^sub>D \\<^sub>0 a" definition map\<^sub>1 where "map\<^sub>1 f = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C f), \\<^sub>0 (trg\<^sub>C f), F f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C f) \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]" sublocale pseudonatural_transformation 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 H \\<^sub>H map\<^sub>0 map\<^sub>1 proof show "\a. C.obj a \ D.ide (map\<^sub>0 a)" unfolding map\<^sub>0_def by fastforce show "\a. C.obj a \ \map\<^sub>0 a : src\<^sub>D (F a) \\<^sub>D src\<^sub>D (H a)\" unfolding map\<^sub>0_def using \.map\<^sub>0_in_hhom \.map\<^sub>0_in_hhom by blast show "\f. C.ide f \ \map\<^sub>1 f : H f \\<^sub>D map\<^sub>0 (src\<^sub>C f) \\<^sub>D map\<^sub>0 (trg\<^sub>C f) \\<^sub>D F f\" by (unfold map\<^sub>1_def map\<^sub>0_def, intro D.comp_in_homI) auto show "\f. C.ide f \ D.iso (map\<^sub>1 f)" unfolding map\<^sub>1_def using \.ide_map\<^sub>0_obj \.iso_map\<^sub>1_ide \.ide_map\<^sub>0_obj \.iso_map\<^sub>1_ide by (intro D.isos_compose) auto show "\\. C.arr \ \ map\<^sub>1 (C.cod \) \\<^sub>D (H \ \\<^sub>D map\<^sub>0 (src\<^sub>C \)) = (map\<^sub>0 (trg\<^sub>C \) \\<^sub>D F \) \\<^sub>D map\<^sub>1 (C.dom \)" proof - fix \ assume \: "C.arr \" have "map\<^sub>1 (C.cod \) \\<^sub>D (H \ \\<^sub>D map\<^sub>0 (src\<^sub>C \)) = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C \), \\<^sub>0 (trg\<^sub>C \), F (C.cod \)] \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>1 (C.cod \)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C \), G (C.cod \), \\<^sub>0 (src\<^sub>C \)] \\<^sub>D (\\<^sub>1 (C.cod \) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (C.cod \), \\<^sub>0 (src\<^sub>C \), \\<^sub>0 (src\<^sub>C \)] \\<^sub>D (H \ \\<^sub>D \\<^sub>0 (src\<^sub>C \) \\<^sub>D \\<^sub>0 (src\<^sub>C \))" unfolding map\<^sub>1_def map\<^sub>0_def using \ D.comp_assoc by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C \), \\<^sub>0 (trg\<^sub>C \), F (C.cod \)] \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>1 (C.cod \)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C \), G (C.cod \), \\<^sub>0 (src\<^sub>C \)] \\<^sub>D ((\\<^sub>1 (C.cod \) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D ((H \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \\<^sub>0 (src\<^sub>C \))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (C.dom \), \\<^sub>0 (src\<^sub>C \), \\<^sub>0 (src\<^sub>C \)]" using \ D.assoc'_naturality [of "H \" "\\<^sub>0 (src\<^sub>C \)" "\\<^sub>0 (src\<^sub>C \)"] D.comp_assoc by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C \), \\<^sub>0 (trg\<^sub>C \), F (C.cod \)] \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>1 (C.cod \)) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C \), G (C.cod \), \\<^sub>0 (src\<^sub>C \)] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \) \\<^sub>D \\<^sub>0 (src\<^sub>C \))) \\<^sub>D (\\<^sub>1 (C.dom \) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (C.dom \), \\<^sub>0 (src\<^sub>C \), \\<^sub>0 (src\<^sub>C \)]" proof - have "(\\<^sub>1 (C.cod \) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D ((H \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) = \\<^sub>1 (C.cod \) \\<^sub>D (H \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \\<^sub>0 (src\<^sub>C \)" using \ D.whisker_right by simp also have "... = (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \) \\<^sub>D \\<^sub>1 (C.dom \) \\<^sub>D \\<^sub>0 (src\<^sub>C \)" using \ \.naturality by simp also have "... = ((\\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D (\\<^sub>1 (C.dom \) \\<^sub>D \\<^sub>0 (src\<^sub>C \))" using \ D.whisker_right by simp finally have "(\\<^sub>1 (C.cod \) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D ((H \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) = ((\\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D (\\<^sub>1 (C.dom \) \\<^sub>D \\<^sub>0 (src\<^sub>C \))" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C \), \\<^sub>0 (trg\<^sub>C \), F (C.cod \)] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>1 (C.cod \)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \))) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C \), G (C.dom \), \\<^sub>0 (src\<^sub>C \)] \\<^sub>D (\\<^sub>1 (C.dom \) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (C.dom \), \\<^sub>0 (src\<^sub>C \), \\<^sub>0 (src\<^sub>C \)]" using \ D.assoc_naturality [of "\\<^sub>0 (trg\<^sub>C \)" "G \" "\\<^sub>0 (src\<^sub>C \)"] D.comp_assoc by simp also have "... = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C \), \\<^sub>0 (trg\<^sub>C \), F (C.cod \)] \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>0 (trg\<^sub>C \) \\<^sub>D F \)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>1 (C.dom \)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C \), G (C.dom \), \\<^sub>0 (src\<^sub>C \)] \\<^sub>D (\\<^sub>1 (C.dom \) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (C.dom \), \\<^sub>0 (src\<^sub>C \), \\<^sub>0 (src\<^sub>C \)]" proof - have "(\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>1 (C.cod \)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) = \\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>1 (C.cod \) \\<^sub>D (G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \))" using \ D.whisker_left by simp also have "... = \\<^sub>0 (trg\<^sub>C \) \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D F \) \\<^sub>D \\<^sub>1 (C.dom \)" using \ \.naturality by simp also have "... = (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>0 (trg\<^sub>C \) \\<^sub>D F \) \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>1 (C.dom \))" using \ D.whisker_left by simp finally have "(\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>1 (C.cod \)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) = (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>0 (trg\<^sub>C \) \\<^sub>D F \) \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>1 (C.dom \))" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>0 (trg\<^sub>C \)) \\<^sub>D F \) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C \), \\<^sub>0 (trg\<^sub>C \), F (C.dom \)] \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D \\<^sub>1 (C.dom \)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C \), G (C.dom \), \\<^sub>0 (src\<^sub>C \)] \\<^sub>D (\\<^sub>1 (C.dom \) \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (C.dom \), \\<^sub>0 (src\<^sub>C \), \\<^sub>0 (src\<^sub>C \)]" using \ D.assoc'_naturality [of "\\<^sub>0 (trg\<^sub>C \)" "\\<^sub>0 (trg\<^sub>C \)" "F \"] D.comp_assoc by simp also have "... = (map\<^sub>0 (trg\<^sub>C \) \\<^sub>D F \) \\<^sub>D map\<^sub>1 (C.dom \)" unfolding map\<^sub>1_def map\<^sub>0_def using \ by simp finally show "map\<^sub>1 (C.cod \) \\<^sub>D (H \ \\<^sub>D map\<^sub>0 (src\<^sub>C \)) = (map\<^sub>0 (trg\<^sub>C \) \\<^sub>D F \) \\<^sub>D map\<^sub>1 (C.dom \)" by blast qed show "\a. C.obj a \ (map\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \\<^sub>D \\<^sub>D[map\<^sub>0 a] = map\<^sub>1 a \\<^sub>D (H.unit a \\<^sub>D map\<^sub>0 a)" proof - fix a assume a: "C.obj a" have "map\<^sub>1 a \\<^sub>D (H.unit a \\<^sub>D map\<^sub>0 a) = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0 a, F a] \\<^sub>D (\\<^sub>0 a \\<^sub>D \\<^sub>1 a) \\<^sub>D \\<^sub>D[\\<^sub>0 a, G a, \\<^sub>0 a] \\<^sub>D (\\<^sub>1 a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H a, \\<^sub>0 a, \\<^sub>0 a] \\<^sub>D (H.unit a \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0 a)" unfolding map\<^sub>1_def map\<^sub>0_def using a C.obj_simps D.comp_assoc by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0 a, F a] \\<^sub>D (\\<^sub>0 a \\<^sub>D \\<^sub>1 a) \\<^sub>D \\<^sub>D[\\<^sub>0 a, G a, \\<^sub>0 a] \\<^sub>D ((\\<^sub>1 a \\<^sub>D \\<^sub>0 a) \\<^sub>D ((H.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0 a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \\<^sub>0 a, \\<^sub>0 a]" using a C.obj_simps D.assoc'_naturality [of "H.unit a" "\\<^sub>0 a" "\\<^sub>0 a"] D.comp_assoc by auto also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0 a, F a] \\<^sub>D (\\<^sub>0 a \\<^sub>D \\<^sub>1 a) \\<^sub>D (\\<^sub>D[\\<^sub>0 a, G a, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0 a \\<^sub>D G.unit a) \\<^sub>D \\<^sub>0 a)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \\<^sub>0 a, \\<^sub>0 a]" proof - have "(\\<^sub>1 a \\<^sub>D \\<^sub>0 a) \\<^sub>D ((H.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0 a) = \\<^sub>1 a \\<^sub>D (H.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0 a" using a C.obj_simps D.whisker_right by (metis C.objE D.hcomp_simps(4) D.hseqI' D.ideD(1) D.ideD(3) D.seqI H.unit_simps(1) H.unit_simps(2) H.unit_simps(5) \.ide_map\<^sub>0_obj \.map\<^sub>0_simps(3) \.map\<^sub>1_simps(1) \.map\<^sub>1_simps(4) \.ide_map\<^sub>0_obj) also have "... = (\\<^sub>0 a \\<^sub>D G.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a" using a C.obj_simps \.respects_unit D.comp_assoc by simp also have "... = ((\\<^sub>0 a \\<^sub>D G.unit a) \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a)" using a C.obj_simps D.whisker_right by (metis C.objE D.hcomp_simps(4) D.hseqI' D.ideD(1) D.ideD(3) D.seqI D.trg_cod H.unit_simps(1-2,5) H.\_simps(1,3,5) \.ide_map\<^sub>0_obj \.map\<^sub>0_simps(3) \.map\<^sub>1_simps(1,4) \.respects_unit \.ide_map\<^sub>0_obj) finally have "(\\<^sub>1 a \\<^sub>D \\<^sub>0 a) \\<^sub>D ((H.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0 a) = ((\\<^sub>0 a \\<^sub>D G.unit a) \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0 a, F a] \\<^sub>D ((\\<^sub>0 a \\<^sub>D \\<^sub>1 a) \\<^sub>D (\\<^sub>0 a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a)) \\<^sub>D \\<^sub>D[\\<^sub>0 a, G.map\<^sub>0 a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \\<^sub>0 a, \\<^sub>0 a]" using a D.assoc_naturality [of "\\<^sub>0 a" "G.unit a" "\\<^sub>0 a"] D.comp_assoc by auto also have "... = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0 a, F a] \\<^sub>D (\\<^sub>0 a \\<^sub>D \\<^sub>0 a \\<^sub>D F.unit a)) \\<^sub>D (\\<^sub>0 a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>D[\\<^sub>0 a, G.map\<^sub>0 a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \\<^sub>0 a, \\<^sub>0 a]" proof - have "(\\<^sub>0 a \\<^sub>D \\<^sub>1 a) \\<^sub>D (\\<^sub>0 a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a) = \\<^sub>0 a \\<^sub>D \\<^sub>1 a \\<^sub>D (G.unit a \\<^sub>D \\<^sub>0 a)" using a D.whisker_left [of "\\<^sub>0 a" "\\<^sub>1 a" "G.unit a \\<^sub>D \\<^sub>0 a"] by force also have "... = \\<^sub>0 a \\<^sub>D (\\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a]" using a \.respects_unit by simp also have "... = (\\<^sub>0 a \\<^sub>D \\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D (\\<^sub>0 a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a])" using a D.whisker_left C.obj_simps by fastforce finally have "(\\<^sub>0 a \\<^sub>D \\<^sub>1 a) \\<^sub>D (\\<^sub>0 a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a) = (\\<^sub>0 a \\<^sub>D \\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D (\\<^sub>0 a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a])" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0 a \\<^sub>D \\<^sub>0 a) \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0 a, F.map\<^sub>0 a] \\<^sub>D (\\<^sub>0 a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>D[\\<^sub>0 a, G.map\<^sub>0 a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \\<^sub>0 a, \\<^sub>0 a]" using a D.assoc'_naturality [of "\\<^sub>0 a" "\\<^sub>0 a" "F.unit a"] D.comp_assoc by fastforce also have "... = ((\\<^sub>0 a \\<^sub>D \\<^sub>0 a) \\<^sub>D F.unit a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0 a, F.map\<^sub>0 a] \\<^sub>D (\\<^sub>0 a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a])) \\<^sub>D (\\<^sub>0 a \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>D[\\<^sub>0 a, G.map\<^sub>0 a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H.map\<^sub>0 a, \\<^sub>0 a, \\<^sub>0 a]" using a D.whisker_left D.whisker_right D.comp_assoc by simp also have "... = ((\\<^sub>0 a \\<^sub>D \\<^sub>0 a) \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a \\<^sub>D \\<^sub>0 a] \\<^sub>D (\\<^sub>D[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, G.map\<^sub>0 a, \\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a, G.map\<^sub>0 a, \\<^sub>0 a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a)) \\<^sub>D \\<^sub>D[\\<^sub>0 a \\<^sub>D \\<^sub>0 a]" using a D.lunit_hcomp(3) [of "\\<^sub>0 a" "\\<^sub>0 a"] D.runit_hcomp(2) [of "\\<^sub>0 a" "\\<^sub>0 a"] D.triangle' [of "\\<^sub>0 a" "\\<^sub>0 a"] D.comp_assoc by auto also have "... = ((\\<^sub>0 a \\<^sub>D \\<^sub>0 a) \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a \\<^sub>D \\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a \\<^sub>D \\<^sub>0 a]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, G.map\<^sub>0 a, \\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a, G.map\<^sub>0 a, \\<^sub>0 a] = (\\<^sub>0 a \\<^sub>D G.map\<^sub>0 a) \\<^sub>D \\<^sub>0 a" using a D.comp_inv_arr' by (metis C.obj_def' D.comp_assoc_assoc'(2) G.map\<^sub>0_def G.map\<^sub>0_simps(1) G.preserves_trg G.weak_arrow_of_homs_axioms \.ide_map\<^sub>0_obj \.map\<^sub>0_simps(2) \.ide_map\<^sub>0_obj \.map\<^sub>0_simps(3) horizontal_homs.objE weak_arrow_of_homs_def) moreover have "((\\<^sub>0 a \\<^sub>D G.map\<^sub>0 a) \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a) = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a)" using a D.comp_cod_arr D.whisker_right by (metis D.runit'_simps(1) D.runit'_simps(5) G.map\<^sub>0_def \.ide_map\<^sub>0_obj \.map\<^sub>0_simps(2) \.ide_map\<^sub>0_obj) moreover have "(\\<^sub>D[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>0 a) = \\<^sub>0 a \\<^sub>D \\<^sub>0 a" using a D.whisker_right D.comp_arr_inv' D.R.components_are_iso by (metis D.ideD(1) D.iso_runit D.runit_simps(5) \.ide_map\<^sub>0_obj \.ide_map\<^sub>0_obj) moreover have "(\\<^sub>0 a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D[\\<^sub>0 a \\<^sub>D \\<^sub>0 a] = \\<^sub>D[\\<^sub>0 a \\<^sub>D \\<^sub>0 a]" using a D.comp_cod_arr \\a. C.obj a \ D.ide (map\<^sub>0 a)\ map\<^sub>0_def by auto ultimately show ?thesis using D.comp_assoc by metis qed also have "... = (map\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \\<^sub>D \\<^sub>D[map\<^sub>0 a]" unfolding map\<^sub>0_def by simp finally show "(map\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \\<^sub>D \\<^sub>D[map\<^sub>0 a] = map\<^sub>1 a \\<^sub>D (H.unit a \\<^sub>D map\<^sub>0 a)" by simp qed show "\f g. \C.ide f; C.ide g; src\<^sub>C g = trg\<^sub>C f\ \ (map\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[map\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (map\<^sub>1 g \\<^sub>D F f) \\<^sub>D D.inv \\<^sub>D[H g, map\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (H g \\<^sub>D map\<^sub>1 f) \\<^sub>D \\<^sub>D[H g, H f, map\<^sub>0 (src\<^sub>C f)] = map\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D map\<^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 "(map\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[map\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (map\<^sub>1 g \\<^sub>D F f) \\<^sub>D D.inv \\<^sub>D[H g, map\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (H g \\<^sub>D map\<^sub>1 f) \\<^sub>D \\<^sub>D[H g, H f, map\<^sub>0 (src\<^sub>C f)] = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D (\\<^sub>1 g \\<^sub>D \\<^sub>0 (src\<^sub>C g)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g) \\<^sub>D \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C f), \\<^sub>0 (trg\<^sub>C f), F f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C f) \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" unfolding map\<^sub>0_def map\<^sub>1_def using f g fg by simp also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D ((\\<^sub>1 g \\<^sub>D \\<^sub>0 (src\<^sub>C g)) \\<^sub>D F f) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g) \\<^sub>D \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C f), \\<^sub>0 (trg\<^sub>C f), F f]) \\<^sub>D (H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D \\<^sub>1 f) \\<^sub>D (H g \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D (H g \\<^sub>D \\<^sub>1 f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" using f g fg D.whisker_left D.whisker_right D.comp_assoc by fastforce also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D (((\\<^sub>1 g \\<^sub>D \\<^sub>0 (src\<^sub>C g)) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g \\<^sub>D \\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f] \\<^sub>D (H g \\<^sub>D \\<^sub>D[\\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C f), \\<^sub>0 (trg\<^sub>C f), F f]) \\<^sub>D (H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D \\<^sub>1 f) \\<^sub>D (H g \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D (H g \\<^sub>D \\<^sub>1 f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" using f g fg D.pentagon' D.comp_assoc D.invert_side_of_triangle(2) [of "\\<^sub>D\<^sup>-\<^sup>1[H g \\<^sub>D \\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f]" "(\\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g) \\<^sub>D \\<^sub>0 (src\<^sub>C g), F f]" "H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f]"] by force also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (\\<^sub>1 g \\<^sub>D \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f] \\<^sub>D (((H g \\<^sub>D \\<^sub>D[\\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C f), \\<^sub>0 (trg\<^sub>C f), F f])) \\<^sub>D (H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D \\<^sub>1 f)) \\<^sub>D (H g \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D (H g \\<^sub>D \\<^sub>1 f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" proof - have "((\\<^sub>1 g \\<^sub>D \\<^sub>0 (src\<^sub>C g)) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g \\<^sub>D \\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (\\<^sub>1 g \\<^sub>D \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f)" using f g fg D.assoc'_naturality [of "\\<^sub>1 g" "\\<^sub>0 (src\<^sub>C g)" "F f"] by simp thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (\\<^sub>1 g \\<^sub>D \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f] \\<^sub>D (H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D \\<^sub>1 f)) \\<^sub>D (H g \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D (H g \\<^sub>D \\<^sub>1 f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" proof - have "((H g \\<^sub>D \\<^sub>D[\\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C f), \\<^sub>0 (trg\<^sub>C f), F f])) \\<^sub>D (H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D \\<^sub>1 f) = H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D \\<^sub>1 f" proof - have "(H g \\<^sub>D \\<^sub>D[\\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f]) = H g \\<^sub>D \\<^sub>0 (src\<^sub>C g) \\<^sub>D \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f" proof - have "(H g \\<^sub>D \\<^sub>D[\\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f]) = H g \\<^sub>D \\<^sub>D[\\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (src\<^sub>C g), \\<^sub>0 (src\<^sub>C g), F f]" using f g fg D.whisker_left by simp also have "... = H g \\<^sub>D \\<^sub>0 (src\<^sub>C g) \\<^sub>D \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f" using f g fg D.comp_arr_inv' by simp finally show ?thesis by simp qed moreover have "(H g \\<^sub>D \\<^sub>0 (src\<^sub>C g) \\<^sub>D \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f) \\<^sub>D (H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D \\<^sub>1 f) = (H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D \\<^sub>1 f)" using f g fg D.comp_cod_arr by simp ultimately show ?thesis using fg by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D ((\\<^sub>1 g \\<^sub>D \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f) \\<^sub>D ((H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f)) \\<^sub>D \\<^sub>1 f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D (H g \\<^sub>D \\<^sub>1 f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" using f g fg D.assoc'_naturality [of "H g" "\\<^sub>0 (src\<^sub>C g)" "\\<^sub>1 f"] D.comp_assoc by simp also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g) \\<^sub>D \\<^sub>1 f) \\<^sub>D (\\<^sub>1 g \\<^sub>D G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)])) \\<^sub>D (H g \\<^sub>D \\<^sub>1 f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" proof - have "(\\<^sub>1 g \\<^sub>D \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f) \\<^sub>D ((H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f)) \\<^sub>D \\<^sub>1 f) = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g) \\<^sub>D \\<^sub>1 f) \\<^sub>D (\\<^sub>1 g \\<^sub>D G f \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg D.interchange D.comp_arr_dom D.comp_cod_arr by (metis C.ideD(1) \.map\<^sub>1_simps(1,5) \.naturality \.map\<^sub>1_simps(1,4) \.naturality C.ideD(2,3)) thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g) \\<^sub>D \\<^sub>1 f) \\<^sub>D (\\<^sub>1 g \\<^sub>D G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)])) \\<^sub>D \\<^sub>D[H g, \\<^sub>0 (trg\<^sub>C f) \\<^sub>D G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D ((H g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f \\<^sub>D \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" using f g fg D.comp_assoc D.hcomp_reassoc(2) by simp also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g) \\<^sub>D \\<^sub>1 f) \\<^sub>D ((\\<^sub>1 g \\<^sub>D G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D[H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (trg\<^sub>C f), G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((H g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f \\<^sub>D \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)])) \\<^sub>D \\<^sub>D[H g, \\<^sub>0 (trg\<^sub>C f) \\<^sub>D G f, \\<^sub>0 (src\<^sub>C f)] = \\<^sub>D[H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (trg\<^sub>C f), G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f))" proof - have "((H g \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, \\<^sub>0 (trg\<^sub>C f) \\<^sub>D G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D (\\<^sub>D[H g, \\<^sub>0 (trg\<^sub>C f), G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = \\<^sub>D[H g, \\<^sub>0 (trg\<^sub>C f), G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>D[H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)]" using f g fg D.pentagon D.comp_assoc by simp moreover have "D.seq \\<^sub>D[H g, \\<^sub>0 (trg\<^sub>C f), G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D[H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)]" using f g fg by (intro D.seqI) auto ultimately show ?thesis using f g fg D.comp_assoc D.invert_opposite_sides_of_square [of "\\<^sub>D[H g, \\<^sub>0 (trg\<^sub>C f), G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" "\\<^sub>D[H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)]" "(H g \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, \\<^sub>0 (trg\<^sub>C f) \\<^sub>D G f, \\<^sub>0 (src\<^sub>C f)]" "\\<^sub>D[H g, \\<^sub>0 (trg\<^sub>C f), G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)"] by fastforce qed thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g) \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (((\\<^sub>1 g \\<^sub>D G f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (trg\<^sub>C f), G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((H g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f \\<^sub>D \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" proof - have "(\\<^sub>1 g \\<^sub>D G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D[H g \\<^sub>D \\<^sub>0 (trg\<^sub>C f), G f, \\<^sub>0 (src\<^sub>C f)] = \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D ((\\<^sub>1 g \\<^sub>D G f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg D.assoc_naturality [of "\\<^sub>1 g" "G f" "\\<^sub>0 (src\<^sub>C f)"] by simp thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g) \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f \\<^sub>D \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)])" proof - have "((\\<^sub>1 g \\<^sub>D G f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (trg\<^sub>C f), G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((H g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = (\\<^sub>1 g \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (trg\<^sub>C f), G f] \\<^sub>D (H g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)" using f g fg D.whisker_right by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>0 (src\<^sub>C f)" proof - have "((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D ((\\<^sub>1 g \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), G f] \\<^sub>D (H g \\<^sub>D \\<^sub>1 f))) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f)] = \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg \.respects_hcomp D.comp_assoc by simp moreover have "D.seq (\\<^sub>1 (g \\<^sub>C f)) (\\<^sub>H (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 by auto ultimately have "(\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D ((\\<^sub>1 g \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), G f] \\<^sub>D (H g \\<^sub>D \\<^sub>1 f)) = (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)]" using f g fg D.invert_side_of_triangle(2) by simp moreover have "D.seq (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)]" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def by auto ultimately have "\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D ((\\<^sub>1 g \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), G f] \\<^sub>D (H g \\<^sub>D \\<^sub>1 f)) = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)]" using f g fg D.invert_side_of_triangle(1) [of "(\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)]" "\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f)" "\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D ((\\<^sub>1 g \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), G f] \\<^sub>D (H g \\<^sub>D \\<^sub>1 f))"] by simp hence "(\\<^sub>1 g \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), G f] \\<^sub>D (H g \\<^sub>D \\<^sub>1 f) = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)]" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def D.invert_side_of_triangle(1) by simp hence "(\\<^sub>1 g \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (src\<^sub>C g), G f] \\<^sub>D (H g \\<^sub>D \\<^sub>1 f) = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)]" using D.comp_assoc by simp thus ?thesis using fg by simp qed also have "... = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg D.whisker_right C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def G.FF_def by force finally have "((\\<^sub>1 g \\<^sub>D G f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[H g, \\<^sub>0 (trg\<^sub>C f), G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((H g \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>0 (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 \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g) \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g \\<^sub>D H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]" proof - have "((\\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f \\<^sub>D \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)])) \\<^sub>D \\<^sub>D[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)] = \\<^sub>D\<^sup>-\<^sup>1[H g \\<^sub>D H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f \\<^sub>D \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]) = \\<^sub>D\<^sup>-\<^sup>1[H g \\<^sub>D H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]" using f g fg D.pentagon' D.comp_assoc by simp moreover have "D.seq (\\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) (\\<^sub>D\<^sup>-\<^sup>1[H g, H f \\<^sub>D \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]))" using f g fg by auto ultimately show ?thesis using f g fg D.comp_assoc D.invert_side_of_triangle(2) [of "(\\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H g, H f \\<^sub>D \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (H g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)])" "\\<^sub>D\<^sup>-\<^sup>1[H g \\<^sub>D H f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)]" "\\<^sub>D\<^sup>-\<^sup>1[H g, H f, \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)]"] by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g) \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g) \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg D.assoc'_naturality [of "\\<^sub>H (g, f)" "\\<^sub>0 (src\<^sub>C f)" "\\<^sub>0 (src\<^sub>C f)"] by fastforce also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g) \\<^sub>D F g, F f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g \\<^sub>D F f) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g \\<^sub>D \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g) \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg D.hcomp_reassoc(1) D.comp_assoc by simp also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g) \\<^sub>D F g, F f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g \\<^sub>D F f) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g) \\<^sub>D \\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg D.comp_assoc D.pentagon D.invert_opposite_sides_of_square [of "\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D[G g, \\<^sub>0 (src\<^sub>C g), F f]" "\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g \\<^sub>D \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g)] \\<^sub>D F f)" "\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, \\<^sub>0 (src\<^sub>C g) \\<^sub>D F f]" "\\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, \\<^sub>0 (src\<^sub>C g), F f]"] by simp also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g) \\<^sub>D F g, F f] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g \\<^sub>D F f) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g \\<^sub>D \\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg D.assoc_naturality [of "\\<^sub>0 (trg\<^sub>C g)" "G g" "\\<^sub>1 f"] D.comp_assoc by simp also have "... = ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g) \\<^sub>D F g, F f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), F g, F f])) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>F (g, f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 (g \\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" proof - have "(\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g \\<^sub>D F f) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g \\<^sub>D \\<^sub>1 f) = \\<^sub>0 (trg\<^sub>C g) \\<^sub>D (\\<^sub>1 g \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f)" using f g fg D.whisker_left by simp also have "... = \\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>F (g, f))) \\<^sub>D \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (src\<^sub>C f)]" proof - have "((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>1 g \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f)) \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 (src\<^sub>C f)] = \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg \.respects_hcomp D.comp_assoc by simp hence "(\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>1 g \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f) = (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (src\<^sub>C f)]" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char D.invert_side_of_triangle(2) by simp hence "\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>1 g \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f) = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>F (g, f))) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (src\<^sub>C f)]" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char G.FF_def D.invert_side_of_triangle(1) [of "(\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (src\<^sub>C f)]" "\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)" "\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>1 g \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f)"] by simp hence "(\\<^sub>1 g \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f) = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>F (g, f))) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (src\<^sub>C f)]" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char G.FF_def F.FF_def D.invert_side_of_triangle(1) by simp thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), F g, F f]) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>F (g, f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 (g \\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (src\<^sub>C f)])" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def G.FF_def D.whisker_left by force finally have "(\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 g \\<^sub>D F f) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>C g), F f]) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g \\<^sub>D \\<^sub>1 f) = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), F g, F f]) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>F (g, f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 (g \\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (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 \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g \\<^sub>D F f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>F (g, f)))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 (g \\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg D.pentagon' D.comp_assoc D.invert_side_of_triangle(1) [of "(\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g) \\<^sub>D F g, F f] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), F g, F f])" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g), F g, F f]" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F g \\<^sub>D F f]" ] by simp also have "... = ((((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D D.inv (\\<^sub>F (g, f)))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 (g \\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg D.assoc'_naturality [of "\\<^sub>0 (trg\<^sub>C g)" "\\<^sub>0 (trg\<^sub>C g)" "D.inv (\\<^sub>F (g, f))"] D.comp_assoc by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F (g \\<^sub>C f)] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 (g \\<^sub>C f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" proof - have "(((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D D.inv (\\<^sub>F (g, f)))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F (g \\<^sub>C f)] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F (g \\<^sub>C f)]" proof - have "((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D D.inv (\\<^sub>F (g, f))) = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f) \\<^sub>D D.inv (\\<^sub>F (g, f))" using f g fg D.whisker_left C.VV.arr_char by simp also have "... = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D F (g \\<^sub>C f)" using f g fg D.comp_arr_inv' F.cmp_components_are_iso by simp finally have "((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D D.inv (\\<^sub>F (g, f))) = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D F (g \\<^sub>C f)" by blast moreover have "((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>0 (trg\<^sub>C g)) \\<^sub>D F (g \\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F (g \\<^sub>C f)] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F (g \\<^sub>C f)]" using f g fg D.comp_cod_arr by simp ultimately show ?thesis by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F (g \\<^sub>C f)] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 (g \\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g \\<^sub>D G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg D.pentagon D.comp_assoc D.invert_side_of_triangle(1) [of "\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)]" "\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 (src\<^sub>C f)]" "\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g \\<^sub>D G f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f))"] D.invert_side_of_triangle(2) [of "(\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)] \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G g, G f, \\<^sub>0 (src\<^sub>C f)]" "\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g \\<^sub>D G f, \\<^sub>0 (src\<^sub>C f)]" "\\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g, G f] \\<^sub>D \\<^sub>0 (src\<^sub>C f)"] by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F (g \\<^sub>C f)] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D ((((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" proof - have "(\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G g \\<^sub>D G f, \\<^sub>0 (src\<^sub>C f)] = \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f g fg D.assoc_naturality [of "\\<^sub>0 (trg\<^sub>C g)" "\\<^sub>G (g, f)" "\\<^sub>0 (src\<^sub>C f)"] by fastforce thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (trg\<^sub>C g), \\<^sub>0 (trg\<^sub>C g), F (g \\<^sub>C f)] \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C g), G (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (g \\<^sub>C f), \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D \\<^sub>0 (src\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f))" proof - have "(((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f))) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = \\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)" proof - have "((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)" using f g fg C.VV.arr_char D.whisker_right by simp also have "... = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)" using f g fg C.VV.arr_char D.whisker_left [of "\\<^sub>0 (trg\<^sub>C g)"] by simp also have "... = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G (g \\<^sub>C f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)" using f g fg D.comp_arr_inv' G.cmp_components_are_iso G.cmp_simps(5) by auto finally have "((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D ((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = (\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G (g \\<^sub>C f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)" by blast moreover have "((\\<^sub>0 (trg\<^sub>C g) \\<^sub>D G (g \\<^sub>C f)) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)) = \\<^sub>1 (g \\<^sub>C f) \\<^sub>D \\<^sub>0 (src\<^sub>C f)" using f g fg D.comp_cod_arr by simp ultimately show ?thesis by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = map\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D map\<^sub>0 (src\<^sub>C f))" unfolding map\<^sub>0_def map\<^sub>1_def using f g fg D.comp_assoc by simp finally show "(map\<^sub>0 (trg\<^sub>C g) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[map\<^sub>0 (trg\<^sub>C g), F g, F f] \\<^sub>D (map\<^sub>1 g \\<^sub>D F f) \\<^sub>D D.inv \\<^sub>D[H g, map\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (H g \\<^sub>D map\<^sub>1 f) \\<^sub>D \\<^sub>D[H g, H f, map\<^sub>0 (src\<^sub>C f)] = map\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>H (g, f) \\<^sub>D map\<^sub>0 (src\<^sub>C f))" using D.comp_assoc by simp qed qed lemma is_pseudonatural_transformation: shows "pseudonatural_transformation 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 H \\<^sub>H map\<^sub>0 map\<^sub>1" .. end subsection "Whiskering of Pseudonatural Transformations" text \ Similarly to ordinary natural transformations, pseudonatural transformations can be whiskered with pseudofunctors on the left and the right. \ locale pseudonatural_transformation_whisker_right = 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 + \.F: 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 + \.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 + H: 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 H \\<^sub>H + \: pseudonatural_transformation 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 \\<^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 \ 'd" and \\<^sub>F :: "'c * 'c \ 'd" and G :: "'c \ 'd" and \\<^sub>G :: "'c * 'c \ 'd" and H :: "'b \ 'c" and \\<^sub>H :: "'b * 'b \ 'c" and \\<^sub>0 :: "'c \ 'd" and \\<^sub>1 :: "'c \ 'd" begin interpretation FoH: 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 H \\<^sub>H F \\<^sub>F .. interpretation GoH: 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 H \\<^sub>H G \\<^sub>G .. definition map\<^sub>0 where "map\<^sub>0 a = \\<^sub>0 (H.map\<^sub>0 a)" definition map\<^sub>1 where "map\<^sub>1 f = \\<^sub>1 (H f)" sublocale pseudonatural_transformation 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 \F o H\ FoH.cmp \G o H\ GoH.cmp map\<^sub>0 map\<^sub>1 proof show "\a. B.obj a \ D.ide (map\<^sub>0 a)" using map\<^sub>0_def by simp show "\a. B.obj a \ \map\<^sub>0 a : src\<^sub>D ((F \ H) a) \\<^sub>D src\<^sub>D ((G \ H) a)\" using map\<^sub>0_def \.map\<^sub>0_in_hhom B.obj_simps C.obj_simps by simp show "\f. B.ide f \ D.iso (map\<^sub>1 f)" using map\<^sub>1_def \.iso_map\<^sub>1_ide by simp show "\f. B.ide f \ \map\<^sub>1 f : (G \ H) f \\<^sub>D map\<^sub>0 (src\<^sub>B f) \\<^sub>D map\<^sub>0 (trg\<^sub>B f) \\<^sub>D (F \ H) f\" using map\<^sub>0_def map\<^sub>1_def by auto show "\\. B.arr \ \ map\<^sub>1 (B.cod \) \\<^sub>D ((G \ H) \ \\<^sub>D map\<^sub>0 (src\<^sub>B \)) = (map\<^sub>0 (trg\<^sub>B \) \\<^sub>D (F \ H) \) \\<^sub>D map\<^sub>1 (B.dom \)" unfolding map\<^sub>0_def map\<^sub>1_def using \.naturality by force show "\a. B.obj a \ (map\<^sub>0 a \\<^sub>D FoH.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \\<^sub>D \\<^sub>D[map\<^sub>0 a] = map\<^sub>1 a \\<^sub>D (GoH.unit a \\<^sub>D map\<^sub>0 a)" proof - fix a assume a: "B.obj a" have "map\<^sub>1 a \\<^sub>D (GoH.unit a \\<^sub>D map\<^sub>0 a) = \\<^sub>1 (H a) \\<^sub>D (G (H.unit a) \\<^sub>D \.G.unit (H.map\<^sub>0 a) \\<^sub>D \\<^sub>0 (H.map\<^sub>0 a))" unfolding map\<^sub>0_def map\<^sub>1_def using a GoH.unit_char' by simp also have "... = (\\<^sub>1 (H a) \\<^sub>D (G (H.unit a) \\<^sub>D \\<^sub>0 (H.map\<^sub>0 a))) \\<^sub>D (\.G.unit (H.map\<^sub>0 a) \\<^sub>D \\<^sub>0 (H.map\<^sub>0 a))" using a D.whisker_right D.comp_assoc by simp also have "... = (\\<^sub>0 (H.map\<^sub>0 a) \\<^sub>D F (H.unit a)) \\<^sub>D \\<^sub>1 (H.map\<^sub>0 a) \\<^sub>D (\.G.unit (H.map\<^sub>0 a) \\<^sub>D \\<^sub>0 (H.map\<^sub>0 a))" using a \.naturality [of "H.unit a"] D.comp_assoc by simp also have "... = ((\\<^sub>0 (H.map\<^sub>0 a) \\<^sub>D F (H.unit a)) \\<^sub>D (\\<^sub>0 (H.map\<^sub>0 a) \\<^sub>D \.F.unit (H.map\<^sub>0 a))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (H.map\<^sub>0 a)] \\<^sub>D \\<^sub>D[\\<^sub>0 (H.map\<^sub>0 a)]" using a \.respects_unit D.comp_assoc by simp also have "... = (\\<^sub>0 (H.map\<^sub>0 a) \\<^sub>D F (H.unit a) \\<^sub>D \.F.unit (H.map\<^sub>0 a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 (H.map\<^sub>0 a)] \\<^sub>D \\<^sub>D[\\<^sub>0 (H.map\<^sub>0 a)]" using a D.whisker_left by simp also have "... = (map\<^sub>0 a \\<^sub>D FoH.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \\<^sub>D \\<^sub>D[map\<^sub>0 a]" unfolding map\<^sub>0_def map\<^sub>1_def using a FoH.unit_char' by simp finally show "(map\<^sub>0 a \\<^sub>D FoH.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \\<^sub>D \\<^sub>D[map\<^sub>0 a] = map\<^sub>1 a \\<^sub>D (GoH.unit a \\<^sub>D map\<^sub>0 a)" by simp qed show "\f g. \B.ide f; B.ide g; src\<^sub>B g = trg\<^sub>B f\ \ (map\<^sub>0 (trg\<^sub>B g) \\<^sub>D FoH.cmp (g, f)) \\<^sub>D \\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (F \ H) g, (F \ H) f] \\<^sub>D (map\<^sub>1 g \\<^sub>D (F \ H) f) \\<^sub>D D.inv \\<^sub>D[(G \ H) g, map\<^sub>0 (src\<^sub>B g), (F \ H) f] \\<^sub>D ((G \ H) g \\<^sub>D map\<^sub>1 f) \\<^sub>D \\<^sub>D[(G \ H) g, (G \ H) f, map\<^sub>0 (src\<^sub>B f)] = map\<^sub>1 (g \\<^sub>B f) \\<^sub>D (GoH.cmp (g, f) \\<^sub>D map\<^sub>0 (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" have "map\<^sub>1 (g \\<^sub>B f) \\<^sub>D (GoH.cmp (g, f) \\<^sub>D map\<^sub>0 (src\<^sub>B f)) = \\<^sub>1 (H (g \\<^sub>B f)) \\<^sub>D (G (H (g \\<^sub>B f)) \\<^sub>D G (\\<^sub>H (g, f)) \\<^sub>D \\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))" unfolding map\<^sub>0_def map\<^sub>1_def using f g fg GoH.cmp_def B.VV.arr_char B.VV.dom_char by simp also have "... = (\\<^sub>1 (H (g \\<^sub>B f)) \\<^sub>D (G (H (g \\<^sub>B f)) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))) \\<^sub>D (G (\\<^sub>H (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f))) \\<^sub>D (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))" using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char H.FF_def C.VV.arr_char C.VV.dom_char C.VV.cod_char D.whisker_right D.comp_assoc by simp also have "... = (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D F (H (g \\<^sub>B f))) \\<^sub>D (\\<^sub>1 (H (g \\<^sub>B f)) \\<^sub>D (G (\\<^sub>H (g, f)) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))) \\<^sub>D (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))" using f g fg \.naturality [of "H (g \\<^sub>B f)"] D.comp_assoc by auto also have "... = (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D F (H (g \\<^sub>B f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D F (\\<^sub>H (g, f))) \\<^sub>D \\<^sub>1 (H g \\<^sub>C H f) \\<^sub>D (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))" using f g fg \.naturality [of "\\<^sub>H (g, f)"] D.comp_assoc by fastforce also have "... = (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D F (H (g \\<^sub>B f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D F (\\<^sub>H (g, f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D \\<^sub>F (H g, H f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \\<^sub>D (\\<^sub>1 (H g) \\<^sub>D F (H f)) \\<^sub>D D.inv \\<^sub>D[G (H g), \\<^sub>0 (src\<^sub>C (H g)), F (H f)] \\<^sub>D (G (H g) \\<^sub>D \\<^sub>1 (H f)) \\<^sub>D \\<^sub>D[G (H g), G (H f), \\<^sub>0 (src\<^sub>C (H f))] \\<^sub>D (D.inv (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f))) \\<^sub>D (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f))))" proof - have "\\<^sub>1 (H g \\<^sub>C H f) = (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D \\<^sub>F (H g, H f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \\<^sub>D (\\<^sub>1 (H g) \\<^sub>D F (H f)) \\<^sub>D D.inv \\<^sub>D[G (H g), \\<^sub>0 (src\<^sub>C (H g)), F (H f)] \\<^sub>D (G (H g) \\<^sub>D \\<^sub>1 (H f)) \\<^sub>D \\<^sub>D[G (H g), G (H f), \\<^sub>0 (src\<^sub>C (H f))] \\<^sub>D D.inv (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))" proof - have "\\<^sub>1 (H g \\<^sub>C H f) \\<^sub>D (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f))) = (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D \\<^sub>F (H g, H f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \\<^sub>D (\\<^sub>1 (H g) \\<^sub>D F (H f)) \\<^sub>D D.inv \\<^sub>D[G (H g), \\<^sub>0 (src\<^sub>C (H g)), F (H f)] \\<^sub>D (G (H g) \\<^sub>D \\<^sub>1 (H f)) \\<^sub>D \\<^sub>D[G (H g), G (H f), \\<^sub>0 (src\<^sub>C (H f))]" using f g fg \.respects_hcomp [of "H f" "H g"] by simp moreover have "D.seq (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D \\<^sub>F (H g, H f)) (\\<^sub>D[\\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \\<^sub>D (\\<^sub>1 (H g) \\<^sub>D F (H f)) \\<^sub>D D.inv \\<^sub>D[G (H g), \\<^sub>0 (src\<^sub>C (H g)), F (H f)] \\<^sub>D (G (H g) \\<^sub>D \\<^sub>1 (H f)) \\<^sub>D \\<^sub>D[G (H g), G (H f), \\<^sub>0 (src\<^sub>C (H f))])" using f g fg C.obj_simps C.VV.arr_char C.VV.dom_char C.VV.cod_char \.F.FF_def by simp moreover have "D.iso (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))" using f g fg \.G.cmp_components_are_iso [of "H g" "H f"] by simp ultimately have "\\<^sub>1 (H g \\<^sub>C H f) = ((\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D \\<^sub>F (H g, H f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \\<^sub>D (\\<^sub>1 (H g) \\<^sub>D F (H f)) \\<^sub>D D.inv \\<^sub>D[G (H g), \\<^sub>0 (src\<^sub>C (H g)), F (H f)] \\<^sub>D (G (H g) \\<^sub>D \\<^sub>1 (H f)) \\<^sub>D \\<^sub>D[G (H g), G (H f), \\<^sub>0 (src\<^sub>C (H f))]) \\<^sub>D D.inv (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))" using D.invert_side_of_triangle(2) by blast thus ?thesis using D.comp_assoc by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D F (H (g \\<^sub>B f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D F (\\<^sub>H (g, f))) \\<^sub>D (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D \\<^sub>F (H g, H f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \\<^sub>D (\\<^sub>1 (H g) \\<^sub>D F (H f)) \\<^sub>D D.inv \\<^sub>D[G (H g), \\<^sub>0 (src\<^sub>C (H g)), F (H f)] \\<^sub>D (G (H g) \\<^sub>D \\<^sub>1 (H f)) \\<^sub>D \\<^sub>D[G (H g), G (H f), \\<^sub>0 (src\<^sub>C (H f))]" proof - have "(D.inv (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f))) \\<^sub>D (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))) = (G (H g) \\<^sub>D G (H f)) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f))" proof - have "(D.inv (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f))) \\<^sub>D (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))) = (D.inv (\\<^sub>G (H g, H f)) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f))) \\<^sub>D (\\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f)))" using f g fg by simp also have "... = D.inv (\\<^sub>G (H g, H f)) \\<^sub>D \\<^sub>G (H g, H f) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f))" using f g fg D.whisker_right C.VV.arr_char by simp also have "... = (G (H g) \\<^sub>D G (H f)) \\<^sub>D \\<^sub>0 (src\<^sub>C (H f))" using f g fg D.comp_inv_arr' \.G.cmp_components_are_iso by simp finally show ?thesis by blast qed moreover have "\\<^sub>D[G (H g), G (H f), \\<^sub>0 (src\<^sub>C (H f))] \\<^sub>D ... = \\<^sub>D[G (H g), G (H f), \\<^sub>0 (src\<^sub>C (H f))]" using f g fg D.comp_arr_dom by simp ultimately show ?thesis by simp qed also have "... = (\\<^sub>0 (trg\<^sub>C (H g)) \\<^sub>D F (H (g \\<^sub>B f)) \\<^sub>D F (\\<^sub>H (g, f)) \\<^sub>D \\<^sub>F (H g, H f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C (H g)), F (H g), F (H f)] \\<^sub>D (\\<^sub>1 (H g) \\<^sub>D F (H f)) \\<^sub>D D.inv \\<^sub>D[G (H g), \\<^sub>0 (src\<^sub>C (H g)), F (H f)] \\<^sub>D (G (H g) \\<^sub>D \\<^sub>1 (H f)) \\<^sub>D \\<^sub>D[G (H g), G (H f), \\<^sub>0 (src\<^sub>C (H f))]" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char B.VV.arr_char B.VV.dom_char B.VV.cod_char H.FF_def D.whisker_left C.VV.arr_char B.VV.arr_char D.comp_assoc by auto also have "... = (map\<^sub>0 (trg\<^sub>B g) \\<^sub>D FoH.cmp (g, f)) \\<^sub>D \\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (F \ H) g, (F \ H) f] \\<^sub>D (map\<^sub>1 g \\<^sub>D (F \ H) f) \\<^sub>D D.inv \\<^sub>D[(G \ H) g, map\<^sub>0 (src\<^sub>B g), (F \ H) f] \\<^sub>D ((G \ H) g \\<^sub>D map\<^sub>1 f) \\<^sub>D \\<^sub>D[(G \ H) g, (G \ H) f, map\<^sub>0 (src\<^sub>B f)]" unfolding map\<^sub>0_def map\<^sub>1_def using f g fg FoH.cmp_def B.VV.arr_char B.VV.dom_char by simp finally show "(map\<^sub>0 (trg\<^sub>B g) \\<^sub>D FoH.cmp (g, f)) \\<^sub>D \\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (F \ H) g, (F \ H) f] \\<^sub>D (map\<^sub>1 g \\<^sub>D (F \ H) f) \\<^sub>D D.inv \\<^sub>D[(G \ H) g, map\<^sub>0 (src\<^sub>B g), (F \ H) f] \\<^sub>D ((G \ H) g \\<^sub>D map\<^sub>1 f) \\<^sub>D \\<^sub>D[(G \ H) g, (G \ H) f, map\<^sub>0 (src\<^sub>B f)] = map\<^sub>1 (g \\<^sub>B f) \\<^sub>D (GoH.cmp (g, f) \\<^sub>D map\<^sub>0 (src\<^sub>B f))" by simp qed qed lemma is_pseudonatural_transformation: shows "pseudonatural_transformation 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 (F o H) FoH.cmp (G o H) GoH.cmp map\<^sub>0 map\<^sub>1" .. end locale pseudonatural_transformation_whisker_left = 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 + \.F: 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 + \.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>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 + \: pseudonatural_transformation 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 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 :: "'b \ 'c" and \\<^sub>F :: "'b * 'b \ 'c" and G :: "'b \ 'c" and \\<^sub>G :: "'b * 'b \ 'c" and H :: "'c \ 'd" and \\<^sub>H :: "'c * 'c \ 'd" and \\<^sub>0 :: "'b \ 'c" and \\<^sub>1 :: "'b \ 'c" begin interpretation HoF: 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 F \\<^sub>F H \\<^sub>H .. interpretation HoG: 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 H \\<^sub>H .. definition map\<^sub>0 where "map\<^sub>0 a = H (\\<^sub>0 a)" definition map\<^sub>1 where "map\<^sub>1 f = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B f), F f)) \\<^sub>D H (\\<^sub>1 f) \\<^sub>D \\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f))" sublocale pseudonatural_transformation 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 HoF.map HoF.cmp HoG.map HoG.cmp map\<^sub>0 map\<^sub>1 proof show "\a. B.obj a \ D.ide (map\<^sub>0 a)" using map\<^sub>0_def by simp show "\f. B.ide f \ D.iso (map\<^sub>1 f)" proof - fix f assume f: "B.ide f" have "D.seq (H (\\<^sub>1 f)) (\\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f)))" using f by (intro D.seqI) auto moreover have "D.seq (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B f), F f))) (H (\\<^sub>1 f) \\<^sub>D \\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f)))" using f \.map\<^sub>1_in_hom [of f] calculation by (intro D.seqI) auto ultimately show "D.iso (map\<^sub>1 f)" using f map\<^sub>1_def H.preserves_iso \.iso_map\<^sub>1_ide H.cmp_components_are_iso C.VV.arr_char D.isos_compose by auto qed show "\a. B.obj a \ \map\<^sub>0 a : src\<^sub>D ((H \ F) a) \\<^sub>D src\<^sub>D ((H \ G) a)\" using map\<^sub>0_def by fastforce show "\f. B.ide f \ \map\<^sub>1 f : (H \ G) f \\<^sub>D map\<^sub>0 (src\<^sub>B f) \\<^sub>D map\<^sub>0 (trg\<^sub>B f) \\<^sub>D (H \ F) f\" using map\<^sub>0_def map\<^sub>1_def by fastforce show "\\. B.arr \ \ map\<^sub>1 (B.cod \) \\<^sub>D ((H \ G) \ \\<^sub>D map\<^sub>0 (src\<^sub>B \)) = (map\<^sub>0 (trg\<^sub>B \) \\<^sub>D (H \ F) \) \\<^sub>D map\<^sub>1 (B.dom \)" proof - fix \ assume \: "B.arr \" have "(map\<^sub>0 (trg\<^sub>B \) \\<^sub>D (H \ F) \) \\<^sub>D map\<^sub>1 (B.dom \) = (H (\\<^sub>0 (trg\<^sub>B \)) \\<^sub>D (H \ F) \) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B (B.dom \)), F (B.dom \))) \\<^sub>D H (\\<^sub>1 (B.dom \)) \\<^sub>D \\<^sub>H (G (B.dom \), \\<^sub>0 (src\<^sub>B (B.dom \)))" unfolding map\<^sub>0_def map\<^sub>1_def using \ by simp also have "... = ((H (\\<^sub>0 (trg\<^sub>B \)) \\<^sub>D H (F \)) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.dom \)))) \\<^sub>D H (\\<^sub>1 (B.dom \)) \\<^sub>D \\<^sub>H (G (B.dom \), \\<^sub>0 (src\<^sub>B \))" using \ D.comp_assoc by simp also have "... = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.cod \))) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B \) \\<^sub>C F \) \\<^sub>D H (\\<^sub>1 (B.dom \))) \\<^sub>D \\<^sub>H (G (B.dom \), \\<^sub>0 (src\<^sub>B (B.dom \)))" proof - have "\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.cod \)) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B \)) \\<^sub>D H (F \)) = H (\\<^sub>0 (trg\<^sub>B \) \\<^sub>C F \) \\<^sub>D \\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.dom \))" using \ H.\.naturality [of "(\\<^sub>0 (trg\<^sub>B \), F \)"] C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def by simp moreover have "D.seq (\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.cod \))) (H (\\<^sub>0 (trg\<^sub>B \)) \\<^sub>D H (F \))" using \ by (intro D.seqI D.hseqI') auto moreover have "D.iso (\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.cod \)))" using \ H.cmp_components_are_iso by simp moreover have "D.iso (\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.dom \)))" using \ H.cmp_components_are_iso by simp ultimately have "(H (\\<^sub>0 (trg\<^sub>B \)) \\<^sub>D H (F \)) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.dom \))) = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.cod \))) \\<^sub>D H (\\<^sub>0 (trg\<^sub>B \) \\<^sub>C F \)" using \ H.cmp_components_are_iso C.VV.arr_char D.invert_opposite_sides_of_square by blast thus ?thesis using \ D.comp_assoc by simp qed also have "... = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.cod \))) \\<^sub>D H ((\\<^sub>0 (trg\<^sub>B \) \\<^sub>C F \) \\<^sub>C \\<^sub>1 (B.dom \)) \\<^sub>D \\<^sub>H (G (B.dom \), \\<^sub>0 (src\<^sub>B (B.dom \)))" using \ by simp also have "... = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.cod \))) \\<^sub>D H (\\<^sub>1 (B.cod \) \\<^sub>C (G \ \\<^sub>C \\<^sub>0 (src\<^sub>B \))) \\<^sub>D \\<^sub>H (G (B.dom \), \\<^sub>0 (src\<^sub>B (B.dom \)))" using \ \.naturality by simp also have "... = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.cod \))) \\<^sub>D H (\\<^sub>1 (B.cod \)) \\<^sub>D H (G \ \\<^sub>C \\<^sub>0 (src\<^sub>B \)) \\<^sub>D \\<^sub>H (G (B.dom \), \\<^sub>0 (src\<^sub>B (B.dom \)))" using \ D.comp_assoc by simp also have "... = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B \), F (B.cod \))) \\<^sub>D H (\\<^sub>1 (B.cod \)) \\<^sub>D \\<^sub>H (G (B.cod \), \\<^sub>0 (src\<^sub>B (B.cod \))) \\<^sub>D (H (G \) \\<^sub>D H (\\<^sub>0 (src\<^sub>B \)))" using \ H.\.naturality [of "(G \, \\<^sub>0 (src\<^sub>B \))"] C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def by force also have "... = map\<^sub>1 (B.cod \) \\<^sub>D ((H \ G) \ \\<^sub>D map\<^sub>0 (src\<^sub>B \))" unfolding map\<^sub>0_def map\<^sub>1_def using \ D.comp_assoc by simp finally show "map\<^sub>1 (B.cod \) \\<^sub>D ((H \ G) \ \\<^sub>D map\<^sub>0 (src\<^sub>B \)) = (map\<^sub>0 (trg\<^sub>B \) \\<^sub>D (H \ F) \) \\<^sub>D map\<^sub>1 (B.dom \)" by simp qed show "\a. B.obj a \ (map\<^sub>0 a \\<^sub>D HoF.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \\<^sub>D \\<^sub>D[map\<^sub>0 a] = map\<^sub>1 a \\<^sub>D (HoG.unit a \\<^sub>D map\<^sub>0 a)" proof - fix a assume a: "B.obj a" have "map\<^sub>1 a \\<^sub>D (HoG.unit a \\<^sub>D map\<^sub>0 a) = D.inv (\\<^sub>H (\\<^sub>0 a, F a)) \\<^sub>D H (\\<^sub>1 a) \\<^sub>D \\<^sub>H (G a, \\<^sub>0 a) \\<^sub>D (H (\.G.unit a) \\<^sub>D H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a))" unfolding map\<^sub>0_def map\<^sub>1_def using a HoG.unit_char' D.comp_assoc by auto also have "... = D.inv (\\<^sub>H (\\<^sub>0 a, F a)) \\<^sub>D H (\\<^sub>1 a) \\<^sub>D (\\<^sub>H (G a, \\<^sub>0 a) \\<^sub>D (H (\.G.unit a) \\<^sub>D H (\\<^sub>0 a))) \\<^sub>D (H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a))" using a D.whisker_right D.comp_assoc by simp also have "... = D.inv (\\<^sub>H (\\<^sub>0 a, F a)) \\<^sub>D (H (\\<^sub>1 a) \\<^sub>D H (\.G.unit a \\<^sub>C \\<^sub>0 a)) \\<^sub>D \\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a) \\<^sub>D (H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a))" using a H.\.naturality [of "(\.G.unit a, \\<^sub>0 a)"] D.comp_assoc C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def by auto also have "... = D.inv (\\<^sub>H (\\<^sub>0 a, F a)) \\<^sub>D H (\\<^sub>1 a \\<^sub>C (\.G.unit a \\<^sub>C \\<^sub>0 a)) \\<^sub>D \\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a) \\<^sub>D (H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a))" proof - have "C.arr (\\<^sub>1 a \\<^sub>C (\.G.unit a \\<^sub>C \\<^sub>0 a))" using a by force hence "H (\\<^sub>1 a) \\<^sub>D H (\.G.unit a \\<^sub>C \\<^sub>0 a) = H (\\<^sub>1 a \\<^sub>C (\.G.unit a \\<^sub>C \\<^sub>0 a))" using a by simp thus ?thesis using D.comp_assoc by simp qed also have "... = D.inv (\\<^sub>H (\\<^sub>0 a, F a)) \\<^sub>D H ((\\<^sub>0 a \\<^sub>C \.F.unit a) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>C \\<^sub>C[\\<^sub>0 a]) \\<^sub>D \\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a) \\<^sub>D (H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a))" using a \.respects_unit by simp also have "... = (D.inv (\\<^sub>H (\\<^sub>0 a, F a)) \\<^sub>D H (\\<^sub>0 a \\<^sub>C \.F.unit a)) \\<^sub>D H (\\<^sub>C\<^sup>-\<^sup>1[\\<^sub>0 a]) \\<^sub>D H (\\<^sub>C[\\<^sub>0 a]) \\<^sub>D \\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a) \\<^sub>D (H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a))" using a D.comp_assoc B.obj_simps by simp also have "... = (H (\\<^sub>0 a) \\<^sub>D H (\.F.unit a)) \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 a, \.F.map\<^sub>0 a)) \\<^sub>D H (\\<^sub>C\<^sup>-\<^sup>1[\\<^sub>0 a])) \\<^sub>D H (\\<^sub>C[\\<^sub>0 a]) \\<^sub>D \\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a) \\<^sub>D (H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a))" proof - have "D.inv (\\<^sub>H (\\<^sub>0 a, F a)) \\<^sub>D H (\\<^sub>0 a \\<^sub>C \.F.unit a) = (H (\\<^sub>0 a) \\<^sub>D H (\.F.unit a)) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 a, \.F.map\<^sub>0 a))" proof - have "D.seq (H (\\<^sub>0 a \\<^sub>C \.F.unit a)) (\\<^sub>H (\\<^sub>0 a, \.F.map\<^sub>0 a))" using a C.VV.arr_char C.VV.dom_char C.VV.cod_char C.obj_simps by auto moreover have "D.iso (\\<^sub>H (\\<^sub>0 a, F a)) \ D.iso (\\<^sub>H (\\<^sub>0 a, \.F.map\<^sub>0 a))" proof - have "C.ide (F a) \ src\<^sub>C (\\<^sub>0 a) = trg\<^sub>C (F a) \ src\<^sub>C (\\<^sub>0 a) = trg\<^sub>C (\.F.map\<^sub>0 a)" using a by auto moreover have "C.ide (\.F.map\<^sub>0 a)" proof - (* TODO: I still haven't been able to configure the simps to do this. *) have "C.obj (\.F.map\<^sub>0 a)" using a by simp thus ?thesis by auto qed ultimately show ?thesis using a H.cmp_components_are_iso B.obj_simps by auto qed ultimately show ?thesis using a H.\.naturality [of "(\\<^sub>0 a, \.F.unit a)"] C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def D.invert_opposite_sides_of_square [of "\\<^sub>H (\\<^sub>0 a, F a)" "H (\\<^sub>0 a) \\<^sub>D H (\.F.unit a)" "H (\\<^sub>0 a \\<^sub>C \.F.unit a)" "\\<^sub>H (\\<^sub>0 a, \.F.map\<^sub>0 a)"] by auto qed thus ?thesis using D.comp_assoc by simp qed also have "... = (H (\\<^sub>0 a) \\<^sub>D H (\.F.unit a)) \\<^sub>D (H (\\<^sub>0 a) \\<^sub>D H.unit (\.F.map\<^sub>0 a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (\\<^sub>0 a)] \\<^sub>D (H (\\<^sub>C[\\<^sub>0 a]) \\<^sub>D \\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a)) \\<^sub>D (H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a))" proof - have "D.inv (\\<^sub>H (\\<^sub>0 a, \.F.map\<^sub>0 a)) \\<^sub>D H (\\<^sub>C\<^sup>-\<^sup>1[\\<^sub>0 a]) = (H (\\<^sub>0 a) \\<^sub>D H.unit (\.F.map\<^sub>0 a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (\\<^sub>0 a)]" using a H.preserves_runit(2) [of "\\<^sub>0 a"] D.comp_assoc by (metis C.ideD(1) C.runit'_simps(1) C.src.preserves_ide C.trg_src D.invert_side_of_triangle(1) \.F.map\<^sub>0_def H.cmp_components_are_iso H.preserves_reflects_arr \.ide_map\<^sub>0_obj \.map\<^sub>0_simps(2)) thus ?thesis using D.comp_assoc by simp qed also have "... = (H (\\<^sub>0 a) \\<^sub>D H (\.F.unit a)) \\<^sub>D (H (\\<^sub>0 a) \\<^sub>D H.unit (\.F.map\<^sub>0 a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (\\<^sub>0 a)] \\<^sub>D \\<^sub>D[H (\\<^sub>0 a)] \\<^sub>D (D.inv (H.unit (\.G.map\<^sub>0 a)) \\<^sub>D H (\\<^sub>0 a)) \\<^sub>D (H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a))" proof - have "H (\\<^sub>C[\\<^sub>0 a]) \\<^sub>D \\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a) = \\<^sub>D[H (\\<^sub>0 a)] \\<^sub>D (D.inv (H.unit (\.G.map\<^sub>0 a)) \\<^sub>D H (\\<^sub>0 a)) \\<^sub>D D.inv (\\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a)) \\<^sub>D \\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a)" using a H.preserves_lunit(1) D.comp_assoc by auto also have "... = \\<^sub>D[H (\\<^sub>0 a)] \\<^sub>D (D.inv (H.unit (\.G.map\<^sub>0 a)) \\<^sub>D H (\\<^sub>0 a))" proof - have "D.inv (\\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a)) \\<^sub>D \\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a) = H (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a)" using a H.cmp_components_are_iso D.comp_inv_arr' by (metis C.isomorphic_implies_hpar(1) \.G.map\<^sub>0_simps(2) \.G.weakly_preserves_objects H.cmp_simps(4) \.ide_map\<^sub>0_obj \.map\<^sub>0_simps(3)) moreover have "(D.inv (H.unit (\.G.map\<^sub>0 a)) \\<^sub>D H (\\<^sub>0 a)) \\<^sub>D (H (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a)) = (D.inv (H.unit (\.G.map\<^sub>0 a)) \\<^sub>D H (\\<^sub>0 a))" using a H.unit_char(2) D.comp_arr_dom by (metis D.arr_inv D.dom_inv D.whisker_right \.G.map\<^sub>0_simps(1) H.unit_simps(5) H.preserves_ide \.ide_map\<^sub>0_obj) ultimately show ?thesis by simp qed finally have "H (\\<^sub>C[\\<^sub>0 a]) \\<^sub>D \\<^sub>H (\.G.map\<^sub>0 a, \\<^sub>0 a) = \\<^sub>D[H (\\<^sub>0 a)] \\<^sub>D (D.inv (H.unit (\.G.map\<^sub>0 a)) \\<^sub>D H (\\<^sub>0 a))" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = ((H (\\<^sub>0 a) \\<^sub>D H (\.F.unit a)) \\<^sub>D (H (\\<^sub>0 a) \\<^sub>D H.unit (\.F.map\<^sub>0 a))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (\\<^sub>0 a)] \\<^sub>D \\<^sub>D[H (\\<^sub>0 a)]" proof - have "(D.inv (H.unit (\.G.map\<^sub>0 a)) \\<^sub>D H (\\<^sub>0 a)) \\<^sub>D (H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a)) = D.inv (H.unit (\.G.map\<^sub>0 a)) \\<^sub>D H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a)" using a D.whisker_right H.unit_char(2) by simp also have "... = H.map\<^sub>0 (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a)" using a H.unit_char(1-2) [of "\.G.map\<^sub>0 a"] D.comp_inv_arr' by simp finally have "(D.inv (H.unit (\.G.map\<^sub>0 a)) \\<^sub>D H (\\<^sub>0 a)) \\<^sub>D (H.unit (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a)) = H.map\<^sub>0 (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a)" by blast moreover have "\\<^sub>D[H (\\<^sub>0 a)] \\<^sub>D (H.map\<^sub>0 (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a)) = \\<^sub>D[H (\\<^sub>0 a)]" using a D.comp_arr_dom [of "\\<^sub>D[H (\\<^sub>0 a)]" "H.map\<^sub>0 (\.G.map\<^sub>0 a) \\<^sub>D H (\\<^sub>0 a)"] by auto ultimately show ?thesis using D.comp_assoc by simp qed also have "... = (H (\\<^sub>0 a) \\<^sub>D HoF.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (\\<^sub>0 a)] \\<^sub>D \\<^sub>D[H (\\<^sub>0 a)]" using a HoF.unit_char' D.whisker_left [of "H (\\<^sub>0 a)"] by simp also have "... = (map\<^sub>0 a \\<^sub>D HoF.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \\<^sub>D \\<^sub>D[map\<^sub>0 a]" unfolding map\<^sub>0_def by simp finally show "(map\<^sub>0 a \\<^sub>D HoF.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 a] \\<^sub>D \\<^sub>D[map\<^sub>0 a] = map\<^sub>1 a \\<^sub>D (HoG.unit a \\<^sub>D map\<^sub>0 a)" by simp qed show "\f g. \B.ide f; B.ide g; src\<^sub>B g = trg\<^sub>B f\ \ (map\<^sub>0 (trg\<^sub>B g) \\<^sub>D HoF.cmp (g, f)) \\<^sub>D \\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (H \ F) g, (H \ F) f] \\<^sub>D (map\<^sub>1 g \\<^sub>D (H \ F) f) \\<^sub>D D.inv \\<^sub>D[(H \ G) g, map\<^sub>0 (src\<^sub>B g), (H \ F) f] \\<^sub>D ((H \ G) g \\<^sub>D map\<^sub>1 f) \\<^sub>D \\<^sub>D[(H \ G) g, (H \ G) f, map\<^sub>0 (src\<^sub>B f)] = map\<^sub>1 (g \\<^sub>B f) \\<^sub>D (HoG.cmp (g, f) \\<^sub>D map\<^sub>0 (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" have "map\<^sub>1 (g \\<^sub>B f) \\<^sub>D (HoG.cmp (g, f) \\<^sub>D map\<^sub>0 (src\<^sub>B f)) = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F (g \\<^sub>B f))) \\<^sub>D H (\\<^sub>1 (g \\<^sub>B f)) \\<^sub>D \\<^sub>H (G (g \\<^sub>B f), \\<^sub>0 (src\<^sub>B f)) \\<^sub>D (H (G (g \\<^sub>B f)) \\<^sub>D H (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" unfolding map\<^sub>0_def map\<^sub>1_def HoG.cmp_def using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char D.comp_assoc by simp also have "... = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F (g \\<^sub>B f))) \\<^sub>D H (\\<^sub>1 (g \\<^sub>B f)) \\<^sub>D \\<^sub>H (G (g \\<^sub>B f), \\<^sub>0 (src\<^sub>B f)) \\<^sub>D ((H (G (g \\<^sub>B f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))) \\<^sub>D (H (\\<^sub>G (g, f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char C.VV.arr_char C.VV.dom_char C.VV.cod_char \.G.FF_def D.comp_assoc D.whisker_right by auto also have "... = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F (g \\<^sub>B f))) \\<^sub>D H (\\<^sub>1 (g \\<^sub>B f)) \\<^sub>D (\\<^sub>H (G (g \\<^sub>B f), \\<^sub>0 (src\<^sub>B f)) \\<^sub>D (H (\\<^sub>G (g, f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" proof - have "(H (G (g \\<^sub>B f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))) \\<^sub>D (H (\\<^sub>G (g, f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))) = H (G (g \\<^sub>B f)) \\<^sub>D H (\\<^sub>G (g, f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))" using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char D.whisker_right by simp also have "... = H (\\<^sub>G (g, f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))" using f g fg D.comp_cod_arr [of "H (\\<^sub>G (g, f))" "H (G (g \\<^sub>B f))"] B.VV.arr_char B.VV.dom_char B.VV.cod_char by simp finally have "(H (G (g \\<^sub>B f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))) \\<^sub>D (H (\\<^sub>G (g, f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))) = H (\\<^sub>G (g, f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F (g \\<^sub>B f))) \\<^sub>D (H (\\<^sub>1 (g \\<^sub>B f)) \\<^sub>D H (\\<^sub>G (g, f) \\<^sub>C \\<^sub>0 (src\<^sub>B f))) \\<^sub>D \\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f)) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" using f g fg H.\.naturality [of "(\\<^sub>G (g, f), \\<^sub>0 (src\<^sub>B f))"] B.VV.arr_char B.VV.dom_char B.VV.cod_char \.G.FF_def C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def D.comp_assoc by simp also have "... = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F (g \\<^sub>B f))) \\<^sub>D H (\\<^sub>1 (g \\<^sub>B f) \\<^sub>C (\\<^sub>G (g, f) \\<^sub>C \\<^sub>0 (src\<^sub>B f))) \\<^sub>D \\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f)) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char D.comp_assoc by simp also have "... = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F (g \\<^sub>B f))) \\<^sub>D H ((\\<^sub>0 (trg\<^sub>B g) \\<^sub>C \\<^sub>F (g, f)) \\<^sub>C \\<^sub>C[\\<^sub>0 (trg\<^sub>B g), F g, F f] \\<^sub>C (\\<^sub>1 g \\<^sub>C F f) \\<^sub>C \\<^sub>C\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>B g), F f] \\<^sub>C (G g \\<^sub>C \\<^sub>1 f) \\<^sub>C \\<^sub>C[G g, G f, \\<^sub>0 (src\<^sub>B f)]) \\<^sub>D \\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f)) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" using f g fg \.respects_hcomp by simp also have "... = (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F (g \\<^sub>B f))) \\<^sub>D H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C \\<^sub>F (g, f))) \\<^sub>D H \\<^sub>C[\\<^sub>0 (trg\<^sub>B g), F g, F f] \\<^sub>D H (\\<^sub>1 g \\<^sub>C F f) \\<^sub>D H (\\<^sub>C\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>B g), F f]) \\<^sub>D H (G g \\<^sub>C \\<^sub>1 f) \\<^sub>D H \\<^sub>C[G g, G f, \\<^sub>0 (src\<^sub>B f)] \\<^sub>D \\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f)) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char \.F.FF_def D.comp_assoc by simp also have "... = (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f))) \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g \\<^sub>C F f)) \\<^sub>D H \\<^sub>C[\\<^sub>0 (trg\<^sub>B g), F g, F f]) \\<^sub>D H (\\<^sub>1 g \\<^sub>C F f) \\<^sub>D H (\\<^sub>C\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>B g), F f]) \\<^sub>D H (G g \\<^sub>C \\<^sub>1 f) \\<^sub>D H \\<^sub>C[G g, G f, \\<^sub>0 (src\<^sub>B f)] \\<^sub>D \\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f)) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" proof - have "D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F (g \\<^sub>B f))) \\<^sub>D H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C \\<^sub>F (g, f)) = (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f))) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g \\<^sub>C F f))" proof - have "\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F (g \\<^sub>B f)) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f))) = H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C \\<^sub>F (g, f)) \\<^sub>D \\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g \\<^sub>C F f)" using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char C.VV.arr_char C.VV.dom_char C.VV.cod_char \.F.FF_def H.FF_def H.\.naturality [of "(\\<^sub>0 (trg\<^sub>B g), \\<^sub>F (g, f))"] by auto moreover have "D.seq (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F (g \\<^sub>B f))) (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f)))" using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def by auto moreover have "D.iso (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F (g \\<^sub>B f))) \ D.iso (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g \\<^sub>C F f))" using f g fg H.cmp_components_are_iso by simp ultimately show ?thesis using f g fg D.invert_opposite_sides_of_square by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f))) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f)) \\<^sub>D \\<^sub>D[H (\\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (F f)) \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C F g, F f)) \\<^sub>D H (\\<^sub>1 g \\<^sub>C F f)) \\<^sub>D H (\\<^sub>C\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>B g), F f]) \\<^sub>D H (G g \\<^sub>C \\<^sub>1 f) \\<^sub>D H \\<^sub>C[G g, G f, \\<^sub>0 (src\<^sub>B f)] \\<^sub>D \\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f)) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" proof - have "D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g \\<^sub>C F f)) \\<^sub>D H \\<^sub>C[\\<^sub>0 (trg\<^sub>B g), F g, F f] = (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f)) \\<^sub>D \\<^sub>D[H (\\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (F f)) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C F g, F f))" proof - have "D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g \\<^sub>C F f)) \\<^sub>D H \\<^sub>C[\\<^sub>0 (trg\<^sub>B g), F g, F f] = ((D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g \\<^sub>C F f)) \\<^sub>D \\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g \\<^sub>C F f)) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f))) \\<^sub>D \\<^sub>D[H (\\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (F f)) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C F g, F f))" using f g fg H.preserves_assoc(1) D.comp_assoc by simp also have "... = ((H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (F g \\<^sub>C F f)) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f))) \\<^sub>D \\<^sub>D[H (\\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (F f)) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C F g, F f))" using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char H.FF_def D.comp_assoc D.comp_inv_arr' H.cmp_components_are_iso by simp also have "... = (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f)) \\<^sub>D \\<^sub>D[H (\\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (F f)) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C F g, F f))" using f g fg C.VV.arr_char H.cmp_simps(5) D.comp_cod_arr D.comp_assoc by auto finally show ?thesis by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f))) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f)) \\<^sub>D \\<^sub>D[H (\\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (F f)) \\<^sub>D (H (\\<^sub>1 g) \\<^sub>D H (F f)) \\<^sub>D (D.inv (\\<^sub>H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g), F f)) \\<^sub>D H (\\<^sub>C\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>B g), F f])) \\<^sub>D H (G g \\<^sub>C \\<^sub>1 f) \\<^sub>D H \\<^sub>C[G g, G f, \\<^sub>0 (src\<^sub>B f)] \\<^sub>D \\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f)) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" proof - have "D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C F g, F f)) \\<^sub>D H (\\<^sub>1 g \\<^sub>C F f) = (H (\\<^sub>1 g) \\<^sub>D H (F f)) \\<^sub>D D.inv (\\<^sub>H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g), F f))" proof - have "\\<^sub>H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C F g, F f) \\<^sub>D (H (\\<^sub>1 g) \\<^sub>D H (F f)) = H (\\<^sub>1 g \\<^sub>C F f) \\<^sub>D \\<^sub>H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g), F f)" using f g fg H.\.naturality [of "(\\<^sub>1 g, F f)"] C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def by simp moreover have "D.seq (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C F g, F f)) (H (\\<^sub>1 g) \\<^sub>D H (F f))" using f g fg by fastforce moreover have "D.iso (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g) \\<^sub>C F g, F f)) \ D.iso (\\<^sub>H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g), F f))" using f g fg H.cmp_components_are_iso by simp ultimately show ?thesis using f g fg D.invert_opposite_sides_of_square by metis qed thus ?thesis using D.comp_assoc by simp qed also have "... = (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f))) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f)) \\<^sub>D \\<^sub>D[H (\\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (F f)) \\<^sub>D (H (\\<^sub>1 g) \\<^sub>D H (F f)) \\<^sub>D (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\\<^sub>0 (src\<^sub>B g)), H (F f)] \\<^sub>D (H (G g) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (src\<^sub>B g), F f))) \\<^sub>D (D.inv (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g) \\<^sub>C F f)) \\<^sub>D H (G g \\<^sub>C \\<^sub>1 f)) \\<^sub>D H \\<^sub>C[G g, G f, \\<^sub>0 (src\<^sub>B f)] \\<^sub>D \\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f)) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" proof - have "D.inv (\\<^sub>H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g), F f)) \\<^sub>D H (\\<^sub>C\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>B g), F f]) = (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\\<^sub>0 (src\<^sub>B g)), H (F f)] \\<^sub>D (H (G g) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (src\<^sub>B g), F f))) \\<^sub>D D.inv (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g) \\<^sub>C F f))" proof - have "D.inv (\\<^sub>H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g), F f)) \\<^sub>D H (\\<^sub>C\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>B g), F f]) = ((D.inv (\\<^sub>H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g), F f)) \\<^sub>D \\<^sub>H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g), F f)) \\<^sub>D (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\\<^sub>0 (src\<^sub>B g)), H (F f)] \\<^sub>D (H (G g) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (src\<^sub>B g), F f))) \\<^sub>D D.inv (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g) \\<^sub>C F f))" using f g fg H.preserves_assoc(2) D.comp_assoc by simp also have "... = (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\\<^sub>0 (src\<^sub>B g)), H (F f)] \\<^sub>D (H (G g) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (src\<^sub>B g), F f))) \\<^sub>D D.inv (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g) \\<^sub>C F f))" proof - have "D.inv (\\<^sub>H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g), F f)) \\<^sub>D \\<^sub>H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g), F f) = H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f)" using f g fg H.cmp_components_are_iso D.comp_inv_arr' H.FF_def C.VV.arr_char C.VV.dom_char C.VV.cod_char by simp moreover have "(H (G g \\<^sub>C \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f)) \\<^sub>D (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f)) = (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f))" using f g fg D.comp_cod_arr C.VV.arr_char C.VV.dom_char C.VV.cod_char by simp ultimately show ?thesis by simp qed finally show ?thesis by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f))) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f)) \\<^sub>D \\<^sub>D[H (\\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (F f)) \\<^sub>D (H (\\<^sub>1 g) \\<^sub>D H (F f)) \\<^sub>D (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\\<^sub>0 (src\<^sub>B g)), H (F f)] \\<^sub>D (H (G g) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (src\<^sub>B g), F f))) \\<^sub>D (H (G g) \\<^sub>D H (\\<^sub>1 f)) \\<^sub>D (D.inv (\\<^sub>H (G g, G f \\<^sub>C \\<^sub>0 (src\<^sub>B f))) \\<^sub>D H \\<^sub>C[G g, G f, \\<^sub>0 (src\<^sub>B f)]) \\<^sub>D \\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f)) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" proof - have "D.inv (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g) \\<^sub>C F f)) \\<^sub>D H (G g \\<^sub>C \\<^sub>1 f) = (H (G g) \\<^sub>D H (\\<^sub>1 f)) \\<^sub>D D.inv (\\<^sub>H (G g, G f \\<^sub>C \\<^sub>0 (src\<^sub>B f)))" proof - have "\\<^sub>H (G g, \\<^sub>0 (trg\<^sub>B f) \\<^sub>C F f) \\<^sub>D (H (G g) \\<^sub>D H (\\<^sub>1 f)) = H (G g \\<^sub>C \\<^sub>1 f) \\<^sub>D \\<^sub>H (G g, G f \\<^sub>C \\<^sub>0 (src\<^sub>B f))" using f g fg H.\.naturality [of "(G g, \\<^sub>1 f)"] C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def by simp thus ?thesis using f g fg H.cmp_components_are_iso C.VV.arr_char C.VV.dom_char C.VV.cod_char D.invert_opposite_sides_of_square [of "\\<^sub>H (G g, \\<^sub>0 (trg\<^sub>B f) \\<^sub>C F f)" "H (G g) \\<^sub>D H (\\<^sub>1 f)" "H (G g \\<^sub>C \\<^sub>1 f)" "\\<^sub>H (G g, G f \\<^sub>C \\<^sub>0 (src\<^sub>B f))"] by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f))) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f)) \\<^sub>D \\<^sub>D[H (\\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (F f)) \\<^sub>D (H (\\<^sub>1 g) \\<^sub>D H (F f)) \\<^sub>D (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\\<^sub>0 (src\<^sub>B g)), H (F f)] \\<^sub>D (H (G g) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (src\<^sub>B g), F f))) \\<^sub>D (H (G g) \\<^sub>D H (\\<^sub>1 f)) \\<^sub>D (H (G g) \\<^sub>D \\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f))) \\<^sub>D \\<^sub>D[H (G g), H (G f), H (\\<^sub>0 (src\<^sub>B f))] \\<^sub>D (D.inv (\\<^sub>H (G g, G f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))) \\<^sub>D (D.inv (\\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f))) \\<^sub>D \\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f))) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f)))" proof - have "D.inv (\\<^sub>H (G g, G f \\<^sub>C \\<^sub>0 (src\<^sub>B f))) \\<^sub>D H \\<^sub>C[G g, G f, \\<^sub>0 (src\<^sub>B f)] = (H (G g) \\<^sub>D \\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f))) \\<^sub>D \\<^sub>D[H (G g), H (G f), H (\\<^sub>0 (src\<^sub>B f))] \\<^sub>D (D.inv (\\<^sub>H (G g, G f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))) \\<^sub>D D.inv (\\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f)))" proof - have "D.inv (\\<^sub>H (G g, G f \\<^sub>C \\<^sub>0 (src\<^sub>B f))) \\<^sub>D H \\<^sub>C[G g, G f, \\<^sub>0 (src\<^sub>B f)] = ((D.inv (\\<^sub>H (G g, G f \\<^sub>C \\<^sub>0 (src\<^sub>B f))) \\<^sub>D \\<^sub>H (G g, G f \\<^sub>C \\<^sub>0 (src\<^sub>B f))) \\<^sub>D (H (G g) \\<^sub>D \\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f)))) \\<^sub>D \\<^sub>D[H (G g), H (G f), H (\\<^sub>0 (src\<^sub>B f))] \\<^sub>D (D.inv (\\<^sub>H (G g, G f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))) \\<^sub>D D.inv (\\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f)))" using f g fg H.preserves_assoc(1) H.cmp_components_are_iso C.VV.arr_char C.VV.dom_char C.VV.cod_char D.comp_assoc by simp moreover have "D.inv (\\<^sub>H (G g, G f \\<^sub>C \\<^sub>0 (src\<^sub>B f))) \\<^sub>D \\<^sub>H (G g, G f \\<^sub>C \\<^sub>0 (src\<^sub>B f)) = H (G g) \\<^sub>D H (G f \\<^sub>C \\<^sub>0 (src\<^sub>B f))" using f g fg H.cmp_components_are_iso C.VV.arr_char C.VV.dom_char C.VV.cod_char D.comp_inv_arr' H.FF_def by simp moreover have "(H (G g) \\<^sub>D H (G f \\<^sub>C \\<^sub>0 (src\<^sub>B f))) \\<^sub>D (H (G g) \\<^sub>D \\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f))) = H (G g) \\<^sub>D \\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f))" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char D.whisker_left D.comp_cod_arr by simp ultimately show ?thesis by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = ((H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f))) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f))) \\<^sub>D \\<^sub>D[H (\\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \\<^sub>D ((D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (F f)) \\<^sub>D (H (\\<^sub>1 g) \\<^sub>D H (F f)) \\<^sub>D (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\\<^sub>0 (src\<^sub>B g)), H (F f)] \\<^sub>D ((H (G g) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (src\<^sub>B g), F f))) \\<^sub>D (H (G g) \\<^sub>D H (\\<^sub>1 f)) \\<^sub>D (H (G g) \\<^sub>D \\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f)))) \\<^sub>D \\<^sub>D[H (G g), H (G f), H (\\<^sub>0 (src\<^sub>B f))]" proof - have "\\<^sub>D[H (G g), H (G f), H (\\<^sub>0 (src\<^sub>B f))] \\<^sub>D (D.inv (\\<^sub>H (G g, G f)) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))) \\<^sub>D (D.inv (\\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f))) \\<^sub>D \\<^sub>H (G g \\<^sub>C G f, \\<^sub>0 (src\<^sub>B f))) \\<^sub>D (\\<^sub>H (G g, G f) \\<^sub>D H (\\<^sub>0 (src\<^sub>B f))) = \\<^sub>D[H (G g), H (G f), H (\\<^sub>0 (src\<^sub>B f))]" using f g fg D.comp_arr_dom D.comp_cod_arr D.whisker_right [of "H (\\<^sub>0 (src\<^sub>B f))" "D.inv (\\<^sub>H (G g, G f))" "\\<^sub>H (G g, G f)"] D.comp_inv_arr' C.VV.arr_char C.VV.dom_char C.VV.cod_char H.FF_def H.cmp_components_are_iso by simp thus ?thesis using D.comp_assoc by simp qed also have "... = (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (F (g \\<^sub>B f)) \\<^sub>D H (\\<^sub>F (g, f)) \\<^sub>D \\<^sub>H (F g, F f)) \\<^sub>D \\<^sub>D[H (\\<^sub>0 (trg\<^sub>B g)), H (F g), H (F f)] \\<^sub>D (D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (\\<^sub>1 g) \\<^sub>D \\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[H (G g), H (\\<^sub>0 (src\<^sub>B g)), H (F f)] \\<^sub>D (H (G g) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (src\<^sub>B g), F f)) \\<^sub>D H (\\<^sub>1 f) \\<^sub>D \\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f))) \\<^sub>D \\<^sub>D[H (G g), H (G f), H (\\<^sub>0 (src\<^sub>B f))]" proof - have "(H (G g) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (src\<^sub>B g), F f))) \\<^sub>D (H (G g) \\<^sub>D H (\\<^sub>1 f)) \\<^sub>D (H (G g) \\<^sub>D \\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f))) = H (G g) \\<^sub>D D.inv (\\<^sub>H (\\<^sub>0 (src\<^sub>B g), F f)) \\<^sub>D H (\\<^sub>1 f) \\<^sub>D \\<^sub>H (G f, \\<^sub>0 (src\<^sub>B f))" using f g fg C.VV.arr_char C.VV.dom_char C.VV.cod_char D.whisker_left H.cmp_components_are_iso by simp moreover have "(H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f))) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f)) = H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (F (g \\<^sub>B f)) \\<^sub>D H (\\<^sub>F (g, f)) \\<^sub>D \\<^sub>H (F g, F f)" proof - have "(H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f))) \\<^sub>D (H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D \\<^sub>H (F g, F f)) = H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (\\<^sub>F (g, f)) \\<^sub>D \\<^sub>H (F g, F f)" proof - have "D.arr (H (\\<^sub>F (g, f)) \\<^sub>D \\<^sub>H (F g, F f))" using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char \.F.FF_def by (intro D.seqI) auto thus ?thesis using f g fg D.whisker_left C.VV.arr_char by simp qed also have "... = H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D (H (F (g \\<^sub>B f) \\<^sub>C \\<^sub>F (g, f))) \\<^sub>D \\<^sub>H (F g, F f)" using f g fg C.comp_cod_arr B.VV.arr_char \.F.cmp_simps(5) by auto also have "... = H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D (H (F (g \\<^sub>B f)) \\<^sub>D H (\\<^sub>F (g, f))) \\<^sub>D \\<^sub>H (F g, F f)" using f g fg B.VV.arr_char \.F.cmp_simps(5) by auto also have "... = H (\\<^sub>0 (trg\<^sub>B g)) \\<^sub>D H (F (g \\<^sub>B f)) \\<^sub>D H (\\<^sub>F (g, f)) \\<^sub>D \\<^sub>H (F g, F f)" using D.comp_assoc by simp finally show ?thesis by simp qed moreover have "(D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (F f)) \\<^sub>D (H (\\<^sub>1 g) \\<^sub>D H (F f)) \\<^sub>D (\\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f)) = D.inv (\\<^sub>H (\\<^sub>0 (trg\<^sub>B g), F g)) \\<^sub>D H (\\<^sub>1 g) \\<^sub>D \\<^sub>H (G g, \\<^sub>0 (src\<^sub>B g)) \\<^sub>D H (F f)" using f g fg D.whisker_right C.VV.arr_char C.VV.dom_char C.VV.cod_char by simp ultimately show ?thesis by simp qed also have "... = (map\<^sub>0 (trg\<^sub>B g) \\<^sub>D HoF.cmp (g, f)) \\<^sub>D \\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (H \ F) g, (H \ F) f] \\<^sub>D (map\<^sub>1 g \\<^sub>D (H \ F) f) \\<^sub>D D.inv \\<^sub>D[(H \ G) g, map\<^sub>0 (src\<^sub>B g), (H \ F) f] \\<^sub>D ((H \ G) g \\<^sub>D map\<^sub>1 f) \\<^sub>D \\<^sub>D[(H \ G) g, (H \ G) f, map\<^sub>0 (src\<^sub>B f)]" unfolding map\<^sub>0_def map\<^sub>1_def HoF.cmp_def using f g fg B.VV.arr_char B.VV.dom_char B.VV.cod_char by simp finally show "(map\<^sub>0 (trg\<^sub>B g) \\<^sub>D HoF.cmp (g, f)) \\<^sub>D \\<^sub>D[map\<^sub>0 (trg\<^sub>B g), (H \ F) g, (H \ F) f] \\<^sub>D (map\<^sub>1 g \\<^sub>D (H \ F) f) \\<^sub>D D.inv \\<^sub>D[(H \ G) g, map\<^sub>0 (src\<^sub>B g), (H \ F) f] \\<^sub>D ((H \ G) g \\<^sub>D map\<^sub>1 f) \\<^sub>D \\<^sub>D[(H \ G) g, (H \ G) f, map\<^sub>0 (src\<^sub>B f)] = map\<^sub>1 (g \\<^sub>B f) \\<^sub>D (HoG.cmp (g, f) \\<^sub>D map\<^sub>0 (src\<^sub>B f))" by simp qed qed lemma is_pseudonatural_transformation: shows "pseudonatural_transformation 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 HoF.map HoF.cmp HoG.map HoG.cmp map\<^sub>0 map\<^sub>1" .. end subsection "Pseudonatural Equivalences" text \ A \emph{pseudonatural equivalence} is a pseudonatural transformation whose components at objects are equivalence maps. Pseudonatural equivalences between pseudofunctors generalize natural isomorphisms between ordinary functors. \ locale pseudonatural_equivalence = pseudonatural_transformation + assumes components_are_equivalences: "C.obj a \ D.equivalence_map (\\<^sub>0 a)" subsubsection "Identity Transformations are Pseudonatural Equivalences" sublocale identity_pseudonatural_transformation \ 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 F \\<^sub>F F \\<^sub>F map\<^sub>0 map\<^sub>1 by unfold_locales (simp add: D.obj_is_equivalence_map) subsubsection "Composition of Pseudonatural Equivalences" locale composite_pseudonatural_equivalence = composite_pseudonatural_transformation + \: 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 F \\<^sub>F G \\<^sub>G \\<^sub>0 \\<^sub>1 + \: 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 H \\<^sub>H \\<^sub>0 \\<^sub>1 begin sublocale 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 F \\<^sub>F H \\<^sub>H map\<^sub>0 map\<^sub>1 apply unfold_locales by (metis D.equivalence_maps_compose D.hcomp_in_hhomE \.components_are_equivalences \.components_are_equivalences map\<^sub>0_def map\<^sub>0_in_hom(1)) lemma is_pseudonatural_equivalence: shows "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 F \\<^sub>F H \\<^sub>H map\<^sub>0 map\<^sub>1" .. end locale pseudonatural_equivalence_whisker_right = pseudonatural_transformation_whisker_right + \: 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 F \\<^sub>F G \\<^sub>G \\<^sub>0 \\<^sub>1 begin interpretation FoH: 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 H \\<^sub>H F \\<^sub>F .. interpretation GoH: 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 H \\<^sub>H G \\<^sub>G .. sublocale pseudonatural_equivalence 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 \F o H\ FoH.cmp \G o H\ GoH.cmp map\<^sub>0 map\<^sub>1 using map\<^sub>0_def \.components_are_equivalences by unfold_locales simp lemma is_pseudonatural_equivalence: shows "pseudonatural_equivalence 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 (F o H) FoH.cmp (G o H) GoH.cmp map\<^sub>0 map\<^sub>1" .. end locale pseudonatural_equivalence_whisker_left = pseudonatural_transformation_whisker_left + \: pseudonatural_equivalence 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 begin interpretation HoF: 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 F \\<^sub>F H \\<^sub>H .. interpretation HoG: 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 H \\<^sub>H .. sublocale pseudonatural_equivalence 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 \H o F\ HoF.cmp \H o G\ HoG.cmp map\<^sub>0 map\<^sub>1 using map\<^sub>0_def \.components_are_equivalences H.preserves_equivalence_maps by unfold_locales simp lemma is_pseudonatural_equivalence: shows "pseudonatural_equivalence 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 (H o F) HoF.cmp (H o G) HoG.cmp map\<^sub>0 map\<^sub>1" .. end subsubsection "Converse of a Pseudonatural Equivalence" text \ It is easy to see that natural isomorphism between ordinary functors is a symmetric relation because a unique inverse to a natural isomorphism is obtained merely by inverting the components. However the situation is more difficult for pseudonatural equivalences because they do not have unique inverses. Instead, we have to choose a quasi-inverse for each of the components. In order to satisfy the required coherence conditions, it is necessary for these quasi-inverses to be part of chosen adjoint equivalences. Some long calculations to establish the coherence conditions seem unavoidable. The purpose of this section is to carry out the construction, given a pseudonatural equivalence, of a ``converse'' pseudonatural equivalence in the opposite direction. \ locale converse_pseudonatural_equivalence = 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>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: 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 + \: pseudonatural_equivalence begin abbreviation (input) F\<^sub>0 where "F\<^sub>0 \ F.map\<^sub>0" abbreviation (input) G\<^sub>0 where "G\<^sub>0 \ G.map\<^sub>0" definition map\<^sub>0 where "map\<^sub>0 a = (SOME g. \\ \. adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (\\<^sub>0 a) g \ \)" abbreviation (input) \\<^sub>0' where "\\<^sub>0' \ map\<^sub>0" definition unit where "unit a = (SOME \. \\. adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (\\<^sub>0 a) (\\<^sub>0' a) \ \)" abbreviation (input) \ where "\ \ unit" definition counit where "counit a = (SOME \. adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (\\<^sub>0 a) (\\<^sub>0' a) (\ a) \)" abbreviation (input) \ where "\ \ counit" lemma chosen_adjoint_equivalence: assumes "C.obj a" shows "adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (\\<^sub>0 a) (\\<^sub>0' a) (\ a) (\ a)" using assms \.components_are_equivalences map\<^sub>0_def unit_def counit_def D.obj_is_equivalence_map D.equivalence_map_extends_to_adjoint_equivalence someI_ex [of "\g. \\ \. adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (\\<^sub>0 a) g \ \"] someI_ex [of "\\. \\. adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (\\<^sub>0 a) (\\<^sub>0' a) \ \"] someI_ex [of "\\. adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (\\<^sub>0 a) (\\<^sub>0' a) (\ a) \"] by simp lemma map\<^sub>0_in_hhom [intro]: assumes "C.obj a" shows "\\\<^sub>0' a : G\<^sub>0 a \\<^sub>D F\<^sub>0 a\" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 a\ \\\<^sub>0' a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence by simp show ?thesis using assms \.map\<^sub>0_in_hhom [of a] antipar by auto qed lemma map\<^sub>0_simps [simp]: assumes "C.obj a" shows "D.ide (\\<^sub>0' a)" and "src\<^sub>D (\\<^sub>0' a) = G\<^sub>0 a" and "trg\<^sub>D (\\<^sub>0' a) = F\<^sub>0 a" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 a\ \\\<^sub>0' a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence by simp show "D.ide (\\<^sub>0' a)" by simp show "src\<^sub>D (\\<^sub>0' a) = G\<^sub>0 a" using assms map\<^sub>0_in_hhom by blast show "trg\<^sub>D (\\<^sub>0' a) = F\<^sub>0 a" using assms map\<^sub>0_in_hhom by blast qed lemma equivalence_map_map\<^sub>0 [simp]: assumes "C.obj a" shows "D.equivalence_map (\\<^sub>0' a)" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 a\ \\\<^sub>0' a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence by simp show "D.equivalence_map (\\<^sub>0' a)" using D.equivalence_map_def adjoint_equivalence_in_bicategory_axioms dual_equivalence by blast qed lemma unit_in_hom [intro]: assumes "C.obj a" shows "\\ a : F\<^sub>0 a \\<^sub>D F\<^sub>0 a\" and "\\ a : F\<^sub>0 a \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a\" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 a\ \\\<^sub>0' a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence by simp show "\\ a : F\<^sub>0 a \\<^sub>D F\<^sub>0 a\" using assms unit_in_hom antipar(1) by auto show "\\ a : F\<^sub>0 a \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a\" using assms unit_in_hom antipar(1) by auto qed lemma unit_simps [simp]: assumes "C.obj a" shows "D.iso (\ a)" and "D.arr (\ a)" and "src\<^sub>D (\ a) = F\<^sub>0 a" and "trg\<^sub>D (\ a) = F\<^sub>0 a" and "D.dom (\ a) = F\<^sub>0 a" and "D.cod (\ a) = \\<^sub>0' a \\<^sub>D \\<^sub>0 a" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 a\ \\\<^sub>0' a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence by simp show "D.iso (\ a)" and "D.arr (\ a)" and "src\<^sub>D (\ a) = F\<^sub>0 a" and "trg\<^sub>D (\ a) = F\<^sub>0 a" and "D.dom (\ a) = F\<^sub>0 a" and "D.cod (\ a) = \\<^sub>0' a \\<^sub>D \\<^sub>0 a" using assms by auto qed lemma iso_unit: assumes "C.obj a" shows "D.iso (\ a)" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 a\ \\\<^sub>0' a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence by simp show ?thesis by simp qed lemma counit_in_hom [intro]: assumes "C.obj a" shows "\\ a : G\<^sub>0 a \\<^sub>D G\<^sub>0 a\" and "\\ a : \\<^sub>0 a \\<^sub>D \\<^sub>0' a \\<^sub>D G\<^sub>0 a\" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 a\ \\\<^sub>0' a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence by simp show "\\ a : G\<^sub>0 a \\<^sub>D G\<^sub>0 a\" using assms counit_in_hom antipar(2) by auto show " \\ a : \\<^sub>0 a \\<^sub>D \\<^sub>0' a \\<^sub>D G\<^sub>0 a\" using assms counit_in_hom antipar(2) by simp qed lemma counit_simps [simp]: assumes "C.obj a" shows "D.iso (\ a)" and "D.arr (\ a)" and "src\<^sub>D (\ a) = G\<^sub>0 a" and "trg\<^sub>D (\ a) = G\<^sub>0 a" and "D.dom (\ a) = \\<^sub>0 a \\<^sub>D \\<^sub>0' a" and "D.cod (\ a) = G\<^sub>0 a" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 a\ \\\<^sub>0' a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence by simp show "D.iso (\ a)" and "D.arr (\ a)" and "src\<^sub>D (\ a) = G\<^sub>0 a" and "trg\<^sub>D (\ a) = G\<^sub>0 a" and "D.dom (\ a) = \\<^sub>0 a \\<^sub>D \\<^sub>0' a" and "D.cod (\ a) = G\<^sub>0 a" using assms by auto qed lemma iso_counit: assumes "C.obj a" shows "D.iso (\ a)" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 a\ \\\<^sub>0' a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence by simp show ?thesis by simp qed lemma quasi_inverts_components: assumes "C.obj a" shows "D.isomorphic (\\<^sub>0' a \\<^sub>D \\<^sub>0 a) (F\<^sub>0 a)" and "D.isomorphic (\\<^sub>0 a \\<^sub>D \\<^sub>0' a) (G\<^sub>0 a)" and "D.quasi_inverses (\\<^sub>0 a) (\\<^sub>0' a)" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 a\ \\\<^sub>0' a\ \\ a\ \\ a\ using assms chosen_adjoint_equivalence by simp show "D.isomorphic (\\<^sub>0' a \\<^sub>D \\<^sub>0 a) (F\<^sub>0 a)" using assms D.isomorphic_def D.isomorphic_symmetric unit_is_iso by blast show "D.isomorphic (\\<^sub>0 a \\<^sub>D \\<^sub>0' a) (G\<^sub>0 a)" using assms D.isomorphic_def counit_is_iso by blast show "D.quasi_inverses (\\<^sub>0 a) (\\<^sub>0' a)" using D.quasi_inverses_def equivalence_in_bicategory_axioms by auto qed definition map\<^sub>1 where "map\<^sub>1 f = (\\<^sub>0' (trg\<^sub>C f) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>0' (trg\<^sub>C f) \\<^sub>D G f \\<^sub>D \ (src\<^sub>C f)) \\<^sub>D (\\<^sub>0' (trg\<^sub>C f) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0' (src\<^sub>C f)]) \\<^sub>D \\<^sub>D[\\<^sub>0' (trg\<^sub>C f), G f \\<^sub>D \\<^sub>0 (src\<^sub>C f), \\<^sub>0' (src\<^sub>C f)] \\<^sub>D ((\\<^sub>0' (trg\<^sub>C f) \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' (src\<^sub>C f)) \\<^sub>D (\\<^sub>D[\\<^sub>0' (trg\<^sub>C f), \\<^sub>0 (trg\<^sub>C f), F f] \\<^sub>D \\<^sub>0' (src\<^sub>C f)) \\<^sub>D ((\ (trg\<^sub>C f) \\<^sub>D F f) \\<^sub>D \\<^sub>0' (src\<^sub>C f)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' (src\<^sub>C f))" abbreviation (input) \\<^sub>1' where "\\<^sub>1' \ map\<^sub>1" lemma map\<^sub>1_in_hom [intro]: assumes "C.ide f" shows "\\\<^sub>1' f : G\<^sub>0 (src\<^sub>C f) \\<^sub>D F\<^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 G f\" proof - show "\\\<^sub>1' f : F f \\<^sub>D \\<^sub>0' (src\<^sub>C f) \\<^sub>D \\<^sub>0' (trg\<^sub>C f) \\<^sub>D G f\" using assms \.iso_map\<^sub>1_ide \.map\<^sub>1_in_vhom by (unfold map\<^sub>1_def, intro D.comp_in_homI) auto thus "\\\<^sub>1' f : G\<^sub>0 (src\<^sub>C f) \\<^sub>D F\<^sub>0 (trg\<^sub>C f)\" using assms D.vconn_implies_hpar(1-2) by auto qed lemma map\<^sub>1_simps [simp]: assumes "C.ide f" shows "D.arr (\\<^sub>1' f)" and "src\<^sub>D (\\<^sub>1' f) = G\<^sub>0 (src\<^sub>C f)" and "trg\<^sub>D (\\<^sub>1' f) = F\<^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 G f" using assms map\<^sub>1_in_hom by auto lemma iso_map\<^sub>1_ide: assumes "C.ide f" shows "D.iso (\\<^sub>1' f)" proof - interpret src: adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 (src\<^sub>C f)\ \\\<^sub>0' (src\<^sub>C f)\ \\ (src\<^sub>C f)\ \\ (src\<^sub>C f)\ using assms chosen_adjoint_equivalence by simp interpret trg: adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 (trg\<^sub>C f)\ \\\<^sub>0' (trg\<^sub>C f)\ \\ (trg\<^sub>C f)\ \\ (trg\<^sub>C f)\ using assms chosen_adjoint_equivalence by simp show ?thesis unfolding map\<^sub>1_def using assms \.iso_map\<^sub>1_ide by (intro D.isos_compose) auto qed interpretation EV: self_evaluation_map V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D .. notation EV.eval ("\_\") sublocale 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 F \\<^sub>F \\<^sub>0' \\<^sub>1' proof show "\a. C.obj a \ D.ide (\\<^sub>0' a)" using map\<^sub>0_simps(1) by simp show "\f. C.ide f \ D.iso (\\<^sub>1' f)" using iso_map\<^sub>1_ide by simp show "\a. C.obj a \ \\\<^sub>0' a : src\<^sub>D (G a) \\<^sub>D src\<^sub>D (F a)\" by fastforce 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 G f\" by auto show "\a. C.obj a \ D.equivalence_map (\\<^sub>0' a)" 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 G \) \\<^sub>D \\<^sub>1' (C.dom \)" proof - fix \ assume \: "C.arr \" let ?a = "src\<^sub>C \" let ?b = "trg\<^sub>C \" let ?f = "C.dom \" let ?g = "C.cod \" have "\\<^sub>1' (C.cod \) \\<^sub>D (F \ \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G ?g \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G ?g \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?g)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?g] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F ?g) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F ?g] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F \ \\<^sub>D \\<^sub>0' ?a)" unfolding map\<^sub>1_def using \ D.comp_assoc by simp also have "... = (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G ?g \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G ?g \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?g)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?g] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\ ?b \\<^sub>D F ?g) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((trg\<^sub>D (F \) \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F ?f] \\<^sub>D \\<^sub>0' ?a)" using \ D.whisker_right D.lunit'_naturality [of "F \"] D.comp_assoc D.whisker_right [of "\\<^sub>0' ?a" "trg\<^sub>D (F \) \\<^sub>D F \" "\\<^sub>D\<^sup>-\<^sup>1[F ?f]"] by simp also have "... = (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G ?g \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G ?g \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?g)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?g] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\ ?b \\<^sub>D F ?f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F ?f] \\<^sub>D \\<^sub>0' ?a)" proof - have "((\ ?b \\<^sub>D F ?g) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((trg\<^sub>D (F \) \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a) = (((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F ?f) \\<^sub>D \\<^sub>0' ?a)" proof - have "((\ ?b \\<^sub>D F ?g) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((trg\<^sub>D (F \) \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a) = (\ ?b \\<^sub>D F ?g) \\<^sub>D (trg\<^sub>D (F \) \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a" using \ D.whisker_right D.obj_simps by simp also have "... = (\ ?b \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a" using \ D.interchange [of "\ ?b" "trg\<^sub>D (F \)" "F ?g" "F \"] D.comp_arr_dom D.comp_cod_arr by simp also have "... = ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F \) \\<^sub>D (\ ?b \\<^sub>D F ?f) \\<^sub>D \\<^sub>0' ?a" using \ D.interchange [of "\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b" "\ ?b" "F \" "F ?f"] D.comp_arr_dom D.comp_cod_arr by simp also have "... = (((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F ?f) \\<^sub>D \\<^sub>0' ?a)" using \ D.whisker_right by simp finally show ?thesis by blast qed thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G ?g \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G ?g \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?g)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F ?f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F ?f] \\<^sub>D \\<^sub>0' ?a)" proof - have "(\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?g] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a" using \ D.whisker_right by simp also have "... = (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F \) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?f] \\<^sub>D \\<^sub>0' ?a" using \ D.assoc_naturality [of "\\<^sub>0' ?b" "\\<^sub>0 ?b" "F \"] by simp also have "... = ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?f] \\<^sub>D \\<^sub>0' ?a)" using \ D.whisker_right by simp finally have "(\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?g] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a) = ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?f] \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G ?g \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, G ?g \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D (G \ \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F ?f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F ?f] \\<^sub>D \\<^sub>0' ?a)" proof - have "((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?g)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?g)) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a" using \ \.iso_map\<^sub>1_ide D.whisker_right by simp also have "... = (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F \)) \\<^sub>D \\<^sub>0' ?a" using \ \.iso_map\<^sub>1_ide D.whisker_left by simp also have "... = (\\<^sub>0' ?b \\<^sub>D (G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D D.inv (\\<^sub>1 ?f)) \\<^sub>D \\<^sub>0' ?a" using \ \.iso_map\<^sub>1_ide \.naturality D.invert_opposite_sides_of_square [of "\\<^sub>1 ?g" "G \ \\<^sub>D \\<^sub>0 ?a" "\\<^sub>0 ?b \\<^sub>D F \" "\\<^sub>1 ?f"] by simp also have "... = (\\<^sub>0' ?b \\<^sub>D (G \ \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?f)) \\<^sub>D \\<^sub>0' ?a" using \ \.iso_map\<^sub>1_ide D.whisker_left by simp also have "... = ((\\<^sub>0' ?b \\<^sub>D (G \ \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?f)) \\<^sub>D \\<^sub>0' ?a)" using \ \.iso_map\<^sub>1_ide D.whisker_right by simp finally have "((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?g)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F \) \\<^sub>D \\<^sub>0' ?a) = ((\\<^sub>0' ?b \\<^sub>D (G \ \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?f)) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G ?g \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G ?f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F ?f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F ?f] \\<^sub>D \\<^sub>0' ?a)" using \ D.assoc_naturality [of "\\<^sub>0' ?b" "G \ \\<^sub>D \\<^sub>0 ?a" "\\<^sub>0' ?a"] D.comp_assoc by simp also have "... = (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g]) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D G ?g \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G ?f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F ?f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F ?f] \\<^sub>D \\<^sub>0' ?a)" proof - have "(\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) = \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g, \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" using \ D.whisker_left by simp also have "... = \\<^sub>0' ?b \\<^sub>D (G \ \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[G ?f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using \ D.assoc_naturality [of "G \" "\\<^sub>0 ?a" "\\<^sub>0' ?a"] by simp also have "... = (\\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?f, \\<^sub>0 ?a, \\<^sub>0' ?a])" using \ D.whisker_left by simp finally have "(\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?f, \\<^sub>0 ?a, \\<^sub>0' ?a])" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D G\<^sub>0 ?a)) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G ?f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G ?f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F ?f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F ?f] \\<^sub>D \\<^sub>0' ?a)" proof - have "(\\<^sub>0' ?b \\<^sub>D G ?g \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) = \\<^sub>0' ?b \\<^sub>D (G ?g \\<^sub>D \ ?a) \\<^sub>D (G \ \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" using \ D.whisker_left by simp also have "... = \\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D \ ?a" using \ D.interchange [of "G ?g" "G \" "\ ?a" "\\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a"] D.comp_arr_dom D.comp_cod_arr by simp also have "... = \\<^sub>0' ?b \\<^sub>D (G \ \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (G ?f \\<^sub>D \ ?a)" using \ D.interchange [of "G \" "G ?f" "G\<^sub>0 ?a" "\ ?a"] D.comp_arr_dom D.comp_cod_arr by simp also have "... = (\\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G ?f \\<^sub>D \ ?a)" using \ D.whisker_left D.obj_simps by auto finally have "(\\<^sub>0' ?b \\<^sub>D G ?g \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G ?f \\<^sub>D \ ?a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?b \\<^sub>D G \) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?f]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G ?f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G ?f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 ?f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F ?f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F ?f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F ?f] \\<^sub>D \\<^sub>0' ?a)" proof - have "(\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D G\<^sub>0 ?a) = \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g] \\<^sub>D (G \ \\<^sub>D G\<^sub>0 ?a)" using \ D.whisker_left D.obj_simps by simp also have "... = \\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D \\<^sub>D[G ?f]" using \ D.runit_naturality [of "G \"] by simp also have "... = (\\<^sub>0' ?b \\<^sub>D G \) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?f])" using \ D.whisker_left by simp finally have "(\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?g]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G \ \\<^sub>D G\<^sub>0 ?a) = (\\<^sub>0' ?b \\<^sub>D G \) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G ?f])" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?b \\<^sub>D G \) \\<^sub>D \\<^sub>1' ?f" unfolding map\<^sub>1_def using \ D.comp_assoc by simp finally show "\\<^sub>1' ?g \\<^sub>D (F \ \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?b \\<^sub>D G \) \\<^sub>D \\<^sub>1' ?f" by blast qed show "\a. C.obj a \ (\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a] \\<^sub>D \\<^sub>D[\\<^sub>0' a] = \\<^sub>1' a \\<^sub>D (F.unit a \\<^sub>D \\<^sub>0' a)" proof - fix a assume "C.obj a" hence a: "C.obj a \ C.arr a \ C.ide a \ src\<^sub>C a = a \ trg\<^sub>C a = a" by auto have 1: "D.ide (F\<^sub>0 a)" using a F.map\<^sub>0_simps(1) by blast have 2: "D.ide (G\<^sub>0 a)" using a G.map\<^sub>0_simps(1) by blast have "\\<^sub>1' a \\<^sub>D (F.unit a \\<^sub>D \\<^sub>0' a) = (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D \\<^sub>D[\\<^sub>0' a, G a \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D ((\\<^sub>0' a \\<^sub>D D.inv (\\<^sub>1 a)) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F a] \\<^sub>D \\<^sub>0' a) \\<^sub>D (F.unit a \\<^sub>D \\<^sub>0' a)" unfolding map\<^sub>1_def using a D.comp_assoc by simp also have "... = (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D \\<^sub>D[\\<^sub>0' a, G a \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D ((\\<^sub>0' a \\<^sub>D D.inv (\\<^sub>1 a)) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F a] \\<^sub>D \\<^sub>0' a) \\<^sub>D (((\ a \\<^sub>D F a) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((F\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F a] \\<^sub>D \\<^sub>0' a) \\<^sub>D (F.unit a \\<^sub>D \\<^sub>0' a) = \\<^sub>D\<^sup>-\<^sup>1[F a] \\<^sub>D F.unit a \\<^sub>D \\<^sub>0' a" using a D.whisker_right by simp also have "... = (F\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a" using a D.lunit'_naturality [of "F.unit a"] by simp also have "... = ((F\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" using a 1 D.whisker_right by simp finally have "(\\<^sub>D\<^sup>-\<^sup>1[F a] \\<^sub>D \\<^sub>0' a) \\<^sub>D (F.unit a \\<^sub>D \\<^sub>0' a) = ((F\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D \\<^sub>D[\\<^sub>0' a, G a \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D ((\\<^sub>0' a \\<^sub>D D.inv (\\<^sub>1 a)) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F a] \\<^sub>D \\<^sub>0' a) \\<^sub>D (((\\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a)) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" proof - have "((\ a \\<^sub>D F a) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((F\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) = (\ a \\<^sub>D F a) \\<^sub>D (F\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a" using a 1 D.whisker_right by simp also have "... = (\ a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a" using a D.interchange [of "\ a" "F\<^sub>0 a" "F a" "F.unit a"] D.comp_cod_arr D.comp_arr_dom by simp also have "... = ((\\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D F.unit a) \\<^sub>D (\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a" using a D.interchange [of "\\<^sub>0' a \\<^sub>D \\<^sub>0 a" "\ a" "F.unit a" "F\<^sub>0 a"] D.comp_cod_arr D.comp_arr_dom by simp also have "... = (((\\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a)" using a 1 D.whisker_right by simp finally have "((\ a \\<^sub>D F a) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((F\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) = (((\\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D \\<^sub>D[\\<^sub>0' a, G a \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (((\\<^sub>0' a \\<^sub>D D.inv (\\<^sub>1 a)) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a)) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" proof - have "(\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F a] \\<^sub>D \\<^sub>0' a) \\<^sub>D (((\\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) = \\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F a] \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a" using a D.whisker_right by simp also have "... = (\\<^sub>0' a \\<^sub>D \\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F\<^sub>0 a] \\<^sub>D \\<^sub>0' a" using a D.assoc_naturality [of "\\<^sub>0' a" "\\<^sub>0 a" "F.unit a"] by simp also have "... = ((\\<^sub>0' a \\<^sub>D \\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" using a 1 D.whisker_right by simp finally have "(\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F a] \\<^sub>D \\<^sub>0' a) \\<^sub>D (((\\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) = ((\\<^sub>0' a \\<^sub>D \\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, G a \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D ((\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0' a)) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" proof - have "((\\<^sub>0' a \\<^sub>D D.inv (\\<^sub>1 a)) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) = (\\<^sub>0' a \\<^sub>D D.inv (\\<^sub>1 a)) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a" using a \.iso_map\<^sub>1_ide D.whisker_right by simp also have "... = (\\<^sub>0' a \\<^sub>D D.inv (\\<^sub>1 a) \\<^sub>D (\\<^sub>0 a \\<^sub>D F.unit a)) \\<^sub>D \\<^sub>0' a" using a \.iso_map\<^sub>1_ide D.whisker_left by simp also have "... = (\\<^sub>0' a \\<^sub>D (G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a" proof - have "(\\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] = \\<^sub>1 a \\<^sub>D (G.unit a \\<^sub>D \\<^sub>0 a)" using a \.respects_unit by simp hence "D.inv (\\<^sub>1 a) \\<^sub>D (\\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a] = G.unit a \\<^sub>D \\<^sub>0 a" using a \.iso_map\<^sub>1_ide D.invert_side_of_triangle(1) by simp hence "D.inv (\\<^sub>1 a) \\<^sub>D (\\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] = (G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]" using a D.iso_lunit D.invert_side_of_triangle(2) D.comp_assoc by simp hence "D.inv (\\<^sub>1 a) \\<^sub>D (\\<^sub>0 a \\<^sub>D F.unit a) = (G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a]" using a D.iso_runit D.comp_assoc D.invert_side_of_triangle(2) [of "(G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]" "D.inv (\\<^sub>1 a) \\<^sub>D (\\<^sub>0 a \\<^sub>D F.unit a)" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]"] by simp thus ?thesis by simp qed also have "... = (\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a" using a D.whisker_left by simp also have "... = ((\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a)" using a D.whisker_right by simp finally have "((\\<^sub>0' a \\<^sub>D D.inv (\\<^sub>1 a)) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>0 a \\<^sub>D F.unit a) \\<^sub>D \\<^sub>0' a) = ((\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G a \\<^sub>D \ a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D[G a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D (G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0' a)) \\<^sub>D \\<^sub>D[\\<^sub>0' a, G\<^sub>0 a \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" using a D.assoc_naturality [of "\\<^sub>0' a" "G.unit a \\<^sub>D \\<^sub>0 a" "\\<^sub>0' a"] D.comp_assoc by simp also have "... = (\\<^sub>0' a \\<^sub>D \\<^sub>D[G a]) \\<^sub>D ((\\<^sub>0' a \\<^sub>D G a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a)) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D \\<^sub>D[\\<^sub>0' a, G\<^sub>0 a \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" proof - have "(\\<^sub>0' a \\<^sub>D \\<^sub>D[G a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D (G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0' a) = \\<^sub>0' a \\<^sub>D \\<^sub>D[G a, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D ((G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0' a)" using a D.whisker_left by simp also have "... = \\<^sub>0' a \\<^sub>D (G.unit a \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[G\<^sub>0 a, \\<^sub>0 a, \\<^sub>0' a]" using a D.assoc_naturality [of "G.unit a" "\\<^sub>0 a" "\\<^sub>0' a"] by simp also have "... = (\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a, \\<^sub>0 a, \\<^sub>0' a])" using a 2 D.whisker_left [of "\\<^sub>0' a"] by simp finally have "(\\<^sub>0' a \\<^sub>D \\<^sub>D[G a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D (G.unit a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0' a) = (\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a, \\<^sub>0 a, \\<^sub>0' a])" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>0' a \\<^sub>D \\<^sub>D[G a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D G\<^sub>0 a)) \\<^sub>D (\\<^sub>0' a \\<^sub>D G\<^sub>0 a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D \\<^sub>D[\\<^sub>0' a, G\<^sub>0 a \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" proof - have "(\\<^sub>0' a \\<^sub>D G a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a) = \\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D \ a" using a D.whisker_left D.comp_arr_dom D.comp_cod_arr D.interchange [of "G a" "G.unit a" "\ a" "\\<^sub>0 a \\<^sub>D \\<^sub>0' a"] by simp also have "... = (\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D G\<^sub>0 a) \\<^sub>D (\\<^sub>0' a \\<^sub>D G\<^sub>0 a \\<^sub>D \ a)" using a 2 D.whisker_left D.comp_arr_dom D.comp_cod_arr D.interchange [of "G.unit a" "G\<^sub>0 a" "G\<^sub>0 a" "\ a"] by simp finally have "(\\<^sub>0' a \\<^sub>D G a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a) = (\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D G\<^sub>0 a) \\<^sub>D (\\<^sub>0' a \\<^sub>D G\<^sub>0 a \\<^sub>D \ a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G\<^sub>0 a \\<^sub>D \ a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D \\<^sub>D[\\<^sub>0' a, G\<^sub>0 a \\<^sub>D \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D (((\\<^sub>0' a \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" proof - have "(\\<^sub>0' a \\<^sub>D \\<^sub>D[G a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D G\<^sub>0 a) = \\<^sub>0' a \\<^sub>D \\<^sub>D[G a] \\<^sub>D (G.unit a \\<^sub>D G\<^sub>0 a)" using a 2 D.whisker_left by simp also have "... = \\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D \\<^sub>D[G\<^sub>0 a]" using a D.runit_naturality [of "G.unit a"] by simp also have "... = (\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a])" using a 2 D.whisker_left by simp finally have "(\\<^sub>0' a \\<^sub>D \\<^sub>D[G a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G.unit a \\<^sub>D G\<^sub>0 a) = (\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a])" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G\<^sub>0 a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D (\\<^sub>D[\\<^sub>0' a, G\<^sub>0 a \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]) \\<^sub>D \\<^sub>0' a)) \\<^sub>D (\\<^sub>D[\\<^sub>0' a \\<^sub>D \\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" using a 1 D.whisker_right D.runit_hcomp(1) D.comp_assoc by simp also have "... = (\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D G\<^sub>0 a \\<^sub>D \ a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>0' a)) \\<^sub>D \\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (\\<^sub>D[\\<^sub>0' a \\<^sub>D \\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" using a D.assoc_naturality [of "\\<^sub>0' a" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a]" "\\<^sub>0' a"] D.comp_assoc by simp also have "... = (\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a]) \\<^sub>D ((\\<^sub>0' a \\<^sub>D G\<^sub>0 a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a \\<^sub>D \\<^sub>0' a])) \\<^sub>D \\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (\\<^sub>D[\\<^sub>0' a \\<^sub>D \\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" using a 2 D.whisker_left D.lunit_hcomp(4) D.comp_assoc by simp also have "... = (\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D \ a) \\<^sub>D \\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D ((\\<^sub>D[\\<^sub>0' a \\<^sub>D \\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" proof - have "(\\<^sub>0' a \\<^sub>D G\<^sub>0 a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a \\<^sub>D \\<^sub>0' a]) = \\<^sub>0' a \\<^sub>D (G\<^sub>0 a \\<^sub>D \ a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a \\<^sub>D \\<^sub>0' a]" using a 2 D.whisker_left by simp also have "... = \\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a] \\<^sub>D \ a" using a D.lunit'_naturality [of "\ a"] by simp also have "... = (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D \ a)" using a 2 D.whisker_left by simp finally have "(\\<^sub>0' a \\<^sub>D G\<^sub>0 a \\<^sub>D \ a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a \\<^sub>D \\<^sub>0' a]) = (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D \ a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \ a) \\<^sub>D \\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (\ a \\<^sub>D \\<^sub>0' a)) \\<^sub>D (\\<^sub>D[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" proof - have "(\\<^sub>D[\\<^sub>0' a \\<^sub>D \\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) = \\<^sub>D[\\<^sub>0' a \\<^sub>D \\<^sub>0 a] \\<^sub>D (\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a" using a 1 D.whisker_right by simp also have "... = \ a \\<^sub>D \\<^sub>D[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a" using a D.runit_naturality [of "\ a"] by simp also have "... = (\ a \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" using a 1 D.whisker_right by simp finally have "(\\<^sub>D[\\<^sub>0' a \\<^sub>D \\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\ a \\<^sub>D F\<^sub>0 a) \\<^sub>D \\<^sub>0' a) = (\ a \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D ((\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a]) \\<^sub>D \\<^sub>D[\\<^sub>0' a] \\<^sub>D (\\<^sub>D[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a)" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 a\ \\\<^sub>0' a\ \\ a\ \\ a\ using a chosen_adjoint_equivalence by simp have "(\\<^sub>0' a \\<^sub>D \ a) \\<^sub>D \\<^sub>D[\\<^sub>0' a, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (\ a \\<^sub>D \\<^sub>0' a) = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a] \\<^sub>D \\<^sub>D[\\<^sub>0' a]" using triangle_right by simp thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a] \\<^sub>D \\<^sub>D[\\<^sub>0' a]" proof - have "\\<^sub>D[\\<^sub>0' a] \\<^sub>D (\\<^sub>D[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a] \\<^sub>D \\<^sub>0' a) = \\<^sub>D[\\<^sub>0' a]" proof - have "D.seq \\<^sub>D[F\<^sub>0 a] \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 a]" using a 1 by auto thus ?thesis using a D.unitor_coincidence [of "F\<^sub>0 a"] D.comp_arr_dom D.whisker_right by (metis D.comp_arr_inv' D.iso_runit D.lunit_simps(1,4) D.objE D.runit_simps(5) F.map\<^sub>0_simps(1) map\<^sub>0_simps(1,3)) qed moreover have "(\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a]" proof - have "(\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a] = ((\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a]) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a])) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a]" using a D.unitor_coincidence D.comp_assoc by simp also have "... = (\\<^sub>0' a \\<^sub>D \\<^sub>D[G\<^sub>0 a] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G\<^sub>0 a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a]" using a 2 D.whisker_left by simp also have "... = (\\<^sub>0' a \\<^sub>D G\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a]" using a 2 D.comp_arr_inv' by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a]" using a 2 D.comp_cod_arr by simp finally show ?thesis by blast qed ultimately show ?thesis using D.comp_assoc by simp qed finally show "(\\<^sub>0' a \\<^sub>D G.unit a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' a] \\<^sub>D \\<^sub>D[\\<^sub>0' a] = \\<^sub>1' a \\<^sub>D (F.unit a \\<^sub>D \\<^sub>0' a)" by simp 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 \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' (trg\<^sub>C g), G g, G f] \\<^sub>D (\\<^sub>1' g \\<^sub>D G f) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0' (src\<^sub>C g), G 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 (\\<^sub>F (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" let ?a = "src\<^sub>C f" let ?b = "trg\<^sub>C f" let ?c = "trg\<^sub>C g" (* * TODO: The following are extremely problematic. * I don't think all the cases are used, but they don't get discovered by auto * and sledgehammer doesn't find them, either. Note that the same issue occurred * for the previous subgoal, where I needed D.ide (F\<^sub>0 a) and D.ide (G\<^sub>0 a). *) have *: "D.ide (G\<^sub>0 ?a) \ D.ide (G\<^sub>0 ?b) \ D.ide (G\<^sub>0 ?b) \ D.ide (G\<^sub>0 ?c)" using f g C.ideD(1) G.map\<^sub>0_simps(1) by blast have **: "D.ide (F\<^sub>0 ?a) \ D.ide (F\<^sub>0 ?b) \ D.ide (F\<^sub>0 ?b) \ D.ide (F\<^sub>0 ?c)" using f g C.ideD(1) F.map\<^sub>0_simps(1) by blast have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D (\\<^sub>1' g \\<^sub>D G f) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0' (src\<^sub>C g), G f] \\<^sub>D (F g \\<^sub>D \\<^sub>1' f) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a] = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0' ?b, G f] \\<^sub>D (F g \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" unfolding map\<^sub>1_def using f g fg D.comp_assoc by simp also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F g, \\<^sub>0' ?b, G f]) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (F g \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f = ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f)" proof - have 1: "D.arr ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b))" using f g fg \.iso_map\<^sub>1_ide by auto thus ?thesis using f g fg D.whisker_right [of "G f"] by (metis D.seqE G.preserves_ide) qed moreover have "F g \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) = (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (F g \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a)" proof - have 1: "D.arr ((\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a))" using f g fg \.iso_map\<^sub>1_ide map\<^sub>1_def map\<^sub>1_simps(1) by presburger thus ?thesis using f g fg D.whisker_left [of "F g"] by (metis D.seqE F.preserves_ide) qed ultimately show ?thesis using f g fg D.comp_assoc by simp qed (* * Move \\<^sub>D\<^sup>-\<^sup>1[F g] and \\<^sub>D\<^sup>-\<^sup>1[F f] to the bottom, followed by \ ?c and D.inv (\\<^sub>1 g). * Move \ ?b down across D.inv (\\<^sub>1 g) to meet \ ?b, where they can be canceled. *) (* Move \\<^sub>D\<^sup>-\<^sup>1[F g] down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f])) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (F g \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F g, \\<^sub>0' ?b, G f] = \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f)" using f g fg D.assoc'_naturality [of "\\<^sub>D\<^sup>-\<^sup>1[F g]" "\\<^sub>0' ?b" "G f"] by simp thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D\<^sup>-\<^sup>1[F g] down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a)) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (F g \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "\\<^sub>0' ?b \\<^sub>D G f" "\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]"] by simp also have "... = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>D\<^sup>-\<^sup>1[F g]" "\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a"] by simp finally have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D\<^sup>-\<^sup>1[F g] down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (F g \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a"] by simp also have "... = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>D\<^sup>-\<^sup>1[F g]" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D\<^sup>-\<^sup>1[F g] down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D (F g \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a" "\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp also have "... = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>D\<^sup>-\<^sup>1[F g]" "\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" "\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D\<^sup>-\<^sup>1[F g] down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp also have "... = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>D\<^sup>-\<^sup>1[F g]" "\\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]" "(\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D\<^sup>-\<^sup>1[F g] down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (F g \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "(\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a" "(\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>D\<^sup>-\<^sup>1[F g]" "(\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a" "(\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D\<^sup>-\<^sup>1[F g] down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "(\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>D\<^sup>-\<^sup>1[F g]" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a" "((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D\<^sup>-\<^sup>1[F g] down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "(\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>D\<^sup>-\<^sup>1[F g]" "(\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "(F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D\<^sup>-\<^sup>1[F g] down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, G f]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (F g \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[F g]" "F g" "(F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed (* Move \ ?c down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D ((((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f)) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f])) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, G f] = \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f)" using f g fg D.assoc'_naturality [of "\ ?c \\<^sub>D F g" "\\<^sub>0' ?b" "G f"] by simp thus ?thesis using D.comp_assoc by simp qed (* Move \ ?c down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a)) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) = (\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\ ?c \\<^sub>D F g" "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>0' ?b \\<^sub>D G f" "\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\ ?c \\<^sub>D F g" "\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a"] by simp finally have "((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?c down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) = (\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\ ?c \\<^sub>D F g" "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\ ?c \\<^sub>D F g" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?c down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = (\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\ ?c \\<^sub>D F g" "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a" "\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\ ?c \\<^sub>D F g" "\\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" "\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?c down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) = (\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\ ?c \\<^sub>D F g" "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\ ?c \\<^sub>D F g" "\\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]" "(\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?c down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\ ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) = (\ ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\ ?c \\<^sub>D F g" "F\<^sub>0 ?c \\<^sub>D F g" "(\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a" "(\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\ ?c \\<^sub>D F g" "(\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a" "(\\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f)) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\ ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f)) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?c down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\ ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) = (\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\ ?c \\<^sub>D F g" "F\<^sub>0 ?c \\<^sub>D F g" "(\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\ ?c \\<^sub>D F g" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a" "((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\ ?c \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?c down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f])) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\ ?c \\<^sub>D F g) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = (\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\ ?c \\<^sub>D F g" "F\<^sub>0 ?c \\<^sub>D F g" "((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "(\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f \\<^sub>D G\<^sub>0 ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a)) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f]) = ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f \\<^sub>D G\<^sub>0 ?a]" using f g fg D.assoc'_naturality [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\\<^sub>0' ?b" "\\<^sub>D[G f]"] by simp thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f \\<^sub>D G\<^sub>0 ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D G f \\<^sub>D \ ?a) = ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a]" using f g fg D.assoc'_naturality [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\\<^sub>0' ?b" "G f \\<^sub>D \ ?a"] by simp thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a]" using f g fg D.assoc'_naturality [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\\<^sub>0' ?b" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) = ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.whisker_left by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]" using f g fg \.iso_map\<^sub>1_ide D.assoc_naturality [of "\\<^sub>0' ?b" "D.inv (\\<^sub>1 f)" "\\<^sub>0' ?a"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a])" using f g fg \.iso_map\<^sub>1_ide D.whisker_left by simp finally have "(((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a])" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f])) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) = ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" using f g fg \.iso_map\<^sub>1_ide D.assoc'_naturality [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\\<^sub>0' ?b" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a)) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) = (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b" "((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "G f" "\\<^sub>D[G f]"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[G f]" "G f \\<^sub>D G\<^sub>0 ?a"] by simp finally have "((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) = (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b" "((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D G\<^sub>0 ?a" "G f \\<^sub>D \ ?a"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D \ ?a" "G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b" "((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f])) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b" "((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a" "(\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a)) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b" "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "G f" "\\<^sub>D[G f]"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[G f]" "G f \\<^sub>D G\<^sub>0 ?a"] by simp finally have "(((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) = (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b" "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D G\<^sub>0 ?a" "G f \\<^sub>D \ ?a"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D \ ?a" "G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) = (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b" "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b" "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a" "(\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) = (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a])) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.assoc'_naturality [of "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g]" "\\<^sub>0' ?b" "(\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) = \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g]" "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g]" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]" "(\\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f)) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) = ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g]" "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "(\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g]" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a" "((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) = ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a])) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.assoc'_naturality [of "\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)" "\\<^sub>0' ?b" "(\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed (* Move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) = (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)" "\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f)) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b" "\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]" "(\\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f)) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>0' ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move D.inv (\\<^sub>1 g) down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)" "\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g" "(\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b" "\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)" "\\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a" "((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?b up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\ ?c \\<^sub>D F g) \\<^sub>D (((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g]" "\ ?c \\<^sub>D F g" "((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "(\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\ ?c \\<^sub>D F g) \\<^sub>D ((\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_cod_arr by simp also have "... = \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\ ?c \\<^sub>D F g) \\<^sub>D ((\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom by simp also have "... = (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg ** D.interchange by auto finally have "(\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?b up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f)) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D ((\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)" "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g]" "((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "(\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.interchange [of "\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)" "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g]" "(\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "(F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D ((\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move \ ?b down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f)) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f" using f g fg D.whisker_right by simp also have "... = \\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f" using f g fg D.pentagon D.invert_side_of_triangle(2) [of "\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b]" "\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b"] D.comp_assoc by simp also have "... = (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f)" using f g fg D.whisker_right by simp finally have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) = (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?b down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f)) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f) = (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f" using f g fg D.whisker_right [of "G f"] by simp also have "... = \\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f" using f g fg D.assoc_naturality [of "\\<^sub>0' ?c" "G g" "\ ?b"] by simp also have "... = (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f)" using f g fg * D.whisker_right [of "G f"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f) = (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f)" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move \ ?b down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) = \\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" using f g fg D.pentagon by simp hence "\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b)" using f g fg D.comp_assoc D.invert_side_of_triangle(1) by simp hence "\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b]" using f g fg D.comp_assoc D.invert_side_of_triangle(2) [of "\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b]" "\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b"] by simp hence "\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f" using f g fg D.whisker_right by simp hence "(\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f) = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f)" using f g fg D.whisker_right by simp thus ?thesis using D.comp_assoc by simp qed (* Move \ ?b down *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D (((\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f) = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f" using f g fg D.whisker_right by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f" using f g fg D.assoc'_naturality [of "\\<^sub>0' ?c" "G g" "\ ?b"] by simp also have "... = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f)" using f g fg * D.whisker_right by simp finally have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f) = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f)" by blast thus ?thesis using D.comp_assoc by simp qed (* There is a cancellation of associativities here. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f])) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f) = (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f" using f g fg * D.comp_arr_inv' D.whisker_right by (metis C.ideD(1) C.obj_trg D.comp_assoc_assoc'(1) D.hcomp_simps(2) D.hseqI' D.ideD(1) G.map\<^sub>0_simps(3) G.preserves_ide G.preserves_src G.preserves_trg map\<^sub>0_simps(1-2)) moreover have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) = (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f" using f g fg D.comp_cod_arr by simp ultimately have "((\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) = (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f" by simp thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D[G f] up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f])) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) = \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D \\<^sub>D[G f]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b]" "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "G f" "\\<^sub>D[G f]"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b]" "\\<^sub>D[G f]" "G f \\<^sub>D G\<^sub>0 ?a"] by simp finally have "(\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) = ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" by blast thus ?thesis using f g fg D.comp_assoc by simp qed (* Move \\<^sub>D[G f] up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f])) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[G f]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" "\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "G f" "\\<^sub>D[G f]"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b" "\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" "\\<^sub>D[G f]" "G f \\<^sub>D G\<^sub>0 ?a"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D[G f] up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a)) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) = (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D \\<^sub>D[G f]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b" "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b" "G f" "\\<^sub>D[G f]"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b" "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b" "\\<^sub>D[G f]" "G f \\<^sub>D G\<^sub>0 ?a"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f]) = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?a up. The useful effect is on the associativity part of the term. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a)) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) = \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \ ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b]" "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D G\<^sub>0 ?a" "G f \\<^sub>D \ ?a"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b]" "G f \\<^sub>D \ ?a" "G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) = ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" by simp thus ?thesis using D.comp_assoc by simp qed (* Move \ ?a up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D \ ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" "\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D G\<^sub>0 ?a" "G f \\<^sub>D \ ?a"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b" "\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" "G f \\<^sub>D \ ?a" "G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?a up and \ ?b down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) = (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \ ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b" "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D G\<^sub>0 ?a" "G f \\<^sub>D \ ?a"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b" "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b" "G f \\<^sub>D \ ?a" "G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D G\<^sub>0 ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \ ?a) = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move \ ?b down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" proof - have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) = \\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" using f g fg D.pentagon by simp hence "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] = \\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b)" using f g fg D.comp_assoc D.invert_side_of_triangle(2) [of "\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b]" "\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b"] by simp hence "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, \\<^sub>0' ?b]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a = \\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a" by simp thus ?thesis using f g fg D.whisker_right by simp qed thus ?thesis using D.comp_assoc by simp qed (* Move \ ?b down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a" using f g fg D.whisker_right by simp also have "... = \\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a" using f g fg D.assoc_naturality [of "\\<^sub>0' ?c" "G g" "\ ?b"] by simp also have "... = (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" using f g fg * D.whisker_right by simp finally have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 f) up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b" "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp also have "... = ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 f) up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Prepare to move D.inv (\\<^sub>1 f) up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b" "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b" "G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b" "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move D.inv (\\<^sub>1 f) up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b" "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a" "(\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) = ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move D.inv (\\<^sub>1 f) up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b" "\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b]" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a" "(\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) = (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move D.inv (\\<^sub>1 f) up across \ ?b. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b" "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b" "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a" "(\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) = (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \ ?b up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b" "\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)" "(\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "(F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed (* * Introduce associativities to align (\ ?b \\<^sub>D \\<^sub>0 ?b) with * (\\<^sub>0 ?b \\<^sub>D \ ?b) in the middle. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, G\<^sub>0 ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G\<^sub>0 ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\ ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, \\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a])) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \ ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, G\<^sub>0 ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G\<^sub>0 ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\ ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a = (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom by simp also have "... = ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?b\<^bold>\, (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?b\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?b\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?b\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?b\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?b\<^bold>\, (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]\" using f g fg D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by simp also have "... = \((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?b\<^bold>\) \<^bold>\ (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\" using f g fg by (intro EV.eval_eqI, auto) also have "... = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" using f g fg D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by simp finally have "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, G\<^sub>0 ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a])) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, G\<^sub>0 ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.assoc'_naturality [of "\\<^sub>0' ?c \\<^sub>D G g" "\ ?b" "(\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, G\<^sub>0 ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a])) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a])" using f g fg * D.comp_arr_dom D.comp_cod_arr D.whisker_left D.interchange [of "\ ?b" "(\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b)" "(\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg * D.comp_arr_dom D.comp_cod_arr D.whisker_left D.interchange [of "G\<^sub>0 ?b" "\ ?b" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]" "\\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, G\<^sub>0 ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G\<^sub>0 ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\ ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) = (\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\ ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]" using f g fg D.whisker_left by simp also have "... = (\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G\<^sub>0 ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\ ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg D.assoc_naturality [of "\ ?b" "\\<^sub>0 ?b" "F f \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G\<^sub>0 ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\ ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg * D.whisker_left by simp finally have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \ ?b \\<^sub>D \\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G\<^sub>0 ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\ ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed finally show ?thesis by simp qed moreover have "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a = (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, \\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \ ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a])" proof - have "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a = ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg ** D.comp_arr_dom by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \\<^sub>0' ?a])) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a])" proof - have "(\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = \(\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0, \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0, \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>])\" using f g fg ** D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by simp also have "... = \(\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 ?b\<^bold>\) \<^bold>\ (\<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\" using f g fg by (intro EV.eval_eqI, auto) also have "... = (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" using f g fg ** D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by simp finally have "(\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a])" proof - have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = \\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D ((\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]" using f g fg ** D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b" "\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b]" "(\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]"] by simp also have "... = \\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a] \\<^sub>D (\ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg ** D.assoc'_naturality [of "\ ?b" "F f" "\\<^sub>0' ?a"] by simp also have "... = (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg ** D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b]" "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]" "\ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a])) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a])" proof - have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg D.assoc'_naturality [of "\\<^sub>0' ?c \\<^sub>D G g" "\\<^sub>0 ?b" "\ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a"] by force thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, \\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \ ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a])" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) = (\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]" using f g fg ** D.whisker_left by auto also have "... = (\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, \\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0 ?b \\<^sub>D \ ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg ** D.assoc_naturality [of "\\<^sub>0 ?b" "\ ?b" "F f \\<^sub>D \\<^sub>0' ?a"] by auto also have "... = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, \\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \ ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg ** D.whisker_left by auto finally have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \ ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, \\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \ ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed finally show ?thesis by blast qed ultimately show ?thesis using D.comp_assoc by simp qed (* Cancel out all the intervening associativities. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, G\<^sub>0 ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G\<^sub>0 ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\ ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, \\<^sub>0' ?b, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \ ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, \\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) = (\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, \\<^sub>0' ?b, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b \\<^sub>D \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, \\<^sub>0' ?b] \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>0' ?b) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b, \\<^sub>0' ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[\\<^sub>0' ?b, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (\\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[\\<^sub>0 ?b, \\<^sub>0' ?b \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) = \((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?b\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?b\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?b\<^bold>\, (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\\\<^sub>0' ?b\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\) \<^bold>\ ((\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0' ?b\<^bold>\) \<^bold>\ (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\\\<^sub>0' ?b\<^bold>\, (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 ?b\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?b\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 ?b\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?b\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\, (\<^bold>\\\<^sub>0' ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 ?b\<^bold>\) \<^bold>\ \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\\\<^sub>0' ?b\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>])\" using f g fg D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by simp also have "... = \(\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\\\<^sub>0' ?b\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\" using f g fg by (intro EV.eval_eqI, auto) also have "... = (\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, \\<^sub>0' ?b, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a" using f g fg D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by simp finally show ?thesis using D.comp_assoc by simp qed thus ?thesis using D.comp_assoc by simp qed (* Hooray! We can finally cancel \ ?b with \ ?b using the triangle identity. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, G\<^sub>0 ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G\<^sub>0 ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b] \\<^sub>D \\<^sub>D[\\<^sub>0 ?b] \\<^sub>D (F f \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a])) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\ ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, \\<^sub>0' ?b, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \ ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D ((\ ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, \\<^sub>0' ?b, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0 ?b \\<^sub>D \ ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg D.whisker_left by simp also have "... = (\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\ ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, \\<^sub>0' ?b, \\<^sub>0 ?b] \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \ ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a" using f g fg D.whisker_right by simp also have "... = (\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b] \\<^sub>D \\<^sub>D[\\<^sub>0 ?b] \\<^sub>D (F f \\<^sub>D \\<^sub>0' ?a)" proof - interpret adjoint_equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\<^sub>0 ?b\ \\\<^sub>0' ?b\ \\ ?b\ \\ ?b\ using f chosen_adjoint_equivalence by simp show ?thesis using fg triangle_left by simp qed finally have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\ ?b \\<^sub>D \\<^sub>0 ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, \\<^sub>0' ?b, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D \ ?b) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b] \\<^sub>D \\<^sub>D[\\<^sub>0 ?b] \\<^sub>D (F f \\<^sub>D \\<^sub>0' ?a)" by simp thus ?thesis using D.comp_assoc by simp qed (* Simplify some more canonical arrows. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f])) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, G\<^sub>0 ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G\<^sub>0 ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b] \\<^sub>D \\<^sub>D[\\<^sub>0 ?b] \\<^sub>D (F f \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" proof - have "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g, G\<^sub>0 ?b, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G\<^sub>0 ?b, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b] \\<^sub>D \\<^sub>D[\\<^sub>0 ?b] \\<^sub>D (F f \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F\<^sub>0 ?b \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\, \<^bold>\G\<^sub>0 ?b\<^bold>\\<^sub>0, (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\G\<^sub>0 ?b\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\G\<^sub>0 ?b\<^bold>\\<^sub>0, \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\)) \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0, \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>])\" using f g fg * ** D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char D.\_ide_simp D.\_ide_simp by simp also have "... = \(\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\F f\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\, (\<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\F\<^sub>0 ?b\<^bold>\\<^sub>0 \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\)\" using f g fg by (intro EV.eval_eqI, auto) also have "... = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg * ** D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char D.\_ide_simp D.\_ide_simp by simp finally show ?thesis by simp qed thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D G f) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f]) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]" "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b" "G f" "\\<^sub>D[G f]"] by simp moreover have "(\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = \\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg * D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b]" "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b" "G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp ultimately show ?thesis using D.comp_assoc by simp qed (* * Move \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] and \\<^sub>D[F f] outside, to get rid of * G\<^sub>0 ?b and F\<^sub>0 ?b. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] = \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a)" using f g fg D.assoc_naturality [of "\\<^sub>0' ?c \\<^sub>D G g" "\\<^sub>0 ?b" "\\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b]" "\\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a" "(F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b]" "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b" "F f \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a])) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b" "\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)" "\\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a" "(F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)" "\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g" "F f \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>0 ?b) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g]" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]"] by simp also have "... = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a])" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g]" "\\<^sub>0' ?c \\<^sub>D G g" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a" "(D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]"] D.whisker_left by simp (* 12 sec *) finally have "(((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G\<^sub>0 ?b) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a])" by blast thus ?thesis using D.comp_assoc by simp qed (* Introduce associativities to achieve (D.inv (\\<^sub>1 g) \\<^sub>D F f) and (G g \\<^sub>D D.inv (\\<^sub>1 f)). *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a])) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom by simp also have "... = (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" using f g fg \.iso_map\<^sub>1_ide D.whisker_left D.comp_assoc by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_inv' by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" using f g fg \.iso_map\<^sub>1_ide D.comp_cod_arr by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_inv_arr' by simp finally show ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a])) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.assoc'_naturality [of "\\<^sub>0' ?c" "G g" "D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) = \\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]" using f g fg \.iso_map\<^sub>1_ide D.whisker_left by simp also have "... = \\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.assoc_naturality [of "G g" "D.inv (\\<^sub>1 f)" "\\<^sub>0' ?a"] by simp also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.whisker_left by simp finally have "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed finally show ?thesis by simp qed moreover have "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.comp_arr_dom by auto also have "... = (((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a]" using f g fg D.whisker_left D.comp_assoc by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (\\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a]" using f g fg D.comp_arr_inv' by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a]" using f g fg D.comp_cod_arr by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_inv_arr' by simp finally show ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a])) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g)) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.assoc'_naturality [of "\\<^sub>0' ?c" "D.inv (\\<^sub>1 g)" "F f \\<^sub>D \\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) = \\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]" using f g fg \.iso_map\<^sub>1_ide D.whisker_left by simp also have "... = \\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a] \\<^sub>D ((D.inv (\\<^sub>1 g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.assoc_naturality [of "D.inv (\\<^sub>1 g)" "F f" "\\<^sub>0' ?a"] by simp also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.whisker_left by simp finally have "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed finally show ?thesis by simp qed ultimately show ?thesis using D.comp_assoc by simp qed (* Cancel intervening associativities. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = \\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a" proof - have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, (\\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?b, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D G g, \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, \\<^sub>0 ?b] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g \\<^sub>D \\<^sub>0 ?b, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D \\<^sub>0 ?b, F f, \\<^sub>0' ?a]) = \(\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\G g\<^bold>\, (\<^bold>\\\<^sub>0 ?b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\G g\<^bold>\,\<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\\<^bold>] \<^bold>\ \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\G g\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\G g\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>])\" using f g fg D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by simp also have "... = \\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\G g\<^bold>\, \<^bold>\\\<^sub>0 ?b\<^bold>\, \<^bold>\F f\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\" using f g fg by (intro EV.eval_eqI, auto) also have "... = \\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a" using f g fg D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by simp finally show ?thesis by blast qed thus ?thesis using D.comp_assoc by simp qed (* Apply "respects composition". *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = \\<^sub>0' ?c \\<^sub>D ((G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((D.inv (\\<^sub>1 g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.whisker_left by simp also have "... = \\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" using f g fg \.iso_map\<^sub>1_ide D.whisker_right by simp also have "... = \\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f]) \\<^sub>D \\<^sub>0' ?a" proof - have 1: "(\\<^sub>1 g \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b, F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f) = D.inv ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f]) \\<^sub>D \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a]" proof - have "((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D (\\<^sub>1 g \\<^sub>D F f) \\<^sub>D D.inv \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f)) \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] = \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a)" using f g fg \.respects_hcomp D.comp_assoc by simp hence "((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f]) \\<^sub>D (\\<^sub>1 g \\<^sub>D F f) \\<^sub>D D.inv \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f) = \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a]" using f g fg D.comp_assoc D.invert_side_of_triangle(2) [of "\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a)" "(\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D (\\<^sub>1 g \\<^sub>D F f) \\<^sub>D D.inv \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f)" "\\<^sub>D[G g, G f, \\<^sub>0 ?a]"] by fastforce moreover have "D.iso ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f])" using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_simp F.cmp_components_are_iso F.FF_def by (intro D.isos_compose D.seqI D.hseqI') auto moreover have "D.seq (\\<^sub>1 (g \\<^sub>C f)) ((\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a])" using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_simp G.FF_def by (intro D.seqI) auto ultimately show ?thesis using f g fg \.iso_map\<^sub>1_ide D.invert_side_of_triangle(1) [of "\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a]" "(\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f]" "(\\<^sub>1 g \\<^sub>D F f) \\<^sub>D D.inv \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f)"] by simp qed have "(G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f) = D.inv ((\\<^sub>1 g \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b, F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f))" proof - have "D.inv ((\\<^sub>1 g \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b, F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f)) = D.inv (\\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b, F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f)) \\<^sub>D D.inv (\\<^sub>1 g \\<^sub>D F f)" proof - have "D.iso ((\\<^sub>1 g \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 ?b, F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f))" using f g fg \.iso_map\<^sub>1_ide by (intro D.isos_compose) auto moreover have "D.iso (\\<^sub>1 g \\<^sub>D F f)" using f g fg \.iso_map\<^sub>1_ide by auto ultimately show ?thesis using \.iso_map\<^sub>1_ide D.inv_comp_left by simp qed also have "... = ((G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, F f]) \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f)" proof - have "D.iso (\\<^sub>D\<^sup>-\<^sup>1[G g, \\<^sub>0 (src\<^sub>C g), F f] \\<^sub>D (G g \\<^sub>D \\<^sub>1 f))" using f g fg \.iso_map\<^sub>1_ide by (intro D.isos_compose) auto thus ?thesis using f g fg \.iso_map\<^sub>1_ide D.inv_comp_left by simp qed finally show ?thesis using D.comp_assoc by simp qed also have "... = D.inv (D.inv ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f]) \\<^sub>D \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a])" using 1 by simp also have "... = D.inv (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a]) \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f])" proof - have 2: "D.iso ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f])" using f g fg \.iso_map\<^sub>1_ide C.VV.arr_char C.VV.dom_simp F.cmp_components_are_iso F.FF_def by (intro D.isos_compose) auto moreover have "D.iso (D.inv ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f]) \\<^sub>D \\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a])" using 2 f g fg \.iso_map\<^sub>1_ide C.VV.arr_char C.VV.dom_simp C.VV.cod_simp F.FF_def G.FF_def D.inv_comp_left by (intro D.isos_compose) auto (* 10 sec *) ultimately show ?thesis using D.inv_comp_left [of "D.inv ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f])" "\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a]"] D.inv_inv D.iso_inv_iso by metis qed also have "... = \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f])" proof - have "D.inv (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a]) = \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f))" proof - have 2: "D.iso (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a])" using f g fg \.iso_map\<^sub>1_ide C.VV.arr_char C.VV.dom_simp C.VV.cod_simp G.FF_def by (intro D.isos_compose D.seqI D.hseqI') auto have "D.inv (\\<^sub>1 (g \\<^sub>C f) \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a]) = D.inv ((\\<^sub>G (g, f) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g, G f, \\<^sub>0 ?a]) \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f))" using 2 f g fg * \.iso_map\<^sub>1_ide G.cmp_components_are_iso C.VV.arr_char C.VV.dom_simp C.VV.cod_simp G.FF_def D.inv_comp_left by simp also have "... = (\\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a)) \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f))" using 2 f g fg * \.iso_map\<^sub>1_ide G.cmp_components_are_iso C.VV.arr_char C.VV.dom_simp C.VV.cod_simp G.FF_def D.inv_comp_left by simp also have "... = \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f))" using D.comp_assoc by simp finally show ?thesis using D.comp_assoc by simp qed thus ?thesis using D.comp_assoc by simp qed finally have "(G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f) = \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D (\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f]" by blast thus ?thesis by simp qed also have "... = \\<^sub>0' ?c \\<^sub>D (\\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.whisker_right C.VV.arr_char C.VV.dom_simp C.VV.cod_simp G.cmp_components_are_iso F.FF_def G.FF_def by simp also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.whisker_left C.VV.arr_char C.VV.dom_simp C.VV.cod_simp G.cmp_components_are_iso F.FF_def G.FF_def by simp finally have "(\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, \\<^sub>0 ?b, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>1 g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D[F f] down, where it will cancel. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g]" "\\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a" "(F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] D.whisker_right by simp also have "... = (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g]" "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "F f \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a"] D.whisker_right by simp finally have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D[F f] down. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = (\ ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g" "\ ?c \\<^sub>D F g" "\\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a" "(F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_arr_dom D.comp_cod_arr D.interchange [of "\ ?c \\<^sub>D F g" "F\<^sub>0 ?c \\<^sub>D F g" "F f \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a"] by simp finally have "(((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D (F\<^sub>0 ?b \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Cancel \\<^sub>D[F f] and \\<^sub>D\<^sup>-\<^sup>1[F f]. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D (\\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a)" using f g fg D.comp_cod_arr D.interchange [of "F\<^sub>0 ?c \\<^sub>D F g" "\\<^sub>D\<^sup>-\<^sup>1[F g]" "\\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a" "\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a" using f g fg D.whisker_right by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a" using f D.comp_arr_inv' by simp finally have "((F\<^sub>0 ?c \\<^sub>D F g) \\<^sub>D \\<^sub>D[F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D\<^sup>-\<^sup>1[G g] up, where it will cancel. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg * D.comp_arr_dom D.interchange [of "\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b]" "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g]" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" "(G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg D.runit_hcomp(4) [of "\\<^sub>0' ?c" "G g"] by simp finally have "(\\<^sub>D[\\<^sub>0' ?c, G g, G\<^sub>0 ?b] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D G g] \\<^sub>D (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" by blast thus ?thesis using D.comp_assoc by simp qed (* Move \\<^sub>D\<^sup>-\<^sup>1[G g] up. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D G f \\<^sub>D \ ?a)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D (G f \\<^sub>D \ ?a) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg D.comp_cod_arr D.interchange [of "\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b" "\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]" "G f \\<^sub>D \ ?a" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])" using f g fg D.comp_arr_dom D.interchange [of "\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]" "\\<^sub>0' ?c \\<^sub>D G g" "G f \\<^sub>D \ ?a" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp finally have "((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G\<^sub>0 ?b) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) = ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])" by blast thus ?thesis using D.comp_assoc by simp qed (* Cancel \\<^sub>D\<^sup>-\<^sup>1[G g] and \\<^sub>D[G g]. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D G f \\<^sub>D \ ?a) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D \\<^sub>D[G f] \\<^sub>D (G f \\<^sub>D \ ?a)" using f g fg D.interchange by simp also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D \\<^sub>D[G f] \\<^sub>D (G f \\<^sub>D \ ?a)" using f g fg D.whisker_left by simp also have "... = (\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f] \\<^sub>D (G f \\<^sub>D \ ?a)" using g D.comp_arr_inv' by simp also have "... = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G f \\<^sub>D \ ?a)" using f g fg D.whisker_left by simp finally have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g]) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G g]) \\<^sub>D G f \\<^sub>D \ ?a) = ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G f \\<^sub>D \ ?a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, (G f \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a])" using f g fg D.assoc'_naturality [of "\\<^sub>0' ?c" "G g" "\\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]"] by simp thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) = \\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a)" using f g fg D.whisker_left by simp also have "... = \\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a] \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg D.pentagon by simp also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a])" using f g fg D.whisker_left by simp finally have "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>D[G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a] \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a])" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G f \\<^sub>D G\<^sub>0 ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G f \\<^sub>D G\<^sub>0 ?a] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G f \\<^sub>D \ ?a)" using f g fg D.assoc'_naturality [of "\\<^sub>0' ?c" "G g" "G f \\<^sub>D \ ?a"] by simp thus ?thesis using D.comp_assoc by simp qed (* Move associativity to where it can be canceled. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G f]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>D[G f])) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D G g) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G f \\<^sub>D G\<^sub>0 ?a] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G f] \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>D[G f])" using f g fg D.assoc'_naturality [of "\\<^sub>0' ?c" "G g" "\\<^sub>D[G f]"] by simp thus ?thesis using D.comp_assoc by simp qed (* Cancel associativities. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a])) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c, G g, G f]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>D[G f]) = (\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>D[G f])" using f g fg D.comp_arr_inv' D.comp_cod_arr by simp thus ?thesis using D.comp_assoc by simp qed (* Permute associativity to continue forming G g \\<^sub>D G f. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, G\<^sub>0 ?a])) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D G f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a]) = \\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a]" using f g fg D.whisker_left by simp also have "... = \\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, G\<^sub>0 ?a] \\<^sub>D ((G g \\<^sub>D G f) \\<^sub>D \ ?a)" using f g fg D.assoc_naturality [of "G g" "G f" "\ ?a"] by simp also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, G\<^sub>0 ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D G f) \\<^sub>D \ ?a)" using f g fg * D.whisker_left by simp finally have "(\\<^sub>0' ?c \\<^sub>D G g \\<^sub>D G f \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a]) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g, G f, G\<^sub>0 ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D G f) \\<^sub>D \ ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Form \\<^sub>D[G g \\<^sub>D G f]. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D G f) \\<^sub>D \ ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" using f g fg * D.runit_hcomp D.whisker_left D.comp_assoc by simp (* Move D.inv (\\<^sub>G (g, f)) to the top and cancel with its inverse. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D G f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) = \\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a)" using f g fg G.cmp_components_are_iso D.whisker_left by simp also have "... = \\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]" using f g fg G.cmp_components_are_iso D.assoc_naturality [of "D.inv (\\<^sub>G (g, f))" "\\<^sub>0 ?a" "\\<^sub>0' ?a"] by simp also have "... = (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a])" using f g fg G.cmp_components_are_iso D.whisker_left by simp finally have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f, \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a])" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)) \\<^sub>D G\<^sub>0 ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D G f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) = \\<^sub>0' ?c \\<^sub>D ((G g \\<^sub>D G f) \\<^sub>D \ ?a) \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a)" using f g fg G.cmp_components_are_iso D.whisker_left by simp also have "... = \\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D \ ?a)" using f g fg * D.comp_arr_dom D.comp_cod_arr D.interchange [of "G g \\<^sub>D G f" "D.inv (\\<^sub>G (g, f))" "\ ?a" "\\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a"] by simp also have "... = \\<^sub>0' ?c \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (G (g \\<^sub>C f) \\<^sub>D \ ?a)" using f g fg * D.comp_arr_dom D.comp_cod_arr D.interchange [of "D.inv (\\<^sub>G (g, f))" "G (g \\<^sub>C f)" "G\<^sub>0 ?a" "\ ?a"] by simp also have "... = (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)) \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a)" using f g fg * G.cmp_components_are_iso D.whisker_left by simp finally have "(\\<^sub>0' ?c \\<^sub>D (G g \\<^sub>D G f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>0 ?a \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)) \\<^sub>D G\<^sub>0 ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (((\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)))) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)])) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)) \\<^sub>D G\<^sub>0 ?a) = \\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f] \\<^sub>D (D.inv (\\<^sub>G (g, f)) \\<^sub>D G\<^sub>0 ?a)" using f g fg * G.cmp_components_are_iso D.whisker_left by simp also have "... = \\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]" using f g fg * G.cmp_components_are_iso D.runit_naturality [of "D.inv (\\<^sub>G (g, f))"] by simp also have "... = (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)])" using f g fg * G.cmp_components_are_iso D.whisker_left by simp finally have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G g \\<^sub>D G f]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)) \\<^sub>D G\<^sub>0 ?a) = (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)])" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a]" proof - have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)))) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) = \\<^sub>0' ?c \\<^sub>D (\\<^sub>G (g, f) \\<^sub>D D.inv (\\<^sub>G (g, f))) \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]" using f g fg D.whisker_left G.cmp_components_are_iso C.VV.arr_char C.VV.dom_simp C.VV.cod_simp by simp also have "... = \\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]" using f g fg G.cmp_components_are_iso D.comp_arr_inv' D.comp_cod_arr by simp finally have "((\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>G (g, f)))) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) = \\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]" by blast thus ?thesis using D.comp_assoc by simp qed (* Perform similar manipulations on the bottom part. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.assoc_naturality [of "\\<^sub>D\<^sup>-\<^sup>1[F g]" "F f" "\\<^sub>0' ?a"] D.comp_assoc by simp also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" proof - have "((\ ?c \\<^sub>D F g) \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[F\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a] = \\<^sub>D[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, F f, \\<^sub>0' ?a] \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg ** D.assoc_naturality [of "\ ?c \\<^sub>D F g" "F f" "\\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\ ?c \\<^sub>D F g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" proof - have "(\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[(\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D F g, F f, \\<^sub>0' ?a] = \\<^sub>D[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg D.assoc_naturality [of "\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g]" "F f" "\\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed (* Replace \\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f by \\<^sub>D\<^sup>-\<^sup>1[F (g \\<^sub>C f)]. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\ ?c \\<^sub>D F g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a)" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a = \\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a" using f g fg D.lunit_hcomp(2) [of "F g" "F f"] by simp also have "... = (\\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a)" using f g fg ** D.whisker_right [of "\\<^sub>0' ?a"] by simp finally have "(\\<^sub>D\<^sup>-\<^sup>1[F g] \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a = (\\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Move the associativity up to the big block of associativities. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\ ?c \\<^sub>D F g \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a)" proof - have "(((\ ?c \\<^sub>D F g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \\<^sub>D \\<^sub>0' ?a) = ((\ ?c \\<^sub>D F g) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \\<^sub>D \\<^sub>0' ?a" using f g fg ** D.whisker_right by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c, F g, F f] \\<^sub>D (\ ?c \\<^sub>D F g \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" using f g fg ** D.assoc'_naturality [of "\ ?c" "F g" "F f"] by simp also have "... = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" using f g fg ** D.whisker_right by simp finally have "(((\ ?c \\<^sub>D F g) \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[trg\<^sub>D (F g), F g, F f] \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* Simplify the block of associativities. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g \\<^sub>D F f, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a)" proof - have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) = \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g \\<^sub>D F f, \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a)" proof - have "(\\<^sub>0' ?c \\<^sub>D \\<^sub>D[\\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g, F f \\<^sub>D \\<^sub>0' ?a] \\<^sub>D \\<^sub>D[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D F g, F f, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g] \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c, F g, F f] \\<^sub>D \\<^sub>0' ?a) = \(\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0 ?c\<^bold>\, \<^bold>\F g\<^bold>\, \<^bold>\F f\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\) \<^bold>\ (\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 ?c\<^bold>\ \<^bold>\ \<^bold>\F g\<^bold>\, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\\\<^sub>0 ?c\<^bold>\ \<^bold>\ \<^bold>\F g\<^bold>\, \<^bold>\F f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 ?c\<^bold>\ \<^bold>\ \<^bold>\F g\<^bold>\, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\\\<^sub>0 ?c\<^bold>\, \<^bold>\F g\<^bold>\\<^bold>] \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\) \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 ?c\<^bold>\, \<^bold>\F g\<^bold>\, \<^bold>\F f\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\)\" using f g fg D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by simp also have "... = \\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\\\<^sub>0 ?c\<^bold>\ \<^bold>\ \<^bold>\F g\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' ?a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' ?c\<^bold>\, \<^bold>\\\<^sub>0 ?c\<^bold>\, \<^bold>\F g\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0' ?a\<^bold>\)\" using f g fg by (intro EV.eval_eqI, auto) also have "... = \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g \\<^sub>D F f, \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a)" using f g fg D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by simp finally show ?thesis by simp qed thus ?thesis using D.comp_assoc by simp qed (* Permute \\<^sub>F (g, f) to the bottom. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F (g \\<^sub>C f), \\<^sub>0' ?a] \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D ((\ ?c \\<^sub>D F g \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a)" proof - have "(\\<^sub>0' ?c \\<^sub>D ((\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F g \\<^sub>D F f, \\<^sub>0' ?a] = \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F (g \\<^sub>C f), \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)" using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.assoc_naturality [of "\\<^sub>0' ?c" "\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)" "\\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F (g \\<^sub>C f), \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F (g \\<^sub>C f)] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a)" proof - have "((\\<^sub>0' ?c \\<^sub>D (\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f))) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a" using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.whisker_right by simp also have "... = \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F (g \\<^sub>C f)] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a" using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.assoc_naturality [of "\\<^sub>0' ?c" "\\<^sub>0 ?c" "\\<^sub>F (g, f)"] by simp also have "... = (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F (g \\<^sub>C f)] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)" using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.whisker_right by simp finally have "((\\<^sub>0' ?c \\<^sub>D (\\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f))) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F (g \\<^sub>C f)] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F (g \\<^sub>C f), \\<^sub>0' ?a] \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F (g \\<^sub>C f)] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a)" proof - have "(((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = ((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D (\ ?c \\<^sub>D F g \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a" using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.whisker_right by simp also have "... = (\ ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a" using f g fg D.comp_arr_dom D.comp_cod_arr C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.interchange [of "\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c" "\ ?c" "\\<^sub>F (g, f)" "F g \\<^sub>D F f"] by simp also have "... = (\ ?c \\<^sub>D F (g \\<^sub>C f)) \\<^sub>D (F\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a" using f g fg D.comp_arr_dom D.comp_cod_arr C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.interchange [of "\ ?c" "F\<^sub>0 ?c" "F (g \\<^sub>C f)" "\\<^sub>F (g, f)"] by simp also have "... = ((\ ?c \\<^sub>D F (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)" using f g fg ** C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.whisker_right by simp finally have "(((\\<^sub>0' ?c \\<^sub>D \\<^sub>0 ?c) \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F g \\<^sub>D F f) \\<^sub>D \\<^sub>0' ?a) = ((\ ?c \\<^sub>D F (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((F\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F (g \\<^sub>C f), \\<^sub>0' ?a]) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F (g \\<^sub>C f)] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F (g \\<^sub>C f)] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>F (g, f) \\<^sub>D \\<^sub>0' ?a)" proof - have "((F\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a) = (F\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a" using f g fg ** C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.whisker_right by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[F (g \\<^sub>C f)] \\<^sub>D \\<^sub>F (g, f) \\<^sub>D \\<^sub>0' ?a" using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.lunit'_naturality [of "\\<^sub>F (g, f)"] by simp also have "... = (\\<^sub>D\<^sup>-\<^sup>1[F (g \\<^sub>C f)] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>F (g, f) \\<^sub>D \\<^sub>0' ?a)" using f g fg C.VV.arr_char C.VV.dom_simp C.VV.cod_char F.FF_def D.whisker_right by simp finally have "((F\<^sub>0 ?c \\<^sub>D \\<^sub>F (g, f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F g \\<^sub>D F f] \\<^sub>D \\<^sub>0' ?a) = (\\<^sub>D\<^sup>-\<^sup>1[F (g \\<^sub>C f)] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>F (g, f) \\<^sub>D \\<^sub>0' ?a)" by blast thus ?thesis using D.comp_assoc by simp qed (* One more associativity to move to its final position. *) also have "... = (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f)]) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D G (g \\<^sub>C f) \\<^sub>D \ ?a) \\<^sub>D (\\<^sub>0' ?c \\<^sub>D \\<^sub>D[G (g \\<^sub>C f), \\<^sub>0 ?a, \\<^sub>0' ?a]) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G (g \\<^sub>C f) \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f))) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c, F (g \\<^sub>C f)] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D ((\ ?c \\<^sub>D F (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F (g \\<^sub>C f)] \\<^sub>D \\<^sub>0' ?a) \\<^sub>D (\\<^sub>F (g, f) \\<^sub>D \\<^sub>0' ?a)" proof - have "(\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f)) \\<^sub>D \\<^sub>0' ?a) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, \\<^sub>0 ?c \\<^sub>D F (g \\<^sub>C f), \\<^sub>0' ?a] = \\<^sub>D[\\<^sub>0' ?c, G (g \\<^sub>C f) \\<^sub>D \\<^sub>0 ?a, \\<^sub>0' ?a] \\<^sub>D ((\\<^sub>0' ?c \\<^sub>D D.inv (\\<^sub>1 (g \\<^sub>C f))) \\<^sub>D \\<^sub>0' ?a)" using f g fg \.iso_map\<^sub>1_ide D.assoc_naturality [of "\\<^sub>0' ?c" "D.inv (\\<^sub>1 (g \\<^sub>C f))" "\\<^sub>0' ?a"] by simp thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>1' (g \\<^sub>C f) \\<^sub>D (\\<^sub>F (g, f) \\<^sub>D \\<^sub>0' ?a)" unfolding map\<^sub>1_def using f g fg D.comp_assoc by simp finally show "(\\<^sub>0' ?c \\<^sub>D \\<^sub>G (g, f)) \\<^sub>D \\<^sub>D[\\<^sub>0' ?c, G g, G f] \\<^sub>D (\\<^sub>1' g \\<^sub>D G f) \\<^sub>D D.inv \\<^sub>D[F g, \\<^sub>0' (src\<^sub>C g), G f] \\<^sub>D (F g \\<^sub>D \\<^sub>1' f) \\<^sub>D \\<^sub>D[F g, F f, \\<^sub>0' ?a] = \\<^sub>1' (g \\<^sub>C f) \\<^sub>D (\\<^sub>F (g, f) \\<^sub>D \\<^sub>0' ?a)" by blast qed qed lemma is_pseudonatural_equivalence: shows "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 F \\<^sub>F \\<^sub>0' \\<^sub>1'" .. end subsection "Pseudonaturally Equivalent Pseudofunctors" text \ Pseudofunctors \F\ and \G\ are \emph{pseudonaturally equivalent} if there is a pseudonatural equivalence between them. \ definition pseudonaturally_equivalent where "pseudonaturally_equivalent 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. pseudonatural_equivalence 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" lemma pseudonaturally_equivalent_reflexive: assumes "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" shows "pseudonaturally_equivalent 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 F \\<^sub>F" proof - interpret F: 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 by simp interpret identity_pseudonatural_transformation 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 .. show ?thesis using pseudonatural_equivalence_axioms pseudonaturally_equivalent_def by blast qed lemma pseudonaturally_equivalent_symmetric: assumes "pseudonaturally_equivalent 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" shows "pseudonaturally_equivalent 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 F \\<^sub>F" proof - obtain \\<^sub>0 \\<^sub>1 where \: "pseudonatural_equivalence 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" using assms pseudonaturally_equivalent_def by blast interpret \: pseudonatural_equivalence 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 using \ by simp interpret \: converse_pseudonatural_equivalence 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 .. show ?thesis using \.is_pseudonatural_equivalence pseudonaturally_equivalent_def by blast qed lemma pseudonaturally_equivalent_transitive: assumes "pseudonaturally_equivalent 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" and "pseudonaturally_equivalent 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 \\<^sub>H" shows "pseudonaturally_equivalent 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 H \\<^sub>H" proof - obtain \\<^sub>0 \\<^sub>1 where \: "pseudonatural_equivalence 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" using assms pseudonaturally_equivalent_def by blast interpret \: pseudonatural_equivalence 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 using \ by simp obtain \\<^sub>0 \\<^sub>1 where \: "pseudonatural_equivalence 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 \\<^sub>H \\<^sub>0 \\<^sub>1" using assms pseudonaturally_equivalent_def by blast interpret \: pseudonatural_equivalence 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 \\<^sub>H \\<^sub>0 \\<^sub>1 using \ by simp interpret \\: composite_pseudonatural_equivalence 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 H \\<^sub>H \\<^sub>0 \\<^sub>1 \\<^sub>0 \\<^sub>1 .. show ?thesis using \\.pseudonatural_equivalence_axioms pseudonaturally_equivalent_def by blast qed text \ If \\\ is a pseudonatural equivalence from pseudofunctor \F\ to pseudofunctor \G\ and \\\ is the converse equivalence, then \F\ is locally naturally isomorphic to the functor that takes a 2-cell \\\ of \C\ to \\\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)\ and symmetrically for \G\. Here we just establish the naturality property and and that each 1-cell \\g : a \\<^sub>C a'\\ is isomorphic to \\\<^sub>0 a' \\<^sub>D (\\<^sub>0 a' \\<^sub>D g \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0 a\. We ultimately need these results to prove that a pseudofunctor extends to an equivalence of bicategories if and only if it is an equivalence pseudofunctor. \ context pseudonatural_equivalence begin 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 F \\<^sub>F G \\<^sub>G \\<^sub>0 \\<^sub>1 .. abbreviation (input) \\<^sub>0 where "\\<^sub>0 \ \.map\<^sub>0" definition \ where "\ f \ \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), F f, \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0 (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 (src\<^sub>C f), \\<^sub>0 (src\<^sub>C f)] \\<^sub>D (G f \\<^sub>D D.inv (\.\ (src\<^sub>C f))) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f]" lemma \_in_hom [intro]: assumes "C.ide f" shows "\\ f : G f \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D F f \\<^sub>D \\<^sub>0 (src\<^sub>C f)\" unfolding \_def using assms by (intro D.comp_in_homI' D.hseqI') auto lemma iso_\: assumes "C.ide f" shows "D.iso (\ f)" unfolding \_def using assms iso_map\<^sub>1_ide by (intro D.isos_compose) auto definition \ where "\ f \ (\\<^sub>0 (trg\<^sub>C f) \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0 (trg\<^sub>C f), \\<^sub>0 (trg\<^sub>C f), F f] \\<^sub>D (\.\ (trg\<^sub>C f) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]" lemma \_in_hom [intro]: assumes "C.ide f" shows "\\ f : F f \\<^sub>D \\<^sub>0 (trg\<^sub>C f) \\<^sub>D G f \\<^sub>D \\<^sub>0 (src\<^sub>C f)\" unfolding \_def using assms iso_map\<^sub>1_ide by (intro D.comp_in_homI' D.hseqI') auto lemma iso_\: assumes "C.ide f" shows "D.iso (\ f)" unfolding \_def using assms iso_map\<^sub>1_ide by (intro D.isos_compose) auto lemma \_naturality: assumes "C.arr \" shows "(\\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \ (C.dom \) = \ (C.cod \) \\<^sub>D F \" and "D.inv (\ (C.cod \)) \\<^sub>D (\\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \ (C.dom \) = F \" proof - let ?a = "src\<^sub>C \" and ?a' = "trg\<^sub>C \" let ?f = "C.dom \" and ?f' = "C.cod \" have "(\\<^sub>0 ?a' \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \ (C.dom \) = ((\\<^sub>0 ?a' \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D (\\<^sub>0 ?a' \\<^sub>D D.inv (\\<^sub>1 ?f))) \\<^sub>D \\<^sub>D[\\<^sub>0 ?a', \\<^sub>0 ?a', F ?f] \\<^sub>D (\.\ ?a' \\<^sub>D F ?f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?f]" unfolding \_def using assms D.comp_assoc by simp also have "... = (\\<^sub>0 ?a' \\<^sub>D D.inv (\\<^sub>1 ?f')) \\<^sub>D ((\\<^sub>0 ?a' \\<^sub>D \\<^sub>0 ?a' \\<^sub>D F \) \\<^sub>D \\<^sub>D[\\<^sub>0 ?a', \\<^sub>0 ?a', F ?f]) \\<^sub>D (\.\ ?a' \\<^sub>D F ?f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?f]" proof - have "(\\<^sub>0 ?a' \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D (\\<^sub>0 ?a' \\<^sub>D D.inv (\\<^sub>1 ?f)) = \\<^sub>0 ?a' \\<^sub>D (G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D D.inv (\\<^sub>1 ?f)" using assms D.whisker_left iso_map\<^sub>1_ide by simp also have "... = \\<^sub>0 ?a' \\<^sub>D D.inv (\\<^sub>1 ?f') \\<^sub>D (\\<^sub>0 ?a' \\<^sub>D F \)" using assms naturality iso_map\<^sub>1_ide D.invert_opposite_sides_of_square [of "\\<^sub>1 ?f'" "G \ \\<^sub>D \\<^sub>0 ?a" "\\<^sub>0 ?a' \\<^sub>D F \" "\\<^sub>1 ?f"] by simp also have "... = (\\<^sub>0 ?a' \\<^sub>D D.inv (\\<^sub>1 ?f')) \\<^sub>D (\\<^sub>0 ?a' \\<^sub>D \\<^sub>0 ?a' \\<^sub>D F \)" using assms D.whisker_left iso_map\<^sub>1_ide by simp finally have "(\\<^sub>0 ?a' \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D (\\<^sub>0 ?a' \\<^sub>D D.inv (\\<^sub>1 ?f)) = (\\<^sub>0 ?a' \\<^sub>D D.inv (\\<^sub>1 ?f')) \\<^sub>D (\\<^sub>0 ?a' \\<^sub>D \\<^sub>0 ?a' \\<^sub>D F \)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0 ?a' \\<^sub>D D.inv (\\<^sub>1 ?f')) \\<^sub>D \\<^sub>D[\\<^sub>0 ?a', \\<^sub>0 ?a', F ?f'] \\<^sub>D (((\\<^sub>0 ?a' \\<^sub>D \\<^sub>0 ?a') \\<^sub>D F \) \\<^sub>D (\.\ ?a' \\<^sub>D F ?f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?f]" proof - have "(\\<^sub>0 ?a' \\<^sub>D \\<^sub>0 ?a' \\<^sub>D F \) \\<^sub>D \\<^sub>D[\\<^sub>0 ?a', \\<^sub>0 ?a', F ?f] = \\<^sub>D[\\<^sub>0 ?a', \\<^sub>0 ?a', F ?f'] \\<^sub>D ((\\<^sub>0 ?a' \\<^sub>D \\<^sub>0 ?a') \\<^sub>D F \)" using assms D.assoc_naturality [of "\\<^sub>0 ?a'" "\\<^sub>0 ?a'" "F \"] by simp thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>0 ?a' \\<^sub>D D.inv (\\<^sub>1 ?f')) \\<^sub>D \\<^sub>D[\\<^sub>0 ?a', \\<^sub>0 ?a', F ?f'] \\<^sub>D (\.\ ?a' \\<^sub>D F ?f') \\<^sub>D (F.map\<^sub>0 ?a' \\<^sub>D F \) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?f]" proof - have "((\\<^sub>0 ?a' \\<^sub>D \\<^sub>0 ?a') \\<^sub>D F \) \\<^sub>D (\.\ ?a' \\<^sub>D F ?f) = (\\<^sub>0 ?a' \\<^sub>D \\<^sub>0 ?a') \\<^sub>D \.\ ?a' \\<^sub>D F \ \\<^sub>D F ?f" using assms D.interchange [of "\\<^sub>0 ?a' \\<^sub>D \\<^sub>0 ?a'" "\.\ ?a'" "F \" "F ?f"] by simp also have "... = \.\ ?a' \\<^sub>D F.map\<^sub>0 ?a' \\<^sub>D F ?f' \\<^sub>D F \" using assms D.comp_arr_dom D.comp_cod_arr by simp also have "... = (\.\ ?a' \\<^sub>D F ?f') \\<^sub>D (F.map\<^sub>0 ?a' \\<^sub>D F \)" using assms D.interchange [of "\.\ ?a'" "F.map\<^sub>0 ?a'" "F ?f'" "F \"] \.unit_in_hom [of ?a'] by fastforce finally have "((\\<^sub>0 ?a' \\<^sub>D \\<^sub>0 ?a') \\<^sub>D F \) \\<^sub>D (\.\ ?a' \\<^sub>D F ?f) = (\.\ ?a' \\<^sub>D F ?f') \\<^sub>D (F.map\<^sub>0 ?a' \\<^sub>D F \)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = \ ?f' \\<^sub>D F \" proof - have "(F.map\<^sub>0 ?a' \\<^sub>D F \) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?f] = \\<^sub>D\<^sup>-\<^sup>1[F ?f'] \\<^sub>D F \" using assms D.lunit'_naturality [of "F \"] by simp thus ?thesis unfolding \_def using assms D.comp_assoc by simp qed finally show "(\\<^sub>0 ?a' \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \ (C.dom \) = \ ?f' \\<^sub>D F \" by blast thus "D.inv (\ ?f') \\<^sub>D (\\<^sub>0 ?a' \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \ (C.dom \) = F \" using assms \_in_hom iso_\ by (metis C.ide_cod D.in_homE D.invert_side_of_triangle(1) D.seqI F.preserves_arr F.preserves_cod) qed lemma isomorphic_expansion: assumes "C.obj a" and "C.obj a'" and "\g : G.map\<^sub>0 a \\<^sub>D G.map\<^sub>0 a'\" and "D.ide g" shows "D.isomorphic (\\<^sub>0 a' \\<^sub>D (\\<^sub>0 a' \\<^sub>D g \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0 a) g" proof - let ?g' = "\\<^sub>0 a' \\<^sub>D g \\<^sub>D \\<^sub>0 a" have g': "\?g' : F.map\<^sub>0 a \\<^sub>D F.map\<^sub>0 a'\" using assms ide_map\<^sub>0_obj \.map\<^sub>0_simps(3) C.obj_simps by (intro D.in_hhomI D.hseqI') auto have ide_g': "D.ide ?g'" using assms g' \.ide_map\<^sub>0_obj ide_map\<^sub>0_obj by blast let ?\ = "(\\<^sub>0 a' \\<^sub>D \\<^sub>D[g]) \\<^sub>D (\\<^sub>0 a' \\<^sub>D g \\<^sub>D \.\ a) \\<^sub>D (\\<^sub>0 a' \\<^sub>D \\<^sub>D[g, \\<^sub>0 a, \\<^sub>0 a]) \\<^sub>D \\<^sub>D[\\<^sub>0 a', g \\<^sub>D \\<^sub>0 a, \\<^sub>0 a]" have \: "\?\ : ?g' \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0 a' \\<^sub>D g\" proof (intro D.comp_in_homI) show "\\\<^sub>D[\\<^sub>0 a', g \\<^sub>D \\<^sub>0 a, \\<^sub>0 a] : (\\<^sub>0 a' \\<^sub>D g \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0 a' \\<^sub>D (g \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0 a\" using assms g' by fastforce show "\\\<^sub>0 a' \\<^sub>D \\<^sub>D[g, \\<^sub>0 a, \\<^sub>0 a] : \\<^sub>0 a' \\<^sub>D (g \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0 a' \\<^sub>D g \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0 a\" using assms g' by fastforce show "\\\<^sub>0 a' \\<^sub>D g \\<^sub>D \.\ a : \\<^sub>0 a' \\<^sub>D g \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0 a' \\<^sub>D g \\<^sub>D G.map\<^sub>0 a\" using assms C.obj_simps by (intro D.hcomp_in_vhom) auto show "\\\<^sub>0 a' \\<^sub>D \\<^sub>D[g] : \\<^sub>0 a' \\<^sub>D g \\<^sub>D G.map\<^sub>0 a \\<^sub>D \\<^sub>0 a' \\<^sub>D g\" using assms apply (intro D.hcomp_in_vhom) apply auto by fastforce qed have iso_\: "D.iso ?\" using assms g' \ ide_g' ide_map\<^sub>0_obj C.obj_simps by (intro D.isos_compose) auto let ?\ = "\\<^sub>D[g] \\<^sub>D (\.\ a' \\<^sub>D g) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a', \\<^sub>0 a', g] \\<^sub>D (\\<^sub>0 a' \\<^sub>D ?\)" have \: "\?\ : \\<^sub>0 a' \\<^sub>D ?g' \\<^sub>D \\<^sub>0 a \\<^sub>D g\" proof (intro D.comp_in_homI) show "\\\<^sub>0 a' \\<^sub>D ?\ : \\<^sub>0 a' \\<^sub>D ?g' \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0 a' \\<^sub>D \\<^sub>0 a' \\<^sub>D g\" using assms g' \ C.obj_simps by (intro D.hcomp_in_vhom) auto show "\\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a', \\<^sub>0 a', g] : \\<^sub>0 a' \\<^sub>D \\<^sub>0 a' \\<^sub>D g \\<^sub>D (\\<^sub>0 a' \\<^sub>D \\<^sub>0 a') \\<^sub>D g\" proof - have "src\<^sub>D (\\<^sub>0 a') = trg\<^sub>D (\\<^sub>0 a')" using assms by auto moreover have "src\<^sub>D (\\<^sub>0 a') = trg\<^sub>D g" using assms by auto ultimately show ?thesis using assms ide_map\<^sub>0_obj \.ide_map\<^sub>0_obj by auto qed show "\\.\ a' \\<^sub>D g : (\\<^sub>0 a' \\<^sub>D \\<^sub>0 a') \\<^sub>D g \\<^sub>D G.map\<^sub>0 a' \\<^sub>D g\" using assms by fastforce show "\\\<^sub>D[g] : G.map\<^sub>0 a' \\<^sub>D g \\<^sub>D g\" using assms by auto qed have iso_\: "D.iso ?\" using assms g' \ \ iso_\ ide_map\<^sub>0_obj iso_map\<^sub>1_ide \.ide_map\<^sub>0_obj apply (intro D.isos_compose) apply (meson D.arrI D.hseqE D.ide_is_iso D.iso_hcomp D.seqE) apply (metis D.hseqE D.in_hhomE D.iso_assoc' D.trg_hcomp F.map\<^sub>0_def map\<^sub>0_simps(2)) by auto show ?thesis using \ iso_\ D.isomorphic_def by auto qed end text \ Here we show that pseudonatural equivalence respects equivalence pseudofunctors, in the sense that a pseudofunctor \G\, pseudonaturally equivalent to an equivalence pseudofunctor \F\, is itself an equivalence pseudofunctor. \ locale pseudofunctor_pseudonaturally_equivalent_to_equivalence_pseudofunctor = 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 + 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 + \: 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 F \\<^sub>F G \\<^sub>G \\<^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 :: "'c \ 'd" and \\<^sub>F :: "'c * 'c \ 'd" and G :: "'c \ 'd" and \\<^sub>G :: "'c * 'c \ 'd" and \\<^sub>0 :: "'c \ 'd" and \\<^sub>1 :: "'c \ 'd" begin 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 F \\<^sub>F G \\<^sub>G \\<^sub>0 \\<^sub>1 .. sublocale 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 proof show "\\ \'. \C.par \ \'; G \ = G \'\ \ \ = \'" proof - fix \ \' assume par: "C.par \ \'" and eq: "G \ = G \'" let ?a = "src\<^sub>C \" and ?a' = "trg\<^sub>C \" let ?f = "C.dom \" and ?f' = "C.cod \" have "\\<^sub>0 ?a' \\<^sub>D F \ = (\\<^sub>0 ?a' \\<^sub>D F \) \\<^sub>D \\<^sub>1 ?f \\<^sub>D D.inv (\\<^sub>1 ?f)" using par \.ide_map\<^sub>0_obj \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_arr_inv' \.map\<^sub>1_in_hom [of \] by (metis C.dom_trg C.ide_dom C.obj_trg C.trg.preserves_dom D.whisker_left F.is_natural_2 F.naturality F.preserves_arr \.map\<^sub>1_simps(5)) also have "... = ((\\<^sub>0 ?a' \\<^sub>D F \) \\<^sub>D \\<^sub>1 ?f) \\<^sub>D D.inv (\\<^sub>1 ?f)" using D.comp_assoc by simp also have "... = ((\\<^sub>0 ?a' \\<^sub>D F \') \\<^sub>D \\<^sub>1 ?f) \\<^sub>D D.inv (\\<^sub>1 ?f)" using eq par \.naturality [of \] \.naturality [of \'] C.src_cod C.trg_cod by metis also have 1: "... = (\\<^sub>0 ?a' \\<^sub>D F \') \\<^sub>D \\<^sub>1 ?f \\<^sub>D D.inv (\\<^sub>1 ?f)" using D.comp_assoc by blast also have "... = \\<^sub>0 ?a' \\<^sub>D F \'" using par \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_arr_inv' \.map\<^sub>1_in_hom [of \'] by (metis 1 C.ide_dom C.trg_cod D.hseqE D.hseqI' D.hseq_char' D.seqE F.preserves_arr F.preserves_trg calculation) finally have "\\<^sub>0 ?a' \\<^sub>D F \ = \\<^sub>0 ?a' \\<^sub>D F \'" by blast hence "F \ = F \'" using par \.components_are_equivalences D.equivalence_cancel_left [of "\\<^sub>0 ?a'" "F \" "F \'"] by simp thus "\ = \'" using par F.is_faithful by blast qed show "\b. D.obj b \ \a. C.obj a \ D.equivalent_objects (map\<^sub>0 a) b" proof - fix b assume b: "D.obj b" obtain a where a: "C.obj a \ D.equivalent_objects (F.map\<^sub>0 a) b" using b F.biessentially_surjective_on_objects by blast have "D.equivalent_objects (F.map\<^sub>0 a) (map\<^sub>0 a)" using a \.components_are_equivalences D.equivalent_objects_def by (metis F.map\<^sub>0_def map\<^sub>0_def \.map\<^sub>0_in_hhom) hence "D.equivalent_objects (map\<^sub>0 a) b" using a D.equivalent_objects_symmetric D.equivalent_objects_transitive by blast thus "\a. C.obj a \ D.equivalent_objects (map\<^sub>0 a) b" using a by auto qed show "\a a' g. \C.obj a; C.obj a'; \g : map\<^sub>0 a \\<^sub>D map\<^sub>0 a'\; D.ide g\ \ \f. \f : a \\<^sub>C a'\ \ C.ide f \ D.isomorphic (G f) g" proof - fix a a' g assume a: "C.obj a" and a': "C.obj a'" assume g: "\g : map\<^sub>0 a \\<^sub>D map\<^sub>0 a'\" and ide_g: "D.ide g" interpret \: 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 F \\<^sub>F G \\<^sub>G \\<^sub>0 \\<^sub>1 .. 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 G \\<^sub>G F \\<^sub>F \'.map\<^sub>0 \'.map\<^sub>1 using \'.is_pseudonatural_equivalence by simp let ?g' = "\'.map\<^sub>0 a' \\<^sub>D g \\<^sub>D \\<^sub>0 a" have g': "\?g' : F.map\<^sub>0 a \\<^sub>D F.map\<^sub>0 a'\" using a a' g \.ide_map\<^sub>0_obj \'.map\<^sub>0_simps(3) C.obj_simps by (intro D.in_hhomI D.hseqI') auto have ide_g': "D.ide ?g'" using a a' g g' ide_g by (meson D.hcomp_in_hhomE D.hseqE D.ide_hcomp \'.ide_map\<^sub>0_obj \.ide_map\<^sub>0_obj) obtain f where f: "\f : a \\<^sub>C a'\ \ C.ide f \ D.isomorphic (F f) ?g'" using a a' g g' ide_g' F.locally_essentially_surjective [of a a' ?g'] by auto obtain \ where \: "\\ : F f \\<^sub>D ?g'\ \ D.iso \" using f D.isomorphic_def by blast have "D.isomorphic (G f) g" proof - have "D.isomorphic (G f) (\\<^sub>0 a' \\<^sub>D F f \\<^sub>D \'.map\<^sub>0 a)" using a a' f \.iso_\ D.isomorphic_def by blast also have "D.isomorphic ... (\\<^sub>0 a' \\<^sub>D ?g' \\<^sub>D \'.map\<^sub>0 a)" using a a' g' \ D.isomorphic_def by (metis D.hcomp_ide_isomorphic D.hcomp_isomorphic_ide D.hseqE D.ideD(1) D.isomorphic_implies_ide(2) \'.ide_map\<^sub>0_obj calculation \.ide_map\<^sub>0_obj) also have "D.isomorphic ... g" using a a' g ide_g \.isomorphic_expansion by simp finally show ?thesis by blast qed thus "\f. \f : a \\<^sub>C a'\ \ C.ide f \ D.isomorphic (G f) g" using f by auto qed show "\f f' \. \C.ide f; C.ide f'; src\<^sub>C f = src\<^sub>C f'; trg\<^sub>C f = trg\<^sub>C f'; \\ : G f \\<^sub>D G f'\\ \ \\. \\ : f \\<^sub>C f'\ \ G \ = \" proof - fix f f' \ assume f: "C.ide f" and f': "C.ide f'" and eq_src: "src\<^sub>C f = src\<^sub>C f'" and eq_trg: "trg\<^sub>C f = trg\<^sub>C f'" and \: "\\ : G f \\<^sub>D G f'\" let ?a = "src\<^sub>C f" and ?a' = "trg\<^sub>C f" let ?\' = "D.inv (\.\ f') \\<^sub>D (\'.map\<^sub>0 ?a' \\<^sub>D \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \.\ f" have \': "\?\' : F f \\<^sub>D F f'\" proof (intro D.comp_in_homI) show "\\.\ f : F f \\<^sub>D \'.map\<^sub>0 ?a' \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a\" using f \.\_in_hom [of f] by blast show "\\'.map\<^sub>0 ?a' \\<^sub>D \ \\<^sub>D \\<^sub>0 ?a : \'.map\<^sub>0 ?a' \\<^sub>D G f \\<^sub>D \\<^sub>0 ?a \\<^sub>D \'.map\<^sub>0 ?a' \\<^sub>D G f' \\<^sub>D \\<^sub>0 ?a\" using f \ by (intro D.hcomp_in_vhom) auto show "\D.inv (\.\ f') : \'.map\<^sub>0 ?a' \\<^sub>D G f' \\<^sub>D \\<^sub>0 ?a \\<^sub>D F f'\" using f' \.iso_\ eq_src eq_trg by auto qed obtain \ where \: "\\ : f \\<^sub>C f'\ \ F \ = ?\'" using f f' eq_src eq_trg \' F.locally_full [of f f' ?\'] by blast have "G \ = \" proof - have "D.inv (\.\ f') \\<^sub>D (\'.map\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \.\ f = D.inv (\.\ f') \\<^sub>D (\'.map\<^sub>0 ?a' \\<^sub>D \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \.\ f" using \ \.\_naturality [of \] by auto hence "\'.map\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \) = \'.map\<^sub>0 ?a' \\<^sub>D \ \\<^sub>D \\<^sub>0 ?a" using f f' \' \.iso_\ D.iso_cancel_left [of "D.inv (\.\ f')" "(\'.map\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)) \\<^sub>D \.\ f" "(\'.map\<^sub>0 ?a' \\<^sub>D \ \\<^sub>D \\<^sub>0 ?a) \\<^sub>D \.\ f"] D.iso_cancel_right [of "\.\ f" "\'.map\<^sub>0 (trg\<^sub>C \) \\<^sub>D G \ \\<^sub>D \\<^sub>0 (src\<^sub>C \)" "\'.map\<^sub>0 ?a' \\<^sub>D \ \\<^sub>D \\<^sub>0 ?a"] by auto hence "\'.map\<^sub>0 ?a' \\<^sub>D G \ \\<^sub>D \\<^sub>0 ?a = \'.map\<^sub>0 ?a' \\<^sub>D \ \\<^sub>D \\<^sub>0 ?a" using \ by auto moreover have "D.par (G \ \\<^sub>D \\<^sub>0 (src\<^sub>C f)) (\ \\<^sub>D \\<^sub>0 (src\<^sub>C f))" using f \ \ \' apply (intro conjI D.hseqI') apply (auto simp add: D.vconn_implies_hpar(1)) apply fastforce by (metis C.ideD(1) C.vconn_implies_hpar(1) D.hcomp_simps(4) D.hseqE D.hseqI' D.seqE D.vconn_implies_hpar(1) preserves_arr preserves_cod preserves_src weak_arrow_of_homs_axioms category.in_homE horizontal_homs_def weak_arrow_of_homs_def) ultimately have "G \ \\<^sub>D \\<^sub>0 ?a = \ \\<^sub>D \\<^sub>0 ?a" using f f' \ \ \.ide_map\<^sub>0_obj \'.ide_map\<^sub>0_obj D.equivalence_cancel_left [of "\'.map\<^sub>0 ?a'" "G \ \\<^sub>D \\<^sub>0 ?a" "\ \\<^sub>D \\<^sub>0 ?a"] \'.components_are_equivalences by auto moreover have "D.par (G \) \" using f f' \ \ by fastforce ultimately show ?thesis using f f' \ \ \.ide_map\<^sub>0_obj D.equivalence_cancel_right [of "\\<^sub>0 ?a" "G \" "\"] \.components_are_equivalences by auto qed thus "\\. \\ : f \\<^sub>C f'\ \ G \ = \" using \ by auto qed qed lemma is_equivalence_pseudofunctor: shows "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" .. end lemma pseudonaturally_equivalent_respects_equivalence_pseudofunctor: assumes "pseudonaturally_equivalent 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" and "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" shows "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 G \\<^sub>G" proof - obtain \\<^sub>0 \\<^sub>1 where \: "pseudonatural_equivalence 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" using assms pseudonaturally_equivalent_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 assms by simp interpret 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 using assms \ by (simp add: pseudonatural_equivalence_def pseudonatural_transformation_def) interpret \: pseudonatural_equivalence 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 using \ by simp interpret G: pseudofunctor_pseudonaturally_equivalent_to_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 G \\<^sub>G \\<^sub>0 \\<^sub>1 .. show "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 G \\<^sub>G" using G.is_equivalence_pseudofunctor by simp qed end diff --git a/thys/Category3/CartesianCategory.thy b/thys/Category3/CartesianCategory.thy --- a/thys/Category3/CartesianCategory.thy +++ b/thys/Category3/CartesianCategory.thy @@ -1,2285 +1,2285 @@ (* Title: CartesianCategory Author: Eugene W. Stark , 2020 Maintainer: Eugene W. Stark *) chapter "Cartesian Category" text\ In this chapter, we explore the notion of a ``cartesian category'', which we define to be a category having binary products and a terminal object. We show that every cartesian category extends to an ``elementary cartesian category'', whose definition assumes that specific choices have been made for projections and terminal object. Conversely, the underlying category of an elementary cartesian category is a cartesian category. We also show that cartesian categories are the same thing as categories with finite products. \ theory CartesianCategory -imports Category3.Limit Category3.SetCat +imports Limit SetCat begin section "Category with Binary Products" subsection "Binary Product Diagrams" text \ The ``shape'' of a binary product diagram is a category having two distinct identity arrows and no non-identity arrows. \ locale binary_product_shape begin sublocale concrete_category \UNIV :: bool set\ \\a b. if a = b then {()} else {}\ \\_. ()\ \\_ _ _ _ _. ()\ apply (unfold_locales, auto) apply (meson empty_iff) by (meson empty_iff) abbreviation comp where "comp \ COMP" abbreviation FF where "FF \ MkIde False" abbreviation TT where "TT \ MkIde True" lemma arr_char: shows "arr f \ f = FF \ f = TT" using arr_char by (cases f, simp_all) lemma ide_char: shows "ide f \ f = FF \ f = TT" using ide_char ide_MkIde by (cases f, auto) lemma is_discrete: shows "ide f \ arr f" using arr_char ide_char by simp lemma dom_simp [simp]: assumes "arr f" shows "dom f = f" using assms is_discrete by simp lemma cod_simp [simp]: assumes "arr f" shows "cod f = f" using assms is_discrete by simp lemma seq_char: shows "seq f g \ arr f \ f = g" by auto lemma comp_simp [simp]: assumes "seq f g" shows "comp f g = f" using assms seq_char by fastforce end locale binary_product_diagram = J: binary_product_shape + C: category C for C :: "'c comp" (infixr "\" 55) and a0 :: 'c and a1 :: 'c + assumes is_discrete: "C.ide a0 \ C.ide a1" begin notation J.comp (infixr "\\<^sub>J" 55) fun map where "map J.FF = a0" | "map J.TT = a1" | "map _ = C.null" sublocale diagram J.comp C map proof show "\f. \ J.arr f \ map f = C.null" using J.arr_char map.elims by auto fix f assume f: "J.arr f" show "C.arr (map f)" using f J.arr_char is_discrete C.ideD(1) map.simps(1-2) by metis show "C.dom (map f) = map (J.dom f)" using f J.arr_char J.dom_char is_discrete by force show "C.cod (map f) = map (J.cod f)" using f J.arr_char J.cod_char is_discrete by force next fix f g assume fg: "J.seq g f" show "map (g \\<^sub>J f) = map g \ map f" using fg J.arr_char J.seq_char J.null_char J.not_arr_null is_discrete by (metis (no_types, lifting) C.comp_ide_self J.comp_simp map.simps(1-2)) qed end subsection "Category with Binary Products" text \ A \emph{binary product} in a category @{term C} is a limit of a binary product diagram in @{term C}. \ context binary_product_diagram begin definition mkCone where "mkCone p0 p1 \ \j. if j = J.FF then p0 else if j = J.TT then p1 else C.null" abbreviation is_rendered_commutative_by where "is_rendered_commutative_by p0 p1 \ C.seq a0 p0 \ C.seq a1 p1 \ C.dom p0 = C.dom p1" abbreviation has_as_binary_product where "has_as_binary_product p0 p1 \ limit_cone (C.dom p0) (mkCone p0 p1)" lemma cone_mkCone: assumes "is_rendered_commutative_by p0 p1" shows "cone (C.dom p0) (mkCone p0 p1)" proof - interpret E: constant_functor J.comp C \C.dom p0\ using assms by unfold_locales auto show "cone (C.dom p0) (mkCone p0 p1)" using assms mkCone_def J.arr_char E.map_simp is_discrete C.comp_ide_arr C.comp_arr_dom by unfold_locales auto qed lemma is_rendered_commutative_by_cone: assumes "cone a \" shows "is_rendered_commutative_by (\ J.FF) (\ J.TT)" proof - interpret \: cone J.comp C map a \ using assms by auto show ?thesis using is_discrete by simp qed lemma mkCone_cone: assumes "cone a \" shows "mkCone (\ J.FF) (\ J.TT) = \" proof - interpret \: cone J.comp C map a \ using assms by auto interpret mkCone_\: cone J.comp C map \C.dom (\ J.FF)\ \mkCone (\ J.FF) (\ J.TT)\ using assms is_rendered_commutative_by_cone cone_mkCone by blast show ?thesis using mkCone_def \.is_extensional J.ide_char mkCone_def NaturalTransformation.eqI [of J.comp C] \.natural_transformation_axioms mkCone_\.natural_transformation_axioms by fastforce qed end locale binary_product_cone = J: binary_product_shape + C: category C + D: binary_product_diagram C f0 f1 + limit_cone J.comp C D.map \C.dom p0\ \D.mkCone p0 p1\ for C :: "'c comp" (infixr "\" 55) and f0 :: 'c and f1 :: 'c and p0 :: 'c and p1 :: 'c begin lemma renders_commutative: shows "D.is_rendered_commutative_by p0 p1" using cone_axioms D.is_rendered_commutative_by_cone D.mkCone_def \.Ya.Cop_S.arr.simps(1) by (metis (no_types, lifting)) (* TODO: pretty opaque *) lemma is_universal': assumes "D.is_rendered_commutative_by p0' p1'" shows "\!h. \h : C.dom p0' \ C.dom p0\ \ p0 \ h = p0' \ p1 \ h = p1'" proof - have "D.cone (C.dom p0') (D.mkCone p0' p1')" using assms D.cone_mkCone by blast hence "\!h. \h : C.dom p0' \ C.dom p0\ \ D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1'" using is_universal by simp moreover have "\h. \h : C.dom p0' \ C.dom p0\ \ D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1' \ p0 \ h = p0' \ p1 \ h = p1'" proof - fix h assume h: "\h : C.dom p0' \ C.dom p0\" show "D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1' \ p0 \ h = p0' \ p1 \ h = p1'" proof assume 1: "D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1'" show "p0 \ h = p0' \ p1 \ h = p1'" proof show "p0 \ h = p0'" proof - have "p0' = D.mkCone p0' p1' J.FF" using D.mkCone_def J.arr_char by simp also have "... = D.cones_map h (D.mkCone p0 p1) J.FF" using 1 by simp also have "... = p0 \ h" using h D.mkCone_def J.arr_char cone_\ by auto finally show ?thesis by auto qed show "p1 \ h = p1'" proof - have "p1' = D.mkCone p0' p1' J.TT" using D.mkCone_def J.arr_char by simp also have "... = D.cones_map h (D.mkCone p0 p1) J.TT" using 1 by simp also have "... = p1 \ h" using h D.mkCone_def J.arr_char cone_\ by auto finally show ?thesis by auto qed qed next assume 1: "p0 \ h = p0' \ p1 \ h = p1'" show "D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1'" using h 1 cone_\ D.mkCone_def by auto qed qed ultimately show ?thesis by blast qed lemma induced_arrowI': assumes "D.is_rendered_commutative_by p0' p1'" shows "\induced_arrow (C.dom p0') (D.mkCone p0' p1') : C.dom p0' \ C.dom p0\" and "p0 \ induced_arrow (C.dom p0') (D.mkCone p0' p1') = p0'" and "p1 \ induced_arrow (C.dom p1') (D.mkCone p0' p1') = p1'" proof - interpret A': constant_functor J.comp C \C.dom p0'\ using assms by (unfold_locales, auto) have cone: "D.cone (C.dom p0') (D.mkCone p0' p1')" using assms D.cone_mkCone [of p0' p1'] by blast show 0: "p0 \ induced_arrow (C.dom p0') (D.mkCone p0' p1') = p0'" proof - have "p0 \ induced_arrow (C.dom p0') (D.mkCone p0' p1') = D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1')) (D.mkCone p0 p1) J.FF" using cone induced_arrowI(1) D.mkCone_def J.arr_char cone_\ by force also have "... = p0'" proof - have "D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1')) (D.mkCone p0 p1) = D.mkCone p0' p1'" using cone induced_arrowI by blast thus ?thesis using J.arr_char D.mkCone_def by simp qed finally show ?thesis by auto qed show "p1 \ induced_arrow (C.dom p1') (D.mkCone p0' p1') = p1'" proof - have "p1 \ induced_arrow (C.dom p1') (D.mkCone p0' p1') = D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1')) (D.mkCone p0 p1) J.TT" using assms cone induced_arrowI(1) D.mkCone_def J.arr_char cone_\ by fastforce also have "... = p1'" proof - have "D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1')) (D.mkCone p0 p1) = D.mkCone p0' p1'" using cone induced_arrowI by blast thus ?thesis using J.arr_char D.mkCone_def by simp qed finally show ?thesis by auto qed show "\induced_arrow (C.dom p0') (D.mkCone p0' p1') : C.dom p0' \ C.dom p0\" using 0 cone induced_arrowI by simp qed end context category begin definition has_as_binary_product where "has_as_binary_product a0 a1 p0 p1 \ ide a0 \ ide a1 \ binary_product_diagram.has_as_binary_product C a0 a1 p0 p1" definition has_binary_products where "has_binary_products = (\a0 a1. ide a0 \ ide a1 \ (\p0 p1. has_as_binary_product a0 a1 p0 p1))" end locale category_with_binary_products = category + assumes has_binary_products: has_binary_products subsection "Elementary Category with Binary Products" text \ An \emph{elementary category with binary products} is a category equipped with a specific way of mapping each pair of objects \a\ and \b\ to a pair of arrows \\\<^sub>1[a, b]\ and \\\<^sub>0[a, b]\ that comprise a universal span. It is useful to assume that the mappings that produce \\\<^sub>1[a, b]\ and \\\<^sub>0[a, b]\ from \a\ and \b\ are extensional; that is, if either \a\ or \b\ is not an identity, then \\\<^sub>1[a, b]\ and \\\<^sub>0[a, b]\ are \null\. \ locale elementary_category_with_binary_products = category C for C :: "'a comp" (infixr "\" 55) and pr0 :: "'a \ 'a \ 'a" ("\\<^sub>0[_, _]") and pr1 :: "'a \ 'a \ 'a" ("\\<^sub>1[_, _]") + assumes pr0_ext: "\ (ide a \ ide b) \ \\<^sub>0[a, b] = null" and pr1_ext: "\ (ide a \ ide b) \ \\<^sub>1[a, b] = null" and span_pr: "\ ide a; ide b \ \ span \\<^sub>1[a, b] \\<^sub>0[a, b]" and cod_pr0: "\ ide a; ide b \ \ cod \\<^sub>0[a, b] = b" and cod_pr1: "\ ide a; ide b \ \ cod \\<^sub>1[a, b] = a" and universal: "span f g \ \!l. \\<^sub>1[cod f, cod g] \ l = f \ \\<^sub>0[cod f, cod g] \ l = g" begin lemma pr0_in_hom': assumes "ide a" and "ide b" shows "\\\<^sub>0[a, b] : dom \\<^sub>0[a, b] \ b\" using assms span_pr cod_pr0 by auto lemma pr1_in_hom': assumes "ide a" and "ide b" shows "\\\<^sub>1[a, b] : dom \\<^sub>0[a, b] \ a\" using assms span_pr cod_pr1 by auto text \ We introduce a notation for tupling, which denotes the arrow into a product that is induced by a span. \ definition tuple ("\_, _\") where "\f, g\ \ if span f g then THE l. \\<^sub>1[cod f, cod g] \ l = f \ \\<^sub>0[cod f, cod g] \ l = g else null" text \ The following defines product of arrows (not just of objects). It will take a little while before we can prove that it is functorial, but for right now it is nice to have it as a notation for the apex of a product cone. We have to go through some slightly unnatural contortions in the development here, though, to avoid having to introduce a separate preliminary notation just for the product of objects. \ (* TODO: I want to use \ but it has already been commandeered for product types. *) definition prod (infixr "\" 51) where "f \ g \ \f \ \\<^sub>1[dom f, dom g], g \ \\<^sub>0[dom f, dom g]\" lemma seq_pr_tuple: assumes "span f g" shows "seq \\<^sub>0[cod f, cod g] \f, g\" proof - have "\\<^sub>0[cod f, cod g] \ \f, g\ = g" unfolding tuple_def using assms universal theI [of "\l. \\<^sub>1[cod f, cod g] \ l = f \ \\<^sub>0[cod f, cod g] \ l = g"] by simp meson thus ?thesis using assms by simp qed lemma tuple_pr_arr: assumes "ide a" and "ide b" and "seq \\<^sub>0[a, b] h" shows "\\\<^sub>1[a, b] \ h, \\<^sub>0[a, b] \ h\ = h" unfolding tuple_def using assms span_pr cod_pr0 cod_pr1 universal [of "\\<^sub>1[a, b] \ h" "\\<^sub>0[a, b] \ h"] theI_unique [of "\l. \\<^sub>1[a, b] \ l = \\<^sub>1[a, b] \ h \ \\<^sub>0[a, b] \ l = \\<^sub>0[a, b] \ h" h] by auto lemma pr_tuple [simp]: assumes "span f g" and "cod f = a" and "cod g = b" shows "\\<^sub>1[a, b] \ \f, g\ = f" and "\\<^sub>0[a, b] \ \f, g\ = g" proof - have 1: "\\<^sub>1[a, b] \ \f, g\ = f \ \\<^sub>0[a, b] \ \f, g\ = g" unfolding tuple_def using assms universal theI [of "\l. \\<^sub>1[a, b] \ l = f \ \\<^sub>0[a, b] \ l = g"] by simp meson show "\\<^sub>1[a, b] \ \f, g\ = f" using 1 by simp show "\\<^sub>0[a, b] \ \f, g\ = g" using 1 by simp qed lemma cod_tuple: assumes "span f g" shows "cod \f, g\ = cod f \ cod g" proof - have "cod f \ cod g = \\\<^sub>1[cod f, cod g], \\<^sub>0[cod f, cod g]\" unfolding prod_def using assms comp_cod_arr span_pr cod_pr0 cod_pr1 by simp also have "... = \\\<^sub>1[cod f, cod g] \ dom \\<^sub>0[cod f, cod g], \\<^sub>0[cod f, cod g] \ dom \\<^sub>0[cod f, cod g]\" using assms span_pr comp_arr_dom by simp also have "... = dom \\<^sub>0[cod f, cod g]" using assms tuple_pr_arr span_pr by simp also have "... = cod \f, g\" using assms seq_pr_tuple by blast finally show ?thesis by simp qed lemma tuple_in_hom [intro]: assumes "\f : a \ b\" and "\g : a \ c\" shows "\\f, g\ : a \ b \ c\" using assms pr_tuple dom_comp cod_tuple apply (elim in_homE, intro in_homI) apply (metis seqE) by metis+ lemma tuple_in_hom' [simp]: assumes "arr f" and "dom f = a" and "cod f = b" and "arr g" and "dom g = a" and "cod g = c" shows "\\f, g\ : a \ b \ c\" using assms by auto lemma tuple_ext: assumes "\ span f g" shows "\f, g\ = null" unfolding tuple_def by (simp add: assms) lemma tuple_simps [simp]: assumes "span f g" shows "arr \f, g\" and "dom \f, g\ = dom f" and "cod \f, g\ = cod f \ cod g" proof - show "arr \f, g\" using assms tuple_in_hom by blast show "dom \f, g\ = dom f" using assms tuple_in_hom by (metis dom_comp pr_tuple(1)) show "cod \f, g\ = cod f \ cod g" using assms cod_tuple by auto qed lemma tuple_pr [simp]: assumes "ide a" and "ide b" shows "\\\<^sub>1[a, b], \\<^sub>0[a, b]\ = a \ b" proof - have 1: "dom \\<^sub>0[a, b] = a \ b" using assms seq_pr_tuple cod_tuple [of "\\<^sub>1[a, b]" "\\<^sub>0[a, b]"] span_pr pr0_in_hom' pr1_in_hom' by (metis cod_pr0 cod_pr1 seqE) hence "\\\<^sub>1[a, b], \\<^sub>0[a, b]\ = \\\<^sub>1[a, b] \ (a \ b), \\<^sub>0[a, b] \ (a \ b)\" using assms pr0_in_hom' pr1_in_hom' comp_arr_dom span_pr by simp thus ?thesis using assms 1 tuple_pr_arr span_pr by (metis comp_arr_dom) qed lemma pr_in_hom [intro, simp]: assumes "ide a" and "ide b" shows "\\\<^sub>0[a, b] : a \ b \ b\" and "\\\<^sub>1[a, b] : a \ b \ a\" proof - show 0: "\\\<^sub>0[a, b] : a \ b \ b\" using assms pr0_in_hom' seq_pr_tuple [of "\\<^sub>1[a, b]" "\\<^sub>0[a, b]"] cod_tuple [of "\\<^sub>1[a, b]" "\\<^sub>0[a, b]"] span_pr cod_pr0 cod_pr1 by (intro in_homI, auto) show "\\\<^sub>1[a, b] : a \ b \ a\" using assms 0 span_pr pr1_in_hom' by fastforce qed lemma pr_simps [simp]: assumes "ide a" and "ide b" shows "arr \\<^sub>0[a, b]" and "dom \\<^sub>0[a, b] = a \ b" and "cod \\<^sub>0[a, b] = b" and "arr \\<^sub>1[a, b]" and "dom \\<^sub>1[a, b] = a \ b" and "cod \\<^sub>1[a, b] = a" using assms pr_in_hom by blast+ lemma arr_pr0_iff [iff]: shows "arr \\<^sub>0[a, b] \ ide a \ ide b" proof show "ide a \ ide b \ arr \\<^sub>0[a, b]" using pr_in_hom by auto show "arr \\<^sub>0[a, b] \ ide a \ ide b" using pr0_ext not_arr_null by metis qed lemma arr_pr1_iff [iff]: shows "arr \\<^sub>1[a, b] \ ide a \ ide b" proof show "ide a \ ide b \ arr \\<^sub>1[a, b]" using pr_in_hom by auto show "arr \\<^sub>1[a, b] \ ide a \ ide b" using pr1_ext not_arr_null by metis qed lemma pr_joint_monic: assumes "seq \\<^sub>0[a, b] h" and "\\<^sub>0[a, b] \ h = \\<^sub>0[a, b] \ h'" and "\\<^sub>1[a, b] \ h = \\<^sub>1[a, b] \ h'" shows "h = h'" using assms by (metis arr_pr0_iff seqE tuple_pr_arr) lemma comp_tuple_arr [simp]: assumes "span f g" and "arr h" and "dom f = cod h" shows "\f, g\ \ h = \f \ h, g \ h\" proof (intro pr_joint_monic [where h = "\f, g\ \ h"]) show "seq \\<^sub>0[cod f, cod g] (\f, g\ \ h)" using assms by fastforce show "\\<^sub>0[cod f, cod g] \ \f, g\ \ h = \\<^sub>0[cod f, cod g] \ \f \ h, g \ h\" proof - have "\\<^sub>0[cod f, cod g] \ \f, g\ \ h = (\\<^sub>0[cod f, cod g] \ \f, g\) \ h" using comp_assoc by simp thus ?thesis using assms by simp qed show "\\<^sub>1[cod f, cod g] \ \f, g\ \ h = \\<^sub>1[cod f, cod g] \ \f \ h, g \ h\" proof - have "\\<^sub>1[cod f, cod g] \ \f, g\ \ h = (\\<^sub>1[cod f, cod g] \ \f, g\) \ h" using comp_assoc by simp thus ?thesis using assms by simp qed qed lemma ide_prod [intro, simp]: assumes "ide a" and "ide b" shows "ide (a \ b)" using assms pr_simps ide_dom [of "\\<^sub>0[a, b]"] by simp lemma prod_in_hom [intro]: assumes "\f : a \ c\" and "\g : b \ d\" shows "\f \ g : a \ b \ c \ d\" using assms prod_def by fastforce lemma prod_in_hom' [simp]: assumes "arr f" and "dom f = a" and "cod f = c" and "arr g" and "dom g = b" and "cod g = d" shows "\f \ g : a \ b \ c \ d\" using assms by blast lemma prod_simps [simp]: assumes "arr f0" and "arr f1" shows "arr (f0 \ f1)" and "dom (f0 \ f1) = dom f0 \ dom f1" and "cod (f0 \ f1) = cod f0 \ cod f1" using assms prod_in_hom by blast+ end subsection "Agreement between the Definitions" text \ We now show that a category with binary products extends (by making a choice) to an elementary category with binary products, and that the underlying category of an elementary category with binary products is a category with binary products. \ context category_with_binary_products begin definition pr1 where "pr1 a b \ if ide a \ ide b then fst (SOME x. has_as_binary_product a b (fst x) (snd x)) else null" definition pr0 where "pr0 a b \ if ide a \ ide b then snd (SOME x. has_as_binary_product a b (fst x) (snd x)) else null" lemma pr_yields_binary_product: assumes "ide a" and "ide b" shows "has_as_binary_product a b (pr1 a b) (pr0 a b)" proof - have "\x. has_as_binary_product a b (fst x) (snd x)" using assms has_binary_products has_binary_products_def has_as_binary_product_def by simp thus ?thesis using assms has_binary_products has_binary_products_def pr0_def pr1_def someI_ex [of "\x. has_as_binary_product a b (fst x) (snd x)"] by simp qed interpretation elementary_category_with_binary_products C pr0 pr1 proof show "\a b. \ (ide a \ ide b) \ pr0 a b = null" using pr0_def by auto show "\a b. \ (ide a \ ide b) \ pr1 a b = null" using pr1_def by auto fix a b assume a: "ide a" and b: "ide b" interpret J: binary_product_shape . interpret D: binary_product_diagram C a b using a b by unfold_locales auto let ?\ = "D.mkCone (pr1 a b) (pr0 a b)" interpret \: limit_cone J.comp C D.map \dom (pr1 a b)\ ?\ using a b pr_yields_binary_product by (simp add: has_as_binary_product_def) have 1: "pr1 a b = ?\ J.FF \ pr0 a b = ?\ J.TT" using D.mkCone_def by simp show "span (pr1 a b) (pr0 a b)" using 1 \.preserves_reflects_arr J.seqE J.arr_char J.seq_char J.is_category D.is_rendered_commutative_by_cone \.cone_axioms by metis show "cod (pr1 a b) = a" using 1 \.preserves_cod [of J.FF] J.cod_char J.arr_char by auto show "cod (pr0 a b) = b" using 1 \.preserves_cod [of J.TT] J.cod_char J.arr_char by auto next fix f g assume fg: "span f g" show "\!l. pr1 (cod f) (cod g) \ l = f \ pr0 (cod f) (cod g) \ l = g" proof - interpret J: binary_product_shape . interpret D: binary_product_diagram C \cod f\ \cod g\ using fg by unfold_locales auto let ?\ = "D.mkCone (pr1 (cod f) (cod g)) (pr0 (cod f) (cod g))" interpret \: limit_cone J.comp C D.map \dom (pr1 (cod f) (cod g))\ ?\ using fg pr_yields_binary_product [of "cod f" "cod g"] has_as_binary_product_def by simp interpret \: binary_product_cone C \cod f\ \cod g\ \pr1 (cod f) (cod g)\ \pr0 (cod f) (cod g)\ .. have 1: "pr1 (cod f) (cod g) = ?\ J.FF \ pr0 (cod f) (cod g) = ?\ J.TT" using D.mkCone_def by simp show "\!l. pr1 (cod f) (cod g) \ l = f \ pr0 (cod f) (cod g) \ l = g" proof - have "\!l. \l : dom f \ dom (pr1 (cod f) (cod g))\ \ pr1 (cod f) (cod g) \ l = f \ pr0 (cod f) (cod g) \ l = g" using fg \.is_universal' by simp moreover have "\l. pr1 (cod f) (cod g) \ l = f \ \l : dom f \ dom (pr1 (cod f) (cod g))\" using fg dom_comp in_homI seqE seqI by metis ultimately show ?thesis by auto qed qed qed proposition extends_to_elementary_category_with_binary_products: shows "elementary_category_with_binary_products C pr0 pr1" .. end context elementary_category_with_binary_products begin interpretation category_with_binary_products C proof show "has_binary_products" proof (unfold has_binary_products_def) have "\a b. ide a \ ide b \ \p0 p1. has_as_binary_product a b p0 p1" proof - fix a b assume ab: "ide a \ ide b" interpret J: binary_product_shape . interpret D: binary_product_diagram C a b using ab by unfold_locales auto have 2: "D.is_rendered_commutative_by \\<^sub>1[a, b] \\<^sub>0[a, b]" using ab by simp let ?\ = "D.mkCone \\<^sub>1[a, b] \\<^sub>0[a, b]" interpret \: cone J.comp C D.map \dom \\<^sub>1[a, b]\ ?\ using D.cone_mkCone 2 by auto interpret \: limit_cone J.comp C D.map \dom \\<^sub>1[a, b]\ ?\ proof fix a' \' assume \': "D.cone a' \'" interpret \': cone J.comp C D.map a' \' using \' by simp show "\!h. \h : a' \ dom \\<^sub>1[a, b]\ \ D.cones_map h (D.mkCone \\<^sub>1[a, b] \\<^sub>0[a, b]) = \'" proof let ?h = "\\' J.FF, \' J.TT\" show h': "\?h : a' \ dom \\<^sub>1[a, b]\ \ D.cones_map ?h (D.mkCone \\<^sub>1[a, b] \\<^sub>0[a, b]) = \'" proof show h: "\?h : a' \ dom \\<^sub>1[a, b]\" using ab tuple_in_hom J.ide_char by fastforce show "D.cones_map ?h (D.mkCone \\<^sub>1[a, b] \\<^sub>0[a, b]) = \'" proof - interpret \'h: cone J.comp C D.map a' \D.cones_map ?h (D.mkCone \\<^sub>1[a, b] \\<^sub>0[a, b])\ proof - have "D.mkCone \\<^sub>1[a, b] \\<^sub>0[a, b] \ D.cones (cod \\' J.FF, \' J.TT\)" using ab h D.cone_mkCone D.is_rendered_commutative_by_cone \.cone_axioms by auto hence "D.cones_map ?h (D.mkCone \\<^sub>1[a, b] \\<^sub>0[a, b]) \ D.cones a'" using ab h D.cones_map_mapsto by blast thus "D.cone a' (D.cones_map ?h (D.mkCone \\<^sub>1[a, b] \\<^sub>0[a, b]))" by simp qed show ?thesis proof - have "\j. J.ide j \ D.cones_map ?h (D.mkCone \\<^sub>1[a, b] \\<^sub>0[a, b]) j = \' j" using ab h J.ide_char D.mkCone_def \.cone_axioms by auto thus ?thesis using NaturalTransformation.eqI \'.natural_transformation_axioms \'h.natural_transformation_axioms by blast qed qed qed show "\h. \h : a' \ dom \\<^sub>1[a, b]\ \ D.cones_map h (D.mkCone \\<^sub>1[a, b] \\<^sub>0[a, b]) = \' \ h = ?h" proof - fix h assume 1: "\h : a' \ dom \\<^sub>1[a, b]\ \ D.cones_map h (D.mkCone \\<^sub>1[a, b] \\<^sub>0[a, b]) = \'" hence "cod h = dom \\<^sub>1[a, b]" by auto show "h = ?h" using 1 ab \.cone_axioms D.mkCone_def h' pr_joint_monic [of a b h ?h] by auto qed qed qed have "has_as_binary_product a b \\<^sub>1[a, b] \\<^sub>0[a, b]" using ab has_as_binary_product_def \.limit_cone_axioms by blast thus "\p0 p1. has_as_binary_product a b p0 p1" by blast qed thus "\a b. ide a \ ide b \ (\p0 p1. has_as_binary_product a b p0 p1)" by simp qed qed proposition is_category_with_binary_products: shows "category_with_binary_products C" .. end subsection "Further Properties" context elementary_category_with_binary_products begin lemma interchange: assumes "seq h f" and "seq k g" shows "(h \ k) \ (f \ g) = h \ f \ k \ g" using assms prod_def comp_tuple_arr comp_assoc by fastforce lemma pr_naturality [simp]: assumes "arr g" and "dom g = b" and "cod g = d" and "arr f" and "dom f = a" and "cod f = c" shows "\\<^sub>0[c, d] \ (f \ g) = g \ \\<^sub>0[a, b]" and "\\<^sub>1[c, d] \ (f \ g) = f \ \\<^sub>1[a, b]" using assms prod_def by fastforce+ abbreviation dup ("\[_]") where "\[f] \ \f, f\" lemma dup_in_hom [intro, simp]: assumes "\f : a \ b\" shows "\\[f] : a \ b \ b\" using assms by fastforce lemma dup_simps [simp]: assumes "arr f" shows "arr \[f]" and "dom \[f] = dom f" and "cod \[f] = cod f \ cod f" using assms dup_in_hom by auto lemma dup_naturality: assumes "\f : a \ b\" shows "\[b] \ f = (f \ f) \ \[a]" using assms prod_def comp_arr_dom comp_cod_arr comp_tuple_arr comp_assoc by fastforce lemma pr_dup [simp]: assumes "ide a" shows "\\<^sub>0[a, a] \ \[a] = a" and "\\<^sub>1[a, a] \ \[a] = a" using assms by simp_all lemma prod_tuple: assumes "span f g" and "seq h f" and "seq k g" shows "(h \ k) \ \f, g\ = \h \ f, k \ g\" using assms prod_def comp_assoc comp_tuple_arr by fastforce lemma tuple_eqI: assumes "seq \\<^sub>0[b, c] f" and "seq \\<^sub>1[b, c] f" and "\\<^sub>0[b, c] \ f = f0" and "\\<^sub>1[b, c] \ f = f1" shows "f = \f1, f0\" using assms pr_joint_monic [of b c "\f1, f0\" f] pr_tuple by auto definition assoc ("\[_, _, _]") where "\[a, b, c] \ \\\<^sub>1[a, b] \ \\<^sub>1[a \ b, c], \\\<^sub>0[a, b] \ \\<^sub>1[a \ b, c], \\<^sub>0[a \ b, c]\\" definition assoc' ("\\<^sup>-\<^sup>1[_, _, _]") where "\\<^sup>-\<^sup>1[a, b, c] \ \\\\<^sub>1[a, b \ c], \\<^sub>1[b, c] \ \\<^sub>0[a, b \ c]\, \\<^sub>0[b, c] \ \\<^sub>0[a, b \ c]\" lemma assoc_in_hom [intro]: assumes "ide a" and "ide b" and "ide c" shows "\\[a, b, c] : (a \ b) \ c \ a \ (b \ c)\" using assms assoc_def by auto lemma assoc_simps [simp]: assumes "ide a" and "ide b" and "ide c" shows "arr \[a, b, c]" and "dom \[a, b, c] = (a \ b) \ c" and "cod \[a, b, c] = a \ (b \ c)" using assms assoc_in_hom by auto lemma assoc'_in_hom [intro]: assumes "ide a" and "ide b" and "ide c" shows "\\\<^sup>-\<^sup>1[a, b, c] : a \ (b \ c) \ (a \ b) \ c\" using assms assoc'_def by auto lemma assoc'_simps [simp]: assumes "ide a" and "ide b" and "ide c" shows "arr \\<^sup>-\<^sup>1[a, b, c]" and "dom \\<^sup>-\<^sup>1[a, b, c] = a \ (b \ c)" and "cod \\<^sup>-\<^sup>1[a, b, c] = (a \ b) \ c" using assms assoc'_in_hom by auto lemma assoc_naturality: assumes "\f0 : a0 \ b0\" and "\f1 : a1 \ b1\" and "\f2 : a2 \ b2\" shows "\[b0, b1, b2] \ ((f0 \ f1) \ f2) = (f0 \ (f1 \ f2)) \ \[a0, a1, a2]" proof - have "\\<^sub>0[b0, b1 \ b2] \ \[b0, b1, b2] \ ((f0 \ f1) \ f2) = \\<^sub>0[b0, b1 \ b2] \ (f0 \ (f1 \ f2)) \ \[a0, a1, a2]" proof - have "\\<^sub>0[b0, b1 \ b2] \ \[b0, b1, b2] \ ((f0 \ f1) \ f2) = (\\<^sub>0[b0, b1 \ b2] \ \[b0, b1, b2]) \ ((f0 \ f1) \ f2)" using comp_assoc by simp also have "... = \\\<^sub>0[b0, b1] \ \\<^sub>1[b0 \ b1, b2], \\<^sub>0[b0 \ b1, b2]\ \ ((f0 \ f1) \ f2)" using assms assoc_def by fastforce also have "... = \(\\<^sub>0[b0, b1] \ \\<^sub>1[b0 \ b1, b2]) \ ((f0 \ f1) \ f2), \\<^sub>0[b0 \ b1, b2] \ ((f0 \ f1) \ f2)\" using assms comp_tuple_arr by fastforce also have "... = \(\\<^sub>0[b0, b1] \ (f0 \ f1)) \ \\<^sub>1[a0 \ a1, a2], f2 \ \\<^sub>0[a0 \ a1, a2]\" using assms comp_assoc by fastforce also have "... = \f1 \ \\<^sub>0[a0, a1] \ \\<^sub>1[a0 \ a1, a2], f2 \ \\<^sub>0[a0 \ a1, a2]\" using assms comp_assoc by (metis in_homE pr_naturality(1)) also have "... = \\<^sub>0[b0, b1 \ b2] \ (f0 \ (f1 \ f2)) \ \[a0, a1, a2]" using assms comp_assoc assoc_def prod_tuple by fastforce finally show ?thesis by blast qed moreover have "\\<^sub>1[b0, b1 \ b2] \ \[b0, b1, b2] \ ((f0 \ f1) \ f2) = \\<^sub>1[b0, b1 \ b2] \ (f0 \ (f1 \ f2)) \ \[a0, a1, a2]" proof - have "\\<^sub>1[b0, b1 \ b2] \ \[b0, b1, b2] \ ((f0 \ f1) \ f2) = (\\<^sub>1[b0, b1 \ b2] \ \[b0, b1, b2]) \ ((f0 \ f1) \ f2)" using comp_assoc by simp also have "... = (\\<^sub>1[b0, b1] \ \\<^sub>1[b0 \ b1, b2]) \ ((f0 \ f1) \ f2)" using assms assoc_def by fastforce also have "... = (\\<^sub>1[b0, b1] \ (f0 \ f1)) \ \\<^sub>1[a0 \ a1, a2]" using assms comp_assoc by fastforce also have "... = f0 \ \\<^sub>1[a0, a1] \ \\<^sub>1[a0 \ a1, a2]" using assms comp_assoc by (metis in_homE pr_naturality(2)) also have "... = \\<^sub>1[b0, b1 \ b2] \ (f0 \ (f1 \ f2)) \ \[a0, a1, a2]" proof - have "\\<^sub>1[b0, b1 \ b2] \ (f0 \ (f1 \ f2)) \ \[a0, a1, a2] = (\\<^sub>1[b0, b1 \ b2] \ (f0 \ (f1 \ f2))) \ \[a0, a1, a2]" using comp_assoc by simp also have "... = f0 \ \\<^sub>1[a0, a1 \ a2] \ \[a0, a1, a2]" using assms comp_assoc by fastforce also have "... = f0 \ \\<^sub>1[a0, a1] \ \\<^sub>1[a0 \ a1, a2]" using assms assoc_def by fastforce finally show ?thesis by simp qed finally show ?thesis by blast qed ultimately show ?thesis using assms pr_joint_monic [of b0 "b1 \ b2" "\[b0, b1, b2] \ ((f0 \ f1) \ f2)" "(f0 \ (f1 \ f2)) \ \[a0, a1, a2]"] by fastforce qed lemma pentagon: assumes "ide a" and "ide b" and "ide c" and "ide d" shows "((a \ \[b, c, d]) \ \[a, b \ c, d]) \ (\[a, b, c] \ d) = \[a, b, c \ d] \ \[a \ b, c, d]" proof (intro pr_joint_monic [where h = "((a \ \[b, c, d]) \ \[a, b \ c, d]) \ (\[a, b, c] \ d)" and h' = "\[a, b, c \ d] \ \[a \ b, c, d]"]) show "seq \\<^sub>0[a, b \ (c \ d)] (((a \ \[b, c, d]) \ \[a, b \ c, d]) \ (\[a, b, c] \ d))" using assms by simp show "\\<^sub>1[a, b \ c \ d] \ ((a \ \[b, c, d]) \ \[a, b \ c, d]) \ (\[a, b, c] \ d) = \\<^sub>1[a, b \ c \ d] \ \[a, b, c \ d] \ \[a \ b, c, d]" proof - have "\\<^sub>1[a, b \ c \ d] \ ((a \ \[b, c, d]) \ \[a, b \ c, d]) \ (\[a, b, c] \ d) = ((\\<^sub>1[a, b \ c \ d] \ (a \ \[b, c, d])) \ \[a, b \ c, d]) \ (\[a, b, c] \ d)" using comp_assoc by simp also have "... = (\\<^sub>1[a, (b \ c) \ d] \ \[a, b \ c, d]) \ (\[a, b, c] \ d)" using assms pr_naturality(2) comp_cod_arr by force also have "... = \\<^sub>1[a, b \ c] \ \\<^sub>1[a \ b \ c, d] \ (\[a, b, c] \ d)" using assms assoc_def comp_assoc by simp also have "... = (\\<^sub>1[a, b \ c] \ \[a, b, c]) \ \\<^sub>1[(a \ b) \ c, d]" using assms pr_naturality(2) comp_assoc by fastforce also have "... = \\<^sub>1[a, b] \ \\<^sub>1[a \ b, c] \ \\<^sub>1[(a \ b) \ c, d]" using assms assoc_def comp_assoc by simp finally have "\\<^sub>1[a, b \ c \ d] \ ((a \ \[b, c, d]) \ \[a, b \ c, d]) \ (\[a, b, c] \ d) = \\<^sub>1[a, b] \ \\<^sub>1[a \ b, c] \ \\<^sub>1[(a \ b) \ c, d]" by blast also have "... = \\<^sub>1[a, b \ c \ d] \ \[a, b, c \ d] \ \[a \ b, c, d]" using assms assoc_def comp_assoc by auto finally show ?thesis by blast qed show "\\<^sub>0[a, b \ (c \ d)] \ ((a \ \[b, c, d]) \ \[a, b \ c, d]) \ (\[a, b, c] \ d) = \\<^sub>0[a, b \ (c \ d)] \ \[a, b, c \ d] \ \[a \ b, c, d]" proof - have "\\<^sub>0[a, b \ (c \ d)] \ ((a \ \[b, c, d]) \ \[a, b \ c, d]) \ (\[a, b, c] \ d) = \\<^sub>0[a, b \ c \ d] \ ((a \ \\\<^sub>1[b, c] \ \\<^sub>1[b \ c, d], \\\<^sub>0[b, c] \ \\<^sub>1[b \ c, d], \\<^sub>0[b \ c, d]\\) \ \\\<^sub>1[a, b \ c] \ \\<^sub>1[a \ b \ c, d], \\\<^sub>0[a, b \ c] \ \\<^sub>1[a \ b \ c, d], \\<^sub>0[a \ b \ c, d]\\) \ (\\\<^sub>1[a, b] \ \\<^sub>1[a \ b, c], \\\<^sub>0[a, b] \ \\<^sub>1[a \ b, c], \\<^sub>0[a \ b, c]\\ \ d)" using assms assoc_def by simp also have "... = \\\<^sub>1[b, c] \ \\<^sub>1[b \ c, d], \\\<^sub>0[b, c] \ \\<^sub>1[b \ c, d], \\<^sub>0[b \ c, d]\\ \ (\\<^sub>0[a, (b \ c) \ d] \ \\\<^sub>1[a, b \ c] \ \\<^sub>1[a \ b \ c, d], \\\<^sub>0[a, b \ c] \ \\<^sub>1[a \ b \ c, d], \\<^sub>0[a \ b \ c, d]\\) \ (\\\<^sub>1[a, b] \ \\<^sub>1[a \ b, c], \\\<^sub>0[a, b] \ \\<^sub>1[a \ b, c], \\<^sub>0[a \ b, c]\\ \ d)" proof - have "\\<^sub>0[a, b \ c \ d] \ (a \ \\\<^sub>1[b, c] \ \\<^sub>1[b \ c, d], \\\<^sub>0[b, c] \ \\<^sub>1[b \ c, d], \\<^sub>0[b \ c, d]\\) = \\\<^sub>1[b, c] \ \\<^sub>1[b \ c, d], \\\<^sub>0[b, c] \ \\<^sub>1[b \ c, d], \\<^sub>0[b \ c, d]\\ \ \\<^sub>0[a, (b \ c) \ d]" using assms assoc_def ide_in_hom pr_naturality(1) by auto thus ?thesis using comp_assoc by metis qed also have "... = \\\<^sub>0[a, b] \ \\<^sub>1[a \ b, c] \ \\<^sub>1[(a \ b) \ c, d], \\\<^sub>0[a \ b, c] \ \\<^sub>1[(a \ b) \ c, d], d \ \\<^sub>0[(a \ b) \ c, d]\\" using assms comp_assoc by simp also have "... = \\\<^sub>0[a, b] \ \\<^sub>1[a \ b, c] \ \\<^sub>1[(a \ b) \ c, d], \\\<^sub>0[a \ b, c] \ \\<^sub>1[(a \ b) \ c, d], \\<^sub>0[(a \ b) \ c, d]\\" using assms comp_cod_arr by simp also have "... = \\<^sub>0[a, b \ (c \ d)] \ \[a, b, c \ d] \ \[a \ b, c, d]" using assms assoc_def comp_assoc by simp finally show ?thesis by simp qed qed lemma inverse_arrows_assoc: assumes "ide a" and "ide b" and "ide c" shows "inverse_arrows \[a, b, c] \\<^sup>-\<^sup>1[a, b, c]" using assms assoc_def assoc'_def comp_assoc by (auto simp add: tuple_pr_arr) interpretation CC: product_category C C .. abbreviation Prod where "Prod fg \ fst fg \ snd fg" abbreviation Prod' where "Prod' fg \ snd fg \ fst fg" interpretation \: binary_functor C C C Prod using tuple_ext CC.comp_char interchange apply unfold_locales apply auto by (metis prod_def seqE)+ interpretation Prod': binary_functor C C C Prod' using tuple_ext CC.comp_char interchange apply unfold_locales apply auto by (metis prod_def seqE)+ lemma binary_functor_Prod: shows "binary_functor C C C Prod" and "binary_functor C C C Prod'" .. definition sym ("\[_, _]") where "\[a1, a0] \ if ide a0 \ ide a1 then \\\<^sub>0[a1, a0], \\<^sub>1[a1, a0]\ else null" lemma sym_in_hom [intro]: assumes "ide a" and "ide b" shows "\\[a, b] : a \ b \ b \ a\" using assms sym_def by auto lemma sym_simps [simp]: assumes "ide a" and "ide b" shows "arr \[a, b]" and "dom \[a, b] = a \ b" and "cod \[a, b] = b \ a" using assms sym_in_hom by auto lemma comp_sym_tuple [simp]: assumes "\f0 : a \ b0\" and "\f1 : a \ b1\" shows "\[b0, b1] \ \f0, f1\ = \f1, f0\" using assms sym_def comp_tuple_arr by fastforce lemma prj_sym [simp]: assumes "ide a0" and "ide a1" shows "\\<^sub>0[a1, a0] \ \[a0, a1] = \\<^sub>1[a0, a1]" and "\\<^sub>1[a1, a0] \ \[a0, a1] = \\<^sub>0[a0, a1]" using assms sym_def by auto lemma comp_sym_sym [simp]: assumes "ide a0" and "ide a1" shows "\[a1, a0] \ \[a0, a1] = (a0 \ a1)" using assms sym_def comp_tuple_arr by auto lemma sym_inverse_arrows: assumes "ide a0" and "ide a1" shows "inverse_arrows \[a0, a1] \[a1, a0]" using assms sym_in_hom comp_sym_sym by auto lemma sym_assoc_coherence: assumes "ide a" and "ide b" and "ide c" shows "\[b, c, a] \ \[a, b \ c] \ \[a, b, c] = (b \ \[a, c]) \ \[b, a, c] \ (\[a, b] \ c)" using assms sym_def assoc_def comp_assoc prod_tuple comp_cod_arr by simp lemma sym_naturality: assumes "\f0 : a0 \ b0\" and "\f1 : a1 \ b1\" shows "\[b0, b1] \ (f0 \ f1) = (f1 \ f0) \ \[a0, a1]" using assms sym_def comp_assoc prod_tuple by fastforce abbreviation \ where "\ fg \ \[cod (fst fg), cod (snd fg)] \ (fst fg \ snd fg)" interpretation \: natural_transformation CC.comp C Prod Prod' \ using sym_def CC.arr_char CC.null_char comp_arr_dom comp_cod_arr apply unfold_locales apply auto using arr_cod_iff_arr ideD(1) apply metis using arr_cod_iff_arr ideD(1) apply metis using prod_tuple by simp lemma \_is_natural_transformation: shows "natural_transformation CC.comp C Prod Prod' \" .. abbreviation Diag where "Diag f \ if arr f then (f, f) else CC.null" interpretation \: "functor" C CC.comp Diag by (unfold_locales, auto) lemma functor_Diag: shows "functor C CC.comp Diag" .. interpretation \o\: composite_functor CC.comp C CC.comp Prod Diag .. interpretation \o\: composite_functor C CC.comp C Diag Prod .. abbreviation \ where "\ \ \(f, g). (\\<^sub>1[cod f, cod g] \ (f \ g), \\<^sub>0[cod f, cod g] \ (f \ g))" interpretation \: transformation_by_components CC.comp CC.comp \o\.map CC.map \ using pr_naturality comp_arr_dom comp_cod_arr by unfold_locales auto lemma \_is_natural_transformation: shows "natural_transformation CC.comp CC.comp \o\.map CC.map \" proof - have "\.map = \" using \.map_def ext \.is_extensional comp_arr_dom comp_cod_arr by auto thus "natural_transformation CC.comp CC.comp \o\.map CC.map \" using \.natural_transformation_axioms by simp qed interpretation \: natural_transformation C C map \o\.map dup using dup_naturality comp_arr_dom comp_cod_arr prod_tuple tuple_ext by unfold_locales auto lemma dup_is_natural_transformation: shows "natural_transformation C C map \o\.map dup" .. interpretation \o\o\: composite_functor C CC.comp CC.comp Diag \o\.map .. interpretation \o\o\: composite_functor CC.comp C C Prod \o\.map .. interpretation \o\: natural_transformation C CC.comp Diag \o\o\.map \Diag \ dup\ proof - have "Diag \ map = Diag" by auto thus "natural_transformation C CC.comp Diag \o\o\.map (Diag \ dup)" using \.natural_transformation_axioms \.natural_transformation_axioms o_assoc horizontal_composite [of C C map \o\.map dup CC.comp Diag Diag Diag] by metis qed interpretation \o\: natural_transformation CC.comp C Prod \o\o\.map \dup \ Prod\ using \.natural_transformation_axioms \.natural_transformation_axioms o_assoc horizontal_composite [of CC.comp C Prod Prod Prod C map \o\.map dup] by simp interpretation \o\: natural_transformation C CC.comp \o\o\.map Diag \\.map \ Diag\ using \.natural_transformation_axioms \.natural_transformation_axioms horizontal_composite [of C CC.comp Diag Diag Diag CC.comp \o\.map CC.map \.map] by simp interpretation \o\: natural_transformation CC.comp C \o\o\.map Prod \Prod \ \.map\ proof - have "Prod \ \o\.map = \o\o\.map" by auto thus "natural_transformation CC.comp C \o\o\.map Prod (Prod \ \.map)" using \.natural_transformation_axioms \.natural_transformation_axioms o_assoc horizontal_composite [of CC.comp CC.comp \o\.map CC.map \.map C Prod Prod Prod] by simp qed interpretation \o\_\o\: vertical_composite C CC.comp Diag \o\o\.map Diag \Diag \ dup\ \\.map \ Diag\ .. interpretation \o\_\o\: vertical_composite CC.comp C Prod \o\o\.map Prod \dup \ Prod\ \Prod \ \.map\ .. interpretation \\: unit_counit_adjunction CC.comp C Diag Prod dup \.map proof show "\o\_\o\.map = Diag" proof fix f have "\ arr f \ \o\_\o\.map f = Diag f" by (simp add: \o\_\o\.is_extensional) moreover have "arr f \ \o\_\o\.map f = Diag f" using comp_cod_arr comp_assoc \o\_\o\.map_def by auto ultimately show "\o\_\o\.map f = Diag f" by blast qed show "\o\_\o\.map = Prod" proof fix fg show "\o\_\o\.map fg = Prod fg" proof - have "\ CC.arr fg \ ?thesis" by (simp add: \.is_extensional \o\_\o\.is_extensional) moreover have "CC.arr fg \ ?thesis" proof - assume fg: "CC.arr fg" have 1: "dup (Prod fg) = \cod (fst fg) \ cod (snd fg), cod (fst fg) \ cod (snd fg)\ \ (fst fg \ snd fg)" using fg \.is_natural_2 apply simp by (metis (no_types, lifting) prod_simps(1) prod_simps(3)) have "\o\_\o\.map fg = (\\<^sub>1[cod (fst fg), cod (snd fg)] \ \\<^sub>0[cod (fst fg), cod (snd fg)]) \ \cod (fst fg) \ cod (snd fg), cod (fst fg) \ cod (snd fg)\ \ (fst fg \ snd fg)" using fg 1 \o\_\o\.map_def comp_cod_arr by simp also have "... = ((\\<^sub>1[cod (fst fg), cod (snd fg)] \ \\<^sub>0[cod (fst fg), cod (snd fg)]) \ \cod (fst fg) \ cod (snd fg), cod (fst fg) \ cod (snd fg)\) \ (fst fg \ snd fg)" using comp_assoc by simp also have "... = \\\<^sub>1[cod (fst fg), cod (snd fg)] \ (cod (fst fg) \ cod (snd fg)), \\<^sub>0[cod (fst fg), cod (snd fg)] \ (cod (fst fg) \ cod (snd fg))\ \ (fst fg \ snd fg)" using fg prod_tuple by simp also have "... = Prod fg" using fg comp_arr_dom \.is_natural_2 by auto finally show ?thesis by simp qed ultimately show ?thesis by blast qed qed qed proposition induces_unit_counit_adjunction: shows "unit_counit_adjunction CC.comp C Diag Prod dup \.map" using \\.unit_counit_adjunction_axioms by simp end section "Category with Terminal Object" locale category_with_terminal_object = category + assumes has_terminal: "\t. terminal t" locale elementary_category_with_terminal_object = category C for C :: "'a comp" (infixr "\" 55) and one :: "'a" ("\") and trm :: "'a \ 'a" ("\[_]") + assumes ide_one: "ide \" and trm_in_hom: "ide a \ \\[a] : a \ \\" and trm_eqI: "\ ide a; \f : a \ \\ \ \ f = \[a]" begin lemma trm_simps: assumes "ide a" shows "arr \[a]" and "dom \[a] = a" and "cod \[a] = \" using assms trm_in_hom by auto lemma trm_one: shows "\[\] = \" using ide_one trm_in_hom trm_eqI ide_in_hom by auto lemma terminal_one: shows "terminal \" using ide_one trm_in_hom trm_eqI terminal_def by metis lemma trm_naturality: assumes "arr f" shows "\[cod f] \ f = \[dom f]" using assms trm_eqI by (metis comp_in_homI' ide_cod ide_dom in_homE trm_in_hom) proposition is_category_with_terminal_object: shows "category_with_terminal_object C" apply unfold_locales using terminal_one by auto end context category_with_terminal_object begin definition some_terminal ("\") where "some_terminal \ SOME t. terminal t" definition "trm" ("\[_]") where "\[f] \ if arr f then THE t. \t : dom f \ \\ else null" lemma terminal_some_terminal [intro]: shows "terminal \" using some_terminal_def has_terminal someI_ex [of "\t. terminal t"] by presburger lemma ide_some_terminal: shows "ide \" using terminal_def by blast lemma trm_in_hom [intro]: assumes "arr f" shows "\\[f] : dom f \ \\" proof - have "ide (dom f)" using assms by fastforce hence "\!t. \t : dom f \ \\" using assms trm_def terminal_def terminal_some_terminal by simp thus ?thesis using assms trm_def [of f] theI' [of "\t. \t : dom f \ \\"] by auto qed lemma trm_simps [simp]: assumes "arr f" shows "arr \[f]" and "dom \[f] = dom f" and "cod \[f] = \" using assms trm_in_hom by auto lemma trm_eqI: assumes "\t : dom f \ \\" shows "t = \[f]" proof - have "ide (dom f)" using assms by (metis ide_dom in_homE) hence "\!t. \t : dom f \ \\" using terminal_def [of \] terminal_some_terminal by auto moreover have "\t : dom f \ \\" using assms by simp ultimately show ?thesis using assms trm_def the1_equality [of "\t. \t : dom f \ \\" t] \ide (dom f)\ arr_dom_iff_arr by fastforce qed sublocale elementary_category_with_terminal_object C \ trm using ide_some_terminal trm_eqI by unfold_locales auto proposition extends_to_elementary_category_with_terminal_object: shows "elementary_category_with_terminal_object C \ trm" .. end section "Cartesian Category" locale cartesian_category = category_with_binary_products + category_with_terminal_object locale elementary_cartesian_category = elementary_category_with_binary_products + elementary_category_with_terminal_object begin proposition is_cartesian_category: shows "cartesian_category C" using cartesian_category.intro is_category_with_binary_products is_category_with_terminal_object by auto end context cartesian_category begin proposition extends_to_elementary_cartesian_category: shows "elementary_cartesian_category C pr0 pr1 \ trm" by (simp add: elementary_cartesian_category_def elementary_category_with_terminal_object_axioms extends_to_elementary_category_with_binary_products) sublocale elementary_cartesian_category C pr0 pr1 \ trm using extends_to_elementary_cartesian_category by simp end text \ Here we prove some facts that will later allow us to show that an elementary cartesian category is a monoidal category. \ context elementary_cartesian_category begin abbreviation \ where "\ \ \\<^sub>0[\, \]" lemma pr_coincidence: shows "\ = \\<^sub>1[\, \]" using ide_one by (simp add: terminal_arr_unique terminal_one) lemma \_is_terminal_arr: shows "terminal_arr \" using ide_one by (simp add: terminal_one) lemma inverse_arrows_\: shows "inverse_arrows \ \\, \\" using ide_one by (metis (no_types, lifting) dup_is_natural_transformation \_is_terminal_arr cod_pr0 comp_cod_arr pr_dup(1) ide_dom inverse_arrows_def map_simp natural_transformation.is_natural_2 pr_simps(2) pr1_in_hom' trm_eqI trm_naturality trm_one tuple_pr) lemma \_is_iso: shows "iso \" using inverse_arrows_\ by auto lemma trm_tensor: assumes "ide a" and "ide b" shows "\[a \ b] = \ \ (\[a] \ \[b])" proof - have "\[a \ b] = \[a] \ \\<^sub>1[a, b]" by (metis assms(1-2) cod_pr1 pr_simps(4-6) trm_naturality) moreover have "\\[b] : b \ \\" using assms(2) trm_in_hom by blast ultimately show ?thesis using assms(1) pr_coincidence trm_in_hom by fastforce qed abbreviation runit ("\[_]") where "\[a] \ \\<^sub>1[a, \]" abbreviation runit' ("\\<^sup>-\<^sup>1[_]") where "\\<^sup>-\<^sup>1[a] \ \a, \[a]\" abbreviation lunit ("\[_]") where "\[a] \ \\<^sub>0[\, a]" abbreviation lunit' ("\\<^sup>-\<^sup>1[_]") where "\\<^sup>-\<^sup>1[a] \ \\[a], a\" lemma runit_in_hom: assumes "ide a" shows "\\[a] : a \ \ \ a\" using assms ide_one by simp lemma runit'_in_hom: assumes "ide a" shows "\\\<^sup>-\<^sup>1[a] : a \ a \ \\" using assms ide_in_hom trm_in_hom by blast lemma lunit_in_hom: assumes "ide a" shows "\\[a] : \ \ a \ a\" using assms ide_one by simp lemma lunit'_in_hom: assumes "ide a" shows "\\\<^sup>-\<^sup>1[a] : a \ \ \ a\" using assms ide_in_hom trm_in_hom by blast lemma runit_naturality: assumes "ide a" shows "\[cod a] \ (a \ \) = a \ \[dom a]" using assms pr_naturality(2) ide_char ide_one by blast lemma inverse_arrows_runit: assumes "ide a" shows "inverse_arrows \[a] \\<^sup>-\<^sup>1[a]" proof show "ide (\[a] \ \\<^sup>-\<^sup>1[a])" proof - have "\[a] \ \\<^sup>-\<^sup>1[a] = a" using assms by (metis in_homE ide_char pr_tuple(1) trm_in_hom) thus ?thesis using assms by presburger qed show "ide (\\<^sup>-\<^sup>1[a] \ \[a])" proof - have "ide (a \ \)" using assms ide_one by blast moreover have "\\<^sup>-\<^sup>1[a] \ \[a] = a \ \" proof (intro pr_joint_monic [of a \ "\\<^sup>-\<^sup>1[a] \ \[a]" "a \ \"]) show "seq \\<^sub>0[a, \] (\\<^sup>-\<^sup>1[a] \ \[a])" using assms ide_one runit'_in_hom [of a] by (intro seqI) auto show "\\<^sub>0[a, \] \ \\<^sup>-\<^sup>1[a] \ \[a] = \\<^sub>0[a, \] \ (a \ \)" proof - have "\\<^sub>0[a, \] \ \\<^sup>-\<^sup>1[a] \ \[a] = (\\<^sub>0[a, \] \ \\<^sup>-\<^sup>1[a]) \ \[a]" using comp_assoc by simp also have "... = \[a] \ \[a]" using assms ide_one by (metis in_homE pr_tuple(2) ide_char trm_in_hom) also have "... = \[a \ \]" using assms ide_one trm_naturality [of "\[a]"] by simp also have "... = \\<^sub>0[a, \] \ (a \ \)" using assms comp_arr_dom ide_one trm_naturality trm_one by fastforce finally show ?thesis by blast qed show "\\<^sub>1[a, \] \ \\<^sup>-\<^sup>1[a] \ \[a] = \\<^sub>1[a, \] \ (a \ \)" using assms by (metis \ide (\[a] \ \\<^sup>-\<^sup>1[a])\ cod_comp cod_pr1 dom_comp ide_compE ide_one comp_assoc runit_naturality) qed ultimately show ?thesis by simp qed qed lemma lunit_naturality: assumes "arr f" shows "C \[cod f] (\ \ f) = C f \[dom f]" using assms pr_naturality(1) ide_char ide_one by blast lemma inverse_arrows_lunit: assumes "ide a" shows "inverse_arrows \[a] \\<^sup>-\<^sup>1[a]" proof show "ide (C \[a] \\<^sup>-\<^sup>1[a])" proof - have "C \[a] \\<^sup>-\<^sup>1[a] = a" using assms by (metis ide_char in_homE pr_tuple(2) trm_in_hom) thus ?thesis using assms by simp qed show "ide (\\<^sup>-\<^sup>1[a] \ \[a])" proof - have "\\<^sup>-\<^sup>1[a] \ \[a] = \ \ a" proof (intro pr_joint_monic [of \ a "\\<^sup>-\<^sup>1[a] \ \[a]" "\ \ a"]) show "seq \[a] (\\<^sup>-\<^sup>1[a] \ \[a])" using assms \ide (\[a] \ \\<^sup>-\<^sup>1[a])\ by blast show "\[a] \ \\<^sup>-\<^sup>1[a] \ \[a] = \[a] \ (\ \ a)" using assms by (metis \ide (\[a] \ \\<^sup>-\<^sup>1[a])\ cod_comp cod_pr0 dom_cod ide_compE ide_one comp_assoc lunit_naturality) show "\\<^sub>1[\, a] \ \\<^sup>-\<^sup>1[a] \ \[a] = \\<^sub>1[\, a] \ (\ \ a)" proof - have "\\<^sub>1[\, a] \ \\<^sup>-\<^sup>1[a] \ \[a] = (\\<^sub>1[\, a] \ \\<^sup>-\<^sup>1[a]) \ \[a]" using comp_assoc by simp also have "... = \[a] \ \[a]" using assms ide_one by (metis pr_tuple(1) ide_char in_homE trm_in_hom) also have "... = \[\ \ a]" using assms ide_one trm_naturality [of "\[a]"] by simp also have "... = \\<^sub>1[\, a] \ (\ \ a)" using assms comp_arr_dom ide_one trm_naturality trm_one by fastforce finally show ?thesis by simp qed qed moreover have "ide (\ \ a)" using assms ide_one by simp finally show ?thesis by blast qed qed lemma comp_lunit_term_dup: assumes "ide a" shows "\[a] \ (\[a] \ a) \ \[a] = a" proof - have "\\[a] : a \ \\" using assms trm_in_hom by blast hence "\[a] \ (\[a] \ a) = a \ \\<^sub>0[a, a]" by (metis assms pr_naturality(1) ide_char in_homE) thus ?thesis by (metis (no_types) assms comp_assoc comp_ide_self pr_dup(1)) qed lemma comp_runit_term_dup: assumes "ide a" shows "\[a] \ (a \ \[a]) \ \[a] = a" proof - have "\\[a] : a \ \\" using assms trm_in_hom by blast hence "\[a] \ (a \ \[a]) = a \ \\<^sub>1[a, a]" using assms by auto thus ?thesis using assms by (metis comp_ide_arr pr_dup(2) ide_char comp_assoc seqI) qed lemma comp_proj_assoc: assumes "ide a0" and "ide a1" and "ide a2" shows "\\<^sub>1[a0, a1 \ a2] \ \[a0, a1, a2] = \\<^sub>1[a0, a1] \ \\<^sub>1[a0 \ a1, a2]" and "\\<^sub>0[a0, a1 \ a2] \ \[a0, a1, a2] = \\\<^sub>0[a0, a1] \ \\<^sub>1[a0 \ a1, a2], \\<^sub>0[a0 \ a1, a2]\" using assms assoc_def by auto lemma dup_coassoc: assumes "ide a" shows "\[a, a, a] \ (\[a] \ a) \ \[a] = (a \ \[a]) \ \[a]" proof (intro pr_joint_monic [of a "a \ a" "\[a, a, a] \ (\[a] \ a) \ \[a]" "(a \ \[a]) \ \[a]"]) show "seq \\<^sub>0[a, a \ a] (\[a, a, a] \ (\[a] \ a) \ \[a])" using assms by simp show "\\<^sub>0[a, a \ a] \ \[a, a, a] \ (\[a] \ a) \ \[a] = \\<^sub>0[a, a \ a] \ (a \ \[a]) \ \[a]" proof - have "\\<^sub>0[a, a \ a] \ \[a, a, a] \ (\[a] \ a) \ \[a] = ((\\<^sub>0[a, a \ a] \ \[a, a, a]) \ (\[a] \ a)) \ \[a]" using comp_assoc by simp also have "... = \((\\<^sub>0[a, a] \ \\<^sub>1[a \ a, a]) \ (\[a] \ a)) \ \[a], (a \ \\<^sub>0[a, a]) \ \[a]\" using assms assoc_def by simp also have "... = \[a]" using assms comp_assoc by simp also have "... = (\\<^sub>0[a, a \ a] \ (a \ \[a])) \ \[a]" using assms assoc_def comp_assoc by simp also have "... = \\<^sub>0[a, a \ a] \ (a \ \[a]) \ \[a]" using comp_assoc by simp finally show ?thesis by blast qed show "\\<^sub>1[a, a \ a] \ \[a, a, a] \ (\[a] \ a) \ \[a] = \\<^sub>1[a, a \ a] \ (a \ \[a]) \ \[a]" proof - have "\\<^sub>1[a, a \ a] \ \[a, a, a] \ (\[a] \ a) \ \[a] = ((\\<^sub>1[a, a \ a] \ \[a, a, a]) \ (\[a] \ a)) \ \[a]" using comp_assoc by simp also have "... = ((\\<^sub>1[a, a] \ \\<^sub>1[a \ a, a]) \ (\[a] \ a)) \ \[a]" using assms assoc_def by simp also have "... = a" using assms comp_assoc by simp also have "... = (a \ \\<^sub>1[a, a]) \ \[a]" using assms comp_assoc by simp also have "... = (\\<^sub>1[a, a \ a] \ (a \ \[a])) \ \[a]" using assms by simp also have "... = \\<^sub>1[a, a \ a] \ (a \ \[a]) \ \[a]" using comp_assoc by simp finally show ?thesis by blast qed qed lemma comp_assoc_tuple: assumes "\f0 : a \ b0\" and "\f1 : a \ b1\" and "\f2 : a \ b2\" shows "\[b0, b1, b2] \ \\f0, f1\, f2\ = \f0, \f1, f2\\" and "\\<^sup>-\<^sup>1[b0, b1, b2] \ \f0, \f1, f2\\ = \\f0, f1\, f2\" using assms assoc_def assoc'_def comp_assoc by fastforce+ lemma dup_tensor: assumes "ide a" and "ide b" shows "\[a \ b] = \\<^sup>-\<^sup>1[a, b, a \ b] \ (a \ \[b, a, b]) \ (a \ \ (a, b) \ b) \ (a \ \\<^sup>-\<^sup>1[a, b, b]) \ \[a, a, b \ b] \ (\[a] \ \[b])" proof (intro pr_joint_monic [of "a \ b" "a \ b" "\[a \ b]"]) show "seq \\<^sub>0[a \ b, a \ b] (\[a \ b])" using assms by simp have 1: "\\<^sup>-\<^sup>1[a, b, a \ b] \ (a \ \[b, a, b]) \ (a \ \ (a, b) \ b) \ (a \ \\<^sup>-\<^sup>1[a, b, b]) \ \[a, a, b \ b] \ (\[a] \ \[b]) = \a \ b, a \ b\" proof - have "\\<^sup>-\<^sup>1[a, b, a \ b] \ (a \ \[b, a, b]) \ (a \ \ (a, b) \ b) \ (a \ \\<^sup>-\<^sup>1[a, b, b]) \ \[a, a, b \ b] \ (\[a] \ \[b]) = \\<^sup>-\<^sup>1[a, b, a \ b] \ (a \ \[b, a, b]) \ (a \ \ (a, b) \ b) \ (a \ \\<^sup>-\<^sup>1[a, b, b]) \ \\\<^sub>1[a, b], \\\<^sub>1[a, b], \[b] \ \\<^sub>0[a, b]\\" proof - have "\[a, a, b \ b] \ (\[a] \ \[b]) = \\\<^sub>1[a, b], \\\<^sub>1[a, b], \[b] \ \\<^sub>0[a, b]\\" using assms assoc_def comp_assoc pr_naturality comp_cod_arr by simp thus ?thesis by presburger qed also have "... = \\<^sup>-\<^sup>1[a, b, a \ b] \ \a \ a \ a \ \\<^sub>1[a, b], \[b, a, b] \ (\[a, b] \ (a \ b) \ b) \ \\<^sup>-\<^sup>1[a, b, b] \ \\\<^sub>1[a, b], \[b \ \\<^sub>0[a, b]]\\" using assms prod_tuple by simp also have "... = \\<^sup>-\<^sup>1[a, b, a \ b] \ \\\<^sub>1[a, b], \[b, a, b] \ (\[a, b] \ b) \ \\<^sup>-\<^sup>1[a, b, b] \ \\\<^sub>1[a, b], \[\\<^sub>0[a, b]]\\" proof - have "a \ a \ a \ \\<^sub>1[a, b] = \\<^sub>1[a, b]" using assms comp_cod_arr by simp moreover have "b \ \\<^sub>0[a, b] = \\<^sub>0[a, b]" using assms comp_cod_arr by simp moreover have "\[a, b] \ (a \ b) \ b = \[a, b] \ b" using assms comp_arr_dom by simp ultimately show ?thesis by simp qed also have "... = \\<^sup>-\<^sup>1[a, b, a \ b] \ \\\<^sub>1[a, b], \[b, a, b] \ (\[a, b] \ b) \ \\\\<^sub>1[a, b], \\<^sub>0[a, b]\, \\<^sub>0[a, b]\\" proof - have "\\<^sup>-\<^sup>1[a, b, b] \ \\\<^sub>1[a, b], \[\\<^sub>0[a, b]]\ = \\\\<^sub>1[a, b], \\<^sub>0[a, b]\, \\<^sub>0[a, b]\" using assms comp_assoc_tuple(2) by blast thus ?thesis by simp qed also have "... = \\<^sup>-\<^sup>1[a, b, a \ b] \ \\\<^sub>1[a, b], \[b, a, b] \ \\[a, b], \\<^sub>0[a, b]\\" using assms prod_tuple comp_arr_dom comp_cod_arr by simp also have "... = \\<^sup>-\<^sup>1[a, b, a \ b] \ \\\<^sub>1[a, b], \\\<^sub>0[a, b], \\\<^sub>1[a, b], \\<^sub>0[a, b]\\\" using assms comp_assoc_tuple(1) by (metis sym_def pr_in_hom) also have "... = \\\\<^sub>1[a, b], \\<^sub>0[a, b]\, \\\<^sub>1[a, b], \\<^sub>0[a, b]\\" using assms comp_assoc_tuple(2) by force also have "... = \[a \ b]" using assms by simp finally show ?thesis by simp qed show "\\<^sub>0[a \ b, a \ b] \ \[a \ b] = \\<^sub>0[a \ b, a \ b] \ \\<^sup>-\<^sup>1[a, b, a \ b] \ (a \ \[b, a, b]) \ (a \ \ (a, b) \ b) \ (a \ \\<^sup>-\<^sup>1[a, b, b]) \ \[a, a, b \ b] \ (\[a] \ \[b])" using assms 1 by force show "\\<^sub>1[a \ b, a \ b] \ \[a \ b] = \\<^sub>1[a \ b, a \ b] \ \\<^sup>-\<^sup>1[a, b, a \ b] \ (a \ \[b, a, b]) \ (a \ \ (a, b) \ b) \ (a \ \\<^sup>-\<^sup>1[a, b, b]) \ \[a, a, b \ b] \ (\[a] \ \[b])" using assms 1 by force qed (* TODO: Not sure if the remaining facts are useful. *) lemma \_eq_trm: shows "\ = \[\ \ \]" proof (intro terminal_arr_unique) show "par \ \[\ \ \]" by (simp add: ide_one trm_one trm_tensor) show "terminal_arr \[\ \ \]" using ide_one \_is_terminal_arr \par \ \[\ \ \]\ by auto show "terminal_arr \" using \_is_terminal_arr by blast qed lemma terminal_tensor_one_one: shows "terminal (\ \ \)" proof show "ide (\ \ \)" using ide_one by simp show "\a. ide a \ \!f. \f : a \ \ \ \\" proof - fix a assume a: "ide a" show "\!f. \f : a \ \ \ \\" proof show "\inv \ \ \[a] : a \ \ \ \\" using a ide_one inverse_arrows_\ inverse_unique trm_in_hom by fastforce show "\f. \f : a \ \ \ \\ \ f = inv \ \ \[a]" proof - fix f assume f: "\f : a \ \ \ \\" have "\ \ f = \[a]" proof (intro terminal_arr_unique) show "par (\ \ f) \[a]" using a f by (metis \_is_iso \_is_terminal_arr \\inv \ \ \[a] : a \ \ \ \\\ cod_comp dom_comp dom_inv ide_one in_homE pr_simps(2-3) seqE seqI) show "terminal_arr (\ \ f)" using a f \_is_terminal_arr cod_comp by force show "terminal_arr \[a]" using a \par (\ \ f) \[a]\ \terminal_arr (\ \ f)\ by auto qed thus "f = inv \ \ \[a]" using a f \_is_iso invert_side_of_triangle(1) \\inv \ \ \[a] : a \ \ \ \\\ by blast qed qed qed qed end section "Category with Finite Products" text \ In this last section, we show that the notion ``cartesian category'', which we defined to be a category with binary products and terminal object, coincides with the notion ``category with finite products''. Due to the inability to quantify over types in HOL, we content ourselves with defining the latter notion as "has \I\-indexed products for every finite set \I\ of natural numbers." We can transfer this property to finite sets at other types using the fact that products are preserved under bijections of the index sets. \ locale category_with_finite_products = category C for C :: "'c comp" + assumes has_finite_products: "finite (I :: nat set) \ has_products I" begin lemma has_finite_products': assumes "I \ UNIV" shows "finite I \ has_products I" proof - assume I: "finite I" obtain n \ where \: "bij_betw \ {k. k < (n :: nat)} I" using I finite_imp_nat_seg_image_inj_on inj_on_imp_bij_betw by fastforce show "has_products I" using assms(1) \ has_finite_products has_products_preserved_by_bijection category_with_finite_products.has_finite_products by blast qed end lemma (in category) has_binary_products_if: assumes "has_products ({0, 1} :: nat set)" shows "has_binary_products" proof (unfold has_binary_products_def) show "\a0 a1. ide a0 \ ide a1 \ (\p0 p1. has_as_binary_product a0 a1 p0 p1)" proof (intro allI impI) fix a0 a1 assume 1: "ide a0 \ ide a1" show "\p0 p1. has_as_binary_product a0 a1 p0 p1" proof - interpret J: binary_product_shape by unfold_locales interpret D: binary_product_diagram C a0 a1 using 1 by unfold_locales auto interpret discrete_diagram J.comp C D.map using J.is_discrete by unfold_locales auto show "\p0 p1. has_as_binary_product a0 a1 p0 p1" proof (unfold has_as_binary_product_def) text \ Here we have to work around the fact that \has_finite_products\ is defined in terms of @{typ "nat set"}, whereas \has_as_binary_product\ is defined in terms of \J.arr set\. \ let ?\ = "(\x :: nat. if x = 0 then J.FF else J.TT)" let ?\ = "\j. if j = J.FF then 0 else 1" have "bij_betw ?\ ({0, 1} :: nat set) {J.FF, J.TT}" using bij_betwI [of ?\ "{0, 1} :: nat set" "{J.FF, J.TT}" ?\] by fastforce hence "has_products {J.FF, J.TT}" using assms has_products_def [of "{J.FF, J.TT}"] has_products_preserved_by_bijection [of "{0, 1} :: nat set" ?\ "{J.FF, J.TT}"] by blast hence "\a. has_as_product J.comp D.map a" using has_products_def [of "{J.FF, J.TT}"] discrete_diagram_axioms J.arr_char by blast hence "\a \. product_cone J.comp C D.map a \" using has_as_product_def by blast hence 2: "\a \. D.limit_cone a \" unfolding product_cone_def by simp obtain a \ where \: "D.limit_cone a \" using 2 by auto interpret \: limit_cone J.comp C D.map a \ using \ by auto have "\ = D.mkCone (\ J.FF) (\ J.TT)" proof - have "\a. J.ide a \ \ a = D.mkCone (\ J.FF) (\ J.TT) a" using D.mkCone_def J.ide_char by auto moreover have "a = dom (\ J.FF)" by simp moreover have "D.cone a (D.mkCone (\ (J.MkIde False)) (\ (J.MkIde True)))" using 1 D.cone_mkCone [of "\ J.FF" "\ J.TT"] by auto ultimately show ?thesis using D.mkCone_def \.natural_transformation_axioms D.cone_mkCone [of "\ J.FF" "\ J.TT"] NaturalTransformation.eqI [of "J.comp" C \.A.map "D.map" \ "D.mkCone (\ J.FF) (\ J.TT)"] cone_def [of J.comp C D.map a "D.mkCone (\ J.FF) (\ J.TT)"] J.ide_char by blast qed hence "D.limit_cone (dom (\ J.FF)) (D.mkCone (\ J.FF) (\ J.TT))" using \.limit_cone_axioms by simp thus "\p0 p1. ide a0 \ ide a1 \ D.has_as_binary_product p0 p1" using 1 by blast qed qed qed qed sublocale category_with_finite_products \ category_with_binary_products C using has_binary_products_if has_finite_products by (unfold_locales, unfold has_binary_products_def) simp proposition (in category_with_finite_products) is_category_with_binary_products: shows "category_with_binary_products C" .. sublocale category_with_finite_products \ category_with_terminal_object C proof interpret J: discrete_category "{} :: nat set" by unfold_locales auto interpret D: empty_diagram J.comp C "\j. null" by unfold_locales auto interpret D: discrete_diagram J.comp C "\j. null" using J.is_discrete by unfold_locales auto have "\a. D.has_as_limit a \ has_as_product J.comp (\j. null) a" using product_cone_def J.category_axioms category_axioms D.discrete_diagram_axioms has_as_product_def product_cone_def by metis moreover have "\a. has_as_product J.comp (\j. null) a" using has_finite_products [of "{} :: nat set"] has_products_def [of "{} :: nat set"] D.discrete_diagram_axioms by blast ultimately have "\a. D.has_as_limit a" by blast thus "\a. terminal a" using D.has_as_limit_iff_terminal by blast qed proposition (in category_with_finite_products) is_category_with_terminal_object: shows "category_with_terminal_object C" .. sublocale category_with_finite_products \ cartesian_category .. proposition (in category_with_finite_products) is_cartesian_category: shows "cartesian_category C" .. context category begin lemma binary_product_of_products_is_product: assumes "has_as_product J0 D0 a0" and "has_as_product J1 D1 a1" and "has_as_binary_product a0 a1 p0 p1" and "Collect (partial_magma.arr J0) \ Collect (partial_magma.arr J1) = {}" and "partial_magma.null J0 = partial_magma.null J1" shows "has_as_product (discrete_category.comp (Collect (partial_magma.arr J0) \ Collect (partial_magma.arr J1)) (partial_magma.null J0)) (\i. if i \ Collect (partial_magma.arr J0) then D0 i else if i \ Collect (partial_magma.arr J1) then D1 i else null) (dom p0)" proof - obtain \0 where \0: "product_cone J0 (\) D0 a0 \0" using assms(1) has_as_product_def by blast obtain \1 where \1: "product_cone J1 (\) D1 a1 \1" using assms(2) has_as_product_def by blast interpret J0: category J0 using \0 product_cone.axioms(1) by metis interpret J1: category J1 using \1 product_cone.axioms(1) by metis interpret D0: discrete_diagram J0 C D0 using \0 product_cone.axioms(3) by metis interpret D1: discrete_diagram J1 C D1 using \1 product_cone.axioms(3) by metis interpret \0: product_cone J0 C D0 a0 \0 using \0 by auto interpret \1: product_cone J1 C D1 a1 \1 using \1 by auto interpret J: discrete_category \Collect J0.arr \ Collect J1.arr\ J0.null using J0.not_arr_null assms(5) by unfold_locales auto interpret X: binary_product_shape . interpret a0xa1: binary_product_diagram C a0 a1 using assms(3) has_as_binary_product_def by (simp add: binary_product_diagram.intro binary_product_diagram_axioms.intro category_axioms) have p0p1: "a0xa1.has_as_binary_product p0 p1" using assms(3) has_as_binary_product_def [of a0 a1 p0 p1] by simp let ?D = "(\i. if i \ Collect J0.arr then D0 i else if i \ Collect J1.arr then D1 i else null)" let ?a = "dom p0" let ?\ = "\i. if i \ Collect J0.arr then \0 i \ p0 else if i \ Collect J1.arr then \1 i \ p1 else null" let ?p0p1 = "a0xa1.mkCone p0 p1" interpret p0p1: limit_cone X.comp C a0xa1.map ?a ?p0p1 using p0p1 by simp have a: "ide ?a" using p0p1.ide_apex by simp have p0: "\p0 : ?a \ a0\" using a0xa1.mkCone_def p0p1.preserves_hom [of X.FF X.FF X.FF] X.ide_char X.ide_in_hom by auto have p1: "\p1 : ?a \ a1\" using a0xa1.mkCone_def p0p1.preserves_hom [of X.TT X.TT X.TT] X.ide_char X.ide_in_hom by auto interpret D: discrete_diagram J.comp C ?D using assms J.arr_char J.dom_char J.cod_char J.is_discrete D0.is_discrete D1.is_discrete J.cod_comp J.seq_char by unfold_locales auto interpret A: constant_functor J.comp C ?a using p0p1.ide_apex by unfold_locales simp interpret \: natural_transformation J.comp C A.map ?D ?\ proof fix j show "\ J.arr j \ ?\ j = null" by simp assume j: "J.arr j" have \0j: "J0.arr j \ \\0 j : a0 \ D0 j\" using D0.is_discrete by auto have \1j: "J1.arr j \ \\1 j : a1 \ D1 j\" using D1.is_discrete by auto show "dom (?\ j) = A.map (J.dom j)" using j J.arr_char p0 p1 \0j \1j by fastforce show "cod (?\ j) = ?D (J.cod j)" using j J.arr_char p0 p1 \0j \1j by fastforce show "?D j \ ?\ (J.dom j) = ?\ j" proof - have 0: "J0.arr j \ D0 j \ \0 j \ p0 = \0 j \ p0" proof - have "J0.arr j \ (D0 j \ \0 j) \ p0 = \0 j \ p0" using p0 \0.is_natural_1 \0.is_natural_2 D0.is_discrete by simp thus "J0.arr j \ D0 j \ \0 j \ p0 = \0 j \ p0" using comp_assoc by simp qed have 1: "J1.arr j \ D1 j \ \1 j \ p1 = \1 j \ p1" proof - have "J1.arr j \ (D1 j \ \1 j) \ p1 = \1 j \ p1" using p1 \1.is_natural_1 \1.is_natural_2 D1.is_discrete by simp thus "J1.arr j \ D1 j \ \1 j \ p1 = \1 j \ p1" using comp_assoc by simp qed show ?thesis using 0 1 by auto qed show "?\ (J.cod j) \ A.map j = ?\ j" using j comp_arr_dom p0 p1 comp_assoc by auto qed interpret \: cone J.comp C ?D ?a ?\ .. interpret \: product_cone J.comp C ?D ?a ?\ proof show "\a' \'. D.cone a' \' \ \!f. \f : a' \ ?a\ \ D.cones_map f ?\ = \'" proof - fix a' \' assume \': "D.cone a' \'" interpret \': cone J.comp C ?D a' \' using \' by simp show "\!f. \f : a' \ ?a\ \ D.cones_map f ?\ = \'" proof let ?\0' = "\i. if i \ Collect J0.arr then \' i else null" let ?\1' = "\i. if i \ Collect J1.arr then \' i else null" have 0: "\i. i \ Collect J0.arr \ \' i \ hom a' (D0 i)" using J.arr_char by auto have 1: "\i. i \ Collect J1.arr \ \' i \ hom a' (D1 i)" using J.arr_char `Collect J0.arr \ Collect J1.arr = {}` by force interpret A0': constant_functor J0 C a' apply unfold_locales using \'.ide_apex by auto interpret A1': constant_functor J1 C a' apply unfold_locales using \'.ide_apex by auto interpret \0': cone J0 C D0 a' ?\0' proof (unfold_locales) fix j show "\ J0.arr j \ (if j \ Collect J0.arr then \' j else null) = null" by simp assume j: "J0.arr j" show 0: "dom (?\0' j) = A0'.map (J0.dom j)" using j by simp show 1: "cod (?\0' j) = D0 (J0.cod j)" using j J.arr_char J.cod_char D0.is_discrete by simp show "D0 j \ (?\0' (J0.dom j)) = ?\0' j" using 1 j J.arr_char D0.is_discrete comp_cod_arr by simp show "?\0' (J0.cod j) \ A0'.map j = ?\0' j" using 0 j J.arr_char D0.is_discrete comp_arr_dom by simp qed interpret \1': cone J1 C D1 a' ?\1' proof (unfold_locales) fix j show "\ J1.arr j \ (if j \ Collect J1.arr then \' j else null) = null" by simp assume j: "J1.arr j" show 0: "dom (?\1' j) = A1'.map (J1.dom j)" using j by simp show 1: "cod (?\1' j) = D1 (J1.cod j)" using assms(4) j J.arr_char J.cod_char D1.is_discrete by auto show "D1 j \ (?\1' (J1.dom j)) = ?\1' j" using 1 j J.arr_char D1.is_discrete comp_cod_arr by simp show "?\1' (J1.cod j) \ A1'.map j = ?\1' j" using 0 j J.arr_char D1.is_discrete comp_arr_dom by simp qed define f0 where "f0 = \0.induced_arrow a' ?\0'" define f1 where "f1 = \1.induced_arrow a' ?\1'" have f0: "\f0 : a' \ a0\" using f0_def \0.induced_arrowI \0'.cone_axioms by simp have f1: "\f1 : a' \ a1\" using f1_def \1.induced_arrowI \1'.cone_axioms by simp have 2: "a0xa1.is_rendered_commutative_by f0 f1" using f0 f1 by auto interpret p0p1: binary_product_cone C a0 a1 p0 p1 .. interpret f0f1: cone X.comp C a0xa1.map a' \a0xa1.mkCone f0 f1\ using 2 f0 f1 a0xa1.cone_mkCone [of f0 f1] by auto define f where "f = p0p1.induced_arrow a' (a0xa1.mkCone f0 f1)" have f: "\f : a' \ ?a\" using f_def 2 f0 f1 p0p1.induced_arrowI'(1) by auto moreover have \': "D.cones_map f ?\ = \'" proof fix j show "D.cones_map f ?\ j = \' j" proof (cases "J0.arr j", cases "J1.arr j") show "\J0.arr j; J1.arr j\ \ D.cones_map f ?\ j = \' j" using assms(4) by auto show "\J0.arr j; \ J1.arr j\ \ D.cones_map f ?\ j = \' j" proof - assume J0: "J0.arr j" and J1: "\ J1.arr j" have "D.cones_map f ?\ j = (\0 j \ p0) \ f" using f J0 J1 \.cone_axioms by auto also have "... = \0 j \ p0 \ f" using comp_assoc by simp also have "... = \0 j \ f0" using 2 f0 f1 f_def p0p1.induced_arrowI' by auto also have "... = \' j" proof - have "\0 j \ f0 = \0 j \ \0.induced_arrow' a' \'" unfolding f0_def by simp also have "... = (\j. if J0.arr j then \0 j \ \0.induced_arrow a' (\i. if i \ Collect J0.arr then \' i else null) else null) j" using J0 by simp also have "... = D0.mkCone \' j" proof - have "(\j. if J0.arr j then \0 j \ \0.induced_arrow a' (\i. if i \ Collect J0.arr then \' i else null) else null) = D0.mkCone \'" using f0 f0_def \0.induced_arrowI(2) [of ?\0' a'] J0 D0.mkCone_cone \0'.cone_axioms \0.cone_axioms J0 by auto thus ?thesis by meson qed also have "... = \' j" using J0 by simp finally show ?thesis by blast qed finally show ?thesis by simp qed show "\ J0.arr j \ D.cones_map f ?\ j = \' j" proof (cases "J1.arr j") show "\\ J0.arr j; \ J1.arr j\ \ D.cones_map f ?\ j = \' j" using f \.cone_axioms \'.is_extensional by auto show "\\ J0.arr j; J1.arr j\ \ D.cones_map f ?\ j = \' j" proof - assume J0: "\ J0.arr j" and J1: "J1.arr j" have "D.cones_map f ?\ j = (\1 j \ p1) \ f" using J0 J1 f \.cone_axioms by auto also have "... = \1 j \ p1 \ f" using comp_assoc by simp also have "... = \1 j \ f1" using 2 f0 f1 f_def p0p1.induced_arrowI' by auto also have "... = \' j" proof - have "\1 j \ f1 = \1 j \ \1.induced_arrow' a' \'" unfolding f1_def by simp also have "... = (\j. if J1.arr j then \1 j \ \1.induced_arrow a' (\i. if i \ Collect J1.arr then \' i else null) else null) j" using J1 by simp also have "... = D1.mkCone \' j" proof - have "(\j. if J1.arr j then \1 j \ \1.induced_arrow a' (\i. if i \ Collect J1.arr then \' i else null) else null) = D1.mkCone \'" using f1 f1_def \1.induced_arrowI(2) [of ?\1' a'] J1 D1.mkCone_cone [of a' \'] \1'.cone_axioms \1.cone_axioms J1 by auto thus ?thesis by meson qed also have "... = \' j" using J1 by simp finally show ?thesis by blast qed finally show ?thesis by simp qed qed qed qed ultimately show "\f : a' \ ?a\ \ D.cones_map f ?\ = \'" by blast show "\f'. \f' : a' \ ?a\ \ D.cones_map f' ?\ = \' \ f' = f" proof - fix f' assume f': "\f' : a' \ ?a\ \ D.cones_map f' ?\ = \'" let ?f0' = "p0 \ f'" let ?f1' = "p1 \ f'" have 1: "a0xa1.is_rendered_commutative_by ?f0' ?f1'" using f' p0 p1 p0p1.renders_commutative seqI' by auto have f0': "\?f0' : a' \ a0\" using f' p0 by auto have f1': "\?f1' : a' \ a1\" using f' p1 by auto have "p0 \ f = p0 \ f'" proof - have "D0.cones_map (p0 \ f) \0 = ?\0'" using f p0 \0.cone_axioms \' \.cone_axioms comp_assoc assms(4) seqI' by fastforce moreover have "D0.cones_map (p0 \ f') \0 = ?\0'" using f' p0 \0.cone_axioms \.cone_axioms comp_assoc assms(4) seqI' by fastforce moreover have "p0 \ f = f0" using 2 f0 f_def p0p1.induced_arrowI'(2) by blast ultimately show ?thesis using f0 f0' \0'.cone_axioms \0.is_universal [of a'] by auto qed moreover have "p1 \ f = p1 \ f'" proof - have "D1.cones_map (p1 \ f) \1 = ?\1'" proof fix j show "D1.cones_map (p1 \ f) \1 j = ?\1' j" using f p1 \1.cone_axioms \' \.cone_axioms comp_assoc assms(4) seqI' apply auto by auto qed moreover have "D1.cones_map (p1 \ f') \1 = ?\1'" proof fix j show "D1.cones_map (p1 \ f') \1 j = ?\1' j" using f' p1 \1.cone_axioms \.cone_axioms comp_assoc assms(4) seqI' apply auto by auto qed moreover have "p1 \ f = f1" using 2 f1 f_def p0p1.induced_arrowI'(3) by blast ultimately show ?thesis using f1 f1' \1'.cone_axioms \1.is_universal [of a'] by auto qed ultimately show "f' = f" using f f' p0p1.is_universal' [of a'] by (metis (no_types, lifting) "1" dom_comp in_homE p0p1.is_universal' p1 seqI') qed qed qed qed show "has_as_product J.comp ?D ?a" unfolding has_as_product_def using \.product_cone_axioms by auto qed end sublocale cartesian_category \ category_with_finite_products proof obtain t where t: "terminal t" using has_terminal by blast { fix n :: nat have "\I :: nat set. finite I \ card I = n \ has_products I" proof (induct n) show "\I :: nat set. finite I \ card I = 0 \ has_products I" proof - fix I :: "nat set" assume "finite I \ card I = 0" hence I: "I = {}" by force thus "has_products I" proof - interpret J: discrete_category I 0 apply unfold_locales using I by auto have "\D. discrete_diagram J.comp C D \ \a. has_as_product J.comp D a" proof - fix D assume D: "discrete_diagram J.comp C D" interpret D: discrete_diagram J.comp C D using D by auto interpret D: empty_diagram J.comp C D using I J.arr_char by unfold_locales simp have "has_as_product J.comp D t" using t D.has_as_limit_iff_terminal has_as_product_def product_cone_def J.category_axioms category_axioms D.discrete_diagram_axioms by metis thus "\a. has_as_product J.comp D a" by blast qed moreover have "I \ UNIV" using I by blast ultimately show ?thesis using I has_products_def by (metis category_with_terminal_object.has_terminal discrete_diagram.product_coneI discrete_diagram_def empty_diagram.has_as_limit_iff_terminal empty_diagram.intro empty_diagram_axioms.intro empty_iff has_as_product_def is_category_with_terminal_object mem_Collect_eq) qed qed show "\n I :: nat set. \ (\I :: nat set. finite I \ card I = n \ has_products I); finite I \ card I = Suc n \ \ has_products I" proof - fix n :: nat fix I :: "nat set" assume IH: "\I :: nat set. finite I \ card I = n \ has_products I" assume I: "finite I \ card I = Suc n" show "has_products I" proof - have "card I = 1 \ has_products I" using I has_unary_products by blast moreover have "card I \ 1 \ has_products I" proof - assume "card I \ 1" hence cardI: "card I > 1" using I by simp obtain i where i: "i \ I" using cardI by fastforce let ?I0 = "{i}" and ?I1 = "I - {i}" have 1: "I = ?I0 \ ?I1 \ ?I0 \ ?I1 = {} \ card ?I0 = 1 \ card ?I1 = n" using i I cardI by auto show "has_products I" proof (unfold has_products_def, intro conjI allI impI) show "I \ UNIV" using I by auto fix J D assume D: "discrete_diagram J C D \ Collect (partial_magma.arr J) = I" interpret D: discrete_diagram J C D using D by simp have Null: "D.J.null \ ?I0 \ D.J.null \ ?I1" using D D.J.not_arr_null i by blast interpret J0: discrete_category ?I0 D.J.null using 1 Null D by unfold_locales auto interpret J1: discrete_category ?I1 D.J.null using Null by unfold_locales auto interpret J0uJ1: discrete_category \Collect J0.arr \ Collect J1.arr\ J0.null using Null 1 J0.null_char J1.null_char by unfold_locales auto interpret D0: discrete_diagram_from_map ?I0 C D D.J.null using 1 J0.ide_char D.preserves_ide D D.is_discrete i by unfold_locales auto interpret D1: discrete_diagram_from_map ?I1 C D D.J.null using 1 J1.ide_char D.preserves_ide D D.is_discrete i by unfold_locales auto obtain a0 where a0: "has_as_product J0.comp D0.map a0" using 1 has_unary_products [of ?I0] has_products_def [of ?I0] D0.discrete_diagram_axioms by fastforce obtain a1 where a1: "has_as_product J1.comp D1.map a1" using 1 I IH [of ?I1] has_products_def [of ?I1] D1.discrete_diagram_axioms by blast have 2: "\p0 p1. has_as_binary_product a0 a1 p0 p1" proof - have "ide a0 \ ide a1" using a0 a1 product_is_ide by auto thus ?thesis using a0 a1 has_binary_products has_binary_products_def by simp qed obtain p0 p1 where a: "has_as_binary_product a0 a1 p0 p1" using 2 by auto let ?a = "dom p0" have "has_as_product J D ?a" proof - have "D = (\j. if j \ Collect J0.arr then D0.map j else if j \ Collect J1.arr then D1.map j else null)" proof fix j show "D j = (if j \ Collect J0.arr then D0.map j else if j \ Collect J1.arr then D1.map j else null)" using 1 D0.map_def D1.map_def D.is_extensional D J0.arr_char J1.arr_char by auto qed moreover have "J = J0uJ1.comp" proof - have "\j j'. J j j' = J0uJ1.comp j j'" proof - fix j j' show "J j j' = J0uJ1.comp j j'" using D J0uJ1.arr_char J0.arr_char J1.arr_char D.is_discrete i apply (cases "j \ ?I0", cases "j' \ ?I0") apply simp_all apply auto[1] apply (metis D.J.comp_arr_ide D.J.comp_ide_arr D.J.ext D.J.seqE D.is_discrete J0.null_char J0uJ1.null_char) by (metis D.J.comp_arr_ide D.J.comp_ide_arr D.J.comp_ide_self D.J.ext D.J.seqE D.is_discrete J0.null_char J0uJ1.null_char mem_Collect_eq) qed thus ?thesis by blast qed moreover have "Collect J0.arr \ Collect J1.arr = {}" by auto moreover have "J0.null = J1.null" using J0.null_char J1.null_char by simp ultimately show "has_as_product J D ?a" using binary_product_of_products_is_product [of J0.comp D0.map a0 J1.comp D1.map a1 p0 p1] J0.arr_char J1.arr_char 1 a0 a1 a by simp qed thus "\a. has_as_product J D a" by blast qed qed ultimately show "has_products I" by blast qed qed qed } hence 1: "\n I :: nat set. finite I \ card I = n \ has_products I" by simp thus "\I :: nat set. finite I \ has_products I" by blast qed proposition (in cartesian_category) is_category_with_finite_products: shows "category_with_finite_products C" .. end diff --git a/thys/Category3/CategoryWithFiniteLimits.thy b/thys/Category3/CategoryWithFiniteLimits.thy --- a/thys/Category3/CategoryWithFiniteLimits.thy +++ b/thys/Category3/CategoryWithFiniteLimits.thy @@ -1,424 +1,424 @@ (* Title: CategoryWithFiniteLimits Author: Eugene W. Stark , 2020 Maintainer: Eugene W. Stark *) chapter "Category with Finite Limits" theory CategoryWithFiniteLimits -imports Category3.CartesianCategory Category3.CategoryWithPullbacks +imports CartesianCategory CategoryWithPullbacks begin text\ In this chapter we define ``category with finite limits'' and show that such categories coincide with those having pullbacks and a terminal object. Since we can't quantify over types in HOL, the best we can do at defining the notion ``category with finite limits'' is to state it for a fixed choice of type (e.g.~@{typ nat}) for the arrows of the ``diagram shape''. However, we then have to go to some trouble to show the existence of finite limits for diagram shapes at other types. \ locale category_with_finite_limits = category + assumes has_finite_limits: "\ category (J :: nat comp); finite (Collect (partial_magma.arr J)) \ \ has_limits_of_shape J" begin text\ We show that a category with finite limits has pullbacks and a terminal object and is therefore also a cartesian category. \ interpretation category_with_pullbacks C proof - interpret J: cospan_shape by unfold_locales have 1: "finite (Collect J.arr)" proof - have "Collect J.arr = {J.AA, J.BB, J.TT, J.AT, J.BT}" using J.arr_char cospan_shape.Dom.cases by auto thus ?thesis by simp qed obtain J' :: "nat comp" where J': "isomorphic_categories J' J.comp" using 1 J.finite_imp_ex_iso_nat_comp by blast interpret J'J: isomorphic_categories J' J.comp using J' by simp obtain \ \ where \\: "inverse_functors J.comp J' \ \" using J'J.iso inverse_functors_sym by blast interpret \\: inverse_functors J.comp J' \ \ using \\ by simp interpret \: invertible_functor J.comp J' \ using \\.inverse_functors_axioms by unfold_locales auto show "category_with_pullbacks C" proof show "has_pullbacks" proof (unfold has_pullbacks_def has_as_pullback_def, intro allI impI) fix f0 f1 assume cospan: "cospan f0 f1" interpret D: cospan_diagram C f0 f1 using cospan by (simp add: category_axioms cospan_diagram_axioms_def cospan_diagram_def) have 2: "has_limits_of_shape J.comp" using 1 bij_betw_finite J'J.A.category_axioms has_finite_limits \.bij_betw_arr_sets has_limits_preserved_by_isomorphism J'J.isomorphic_categories_axioms by blast obtain a \ where \: "limit_cone J.comp C D.map a \" using 2 D.diagram_axioms has_limits_of_shape_def by blast interpret \: limit_cone J.comp C D.map a \ using \ by simp have "D.map = cospan_diagram.map C f0 f1" by simp moreover have "a = dom (\ J.AA)" using J.arr_char \.component_in_hom by force moreover have "\ = cospan_diagram.mkCone (\) f0 f1 (\ J.AA) (\ J.BB)" using D.mkCone_cone \.cone_axioms by auto ultimately have "limit_cone (\\<^sub>J) (\) (cospan_diagram.map (\) f0 f1) (dom (\ J.AA)) (cospan_diagram.mkCone (\) f0 f1 (\ J.AA) (\ J.BB))" using \.limit_cone_axioms by simp thus "\p0 p1. cospan f0 f1 \ limit_cone (\\<^sub>J) (\) (cospan_diagram.map (\) f0 f1) (dom p0) (cospan_diagram.mkCone (\) f0 f1 p0 p1)" using cospan by auto qed qed qed lemma is_category_with_pullbacks: shows "category_with_pullbacks C" .. sublocale category_with_pullbacks C .. interpretation category_with_terminal_object C proof show "\a. terminal a" proof - interpret J: discrete_category \{} :: nat set\ 0 by unfold_locales simp have 1: "has_limits_of_shape J.comp" using has_finite_limits by (metis Collect_empty_eq J.arr_char J.is_category empty_iff finite.emptyI) interpret D: diagram J.comp C \\_. null\ by unfold_locales auto obtain t \ where \: "D.limit_cone t \" using 1 D.diagram_axioms has_limits_of_shape_def by blast interpret \: limit_cone J.comp C \\_. null\ t \ using \ by simp have "terminal t" proof show "ide t" using \.ide_apex by simp fix a assume a: "ide a" show "\!f. \f : a \ t\" proof - interpret a: constant_functor J.comp C a using a by unfold_locales interpret \: cone J.comp C \\_.null\ a \\_.null\ apply unfold_locales apply simp using dom_null cod_null comp_null by blast+ have "\!f. \f : a \ t\ \ D.cones_map f \ = (\_. null)" using \.induced_arrowI [of "\_.null" a] \.cone_axioms \.is_universal [of a "\_. null"] by simp moreover have "\f. \f : a \ t\ \ D.cones_map f \ = (\_. null)" using \.cone_axioms by auto ultimately show ?thesis by auto qed qed thus ?thesis by blast qed qed lemma is_category_with_terminal_object: shows "category_with_terminal_object C" .. sublocale category_with_terminal_object C .. sublocale category_with_finite_products using has_finite_limits has_finite_products_if_has_finite_limits has_limits_of_shape_def diagram_def by unfold_locales blast sublocale cartesian_category .. end locale category_with_pullbacks_and_terminal = category_with_pullbacks + category_with_terminal_object sublocale category_with_finite_limits \ category_with_pullbacks_and_terminal .. text\ Conversely, we show that a category with pullbacks and a terminal object also has finite products and equalizers, and therefore has finite limits. \ context category_with_pullbacks_and_terminal begin interpretation ECP: elementary_category_with_pullbacks C prj0 prj1 using extends_to_elementary_category_with_pullbacks by simp abbreviation prj0' where "prj0' a b \ (if ide a \ ide b then prj0 (trm a) (trm b) else null)" abbreviation prj1' where "prj1' a b \ (if ide a \ ide b then prj1 (trm a) (trm b) else null)" interpretation ECC: elementary_cartesian_category C prj0' prj1' \ trm using trm_naturality ECP.universal by unfold_locales auto interpretation category_with_equalizers C proof (unfold_locales, unfold has_equalizers_def, intro allI impI) fix f0 f1 assume par: "par f0 f1" interpret J: parallel_pair by unfold_locales interpret D: parallel_pair_diagram C f0 f1 using par by unfold_locales auto have 1: "cospan (ECC.prod f1 (dom f0)) (ECC.prod f0 (dom f0))" using par by simp let ?g0 = "ECC.prod f0 (dom f0) \ ECC.dup (dom f0)" let ?g1 = "ECC.prod f1 (dom f1) \ ECC.dup (dom f1)" have g0: "\?g0 : dom f0 \ ECC.prod (cod f0) (dom f0)\" using par by simp have g1: "\?g1 : dom f1 \ ECC.prod (cod f1) (dom f1)\" using par by simp define e0 where "e0 = prj0 ?g1 ?g0" define e1 where "e1 = prj1 ?g1 ?g0" have e0: "\e0 : dom e0 \ dom f0\" using par 1 e0_def by auto have e1: "\e1 : dom e0 \ dom f1\" using par 1 e1_def e0_def by auto have eq: "e0 = e1" proof - have "e1 = prj0' (cod f1) (dom f1) \ ?g1 \ e1" proof - have "((prj0' (cod f1) (dom f1) \ (ECC.prod f1 (dom f1))) \ ECC.dup (dom f1)) \ e1 = dom f1 \ e1" using par ECC.pr_naturality(1) [of "dom f1" "dom f1" "dom f1" f1 "dom f1" "cod f1"] comp_cod_arr ECC.pr_dup(1) by auto also have "... = e1" using par e1 comp_cod_arr by blast finally show ?thesis using comp_assoc by simp qed also have "... = prj0' (cod f1) (dom f1) \ ?g0 \ e0" using par ECP.pullback_commutes unfolding commutative_square_def e0_def e1_def by simp also have "... = e0" proof - have "((prj0' (cod f1) (dom f1) \ (ECC.prod f0 (dom f0))) \ ECC.dup (dom f0)) \ e0 = dom f0 \ e0" using par ECC.pr_naturality(1) [of "dom f0" "dom f0" "dom f1" f0 "dom f0" "cod f0"] comp_cod_arr ECC.pr_dup(1) ide_dom by auto also have "... = e0" using e0 comp_cod_arr by blast finally show ?thesis using comp_assoc by simp qed finally show ?thesis by auto qed have equalizes: "D.is_equalized_by e0" proof show "seq f0 e0" using par e0 by auto show "f0 \ e0 = f1 \ e0" proof - have "f0 \ e0 = (f0 \ dom f0) \ e0" using par comp_arr_dom by simp also have "... = (f0 \ (prj1' (dom f0) (dom f0) \ ECC.dup (dom f0))) \ e0" using par ECC.pr_dup(2) by auto also have "... = ((f0 \ prj1' (dom f0) (dom f0)) \ ECC.dup (dom f0)) \ e0" using comp_assoc by auto also have "... = prj1' (cod f1) (dom f1) \ ?g0 \ e0" using par ECC.pr_naturality(2) [of "dom f0" "dom f0" "dom f1" f0 "dom f0" "cod f0"] by (metis (no_types, lifting) arr_dom cod_dom dom_dom comp_assoc) also have "... = prj1' (cod f1) (dom f1) \ ?g1 \ e1" using par ECP.pullback_commutes [of ?g1 ?g0] unfolding commutative_square_def e0_def e1_def by simp also have "... = (prj1' (cod f1) (dom f1) \ ?g1) \ e1" using comp_assoc by simp also have "... = (f1 \ (prj1' (dom f1) (dom f1) \ ECC.dup (dom f1))) \ e1" using par ECC.pr_naturality(2) [of "dom f1" "dom f1" "dom f1" f1 "dom f1" "cod f1"] by (metis (no_types, lifting) arr_dom cod_dom dom_dom comp_assoc) also have "... = (f1 \ dom f1) \ e1" using par ECC.pr_dup(2) by auto also have "... = f1 \ e1" using par comp_arr_dom by simp also have "... = f1 \ e0" using eq by simp finally show ?thesis by simp qed qed show "\e. has_as_equalizer f0 f1 e" proof interpret E: constant_functor J.comp C \dom e0\ using par e0 by unfold_locales auto interpret \: cone J.comp C D.map \dom e0\ \D.mkCone e0\ using equalizes D.cone_mkCone e0_def by auto interpret \: limit_cone J.comp C D.map \dom e0\ \D.mkCone e0\ proof show "\a' \'. D.cone a' \' \ \!f. \f : a' \ dom e0\ \ D.cones_map f (D.mkCone e0) = \'" proof - fix a' \' assume \': "D.cone a' \'" interpret \': cone J.comp C D.map a' \' using \' by simp have 3: "commutative_square ?g1 ?g0 (\' J.Zero) (\' J.Zero)" proof show "cospan ?g1 ?g0" using par g0 g1 by simp show 4: "span (\' J.Zero) (\' J.Zero)" using J.arr_char by simp show 5: "dom ?g1 = cod (\' J.Zero)" using par g1 J.arr_char D.map_def by simp show "?g1 \ \' J.Zero = ?g0 \ \' J.Zero" proof - have "?g1 \ \' J.Zero = ECC.prod f1 (dom f1) \ ECC.dup (dom f1) \ \' J.Zero" using comp_assoc by simp also have "... = ECC.prod f1 (dom f1) \ ECC.tuple (\' J.Zero) (\' J.Zero)" using par D.map_def J.arr_char comp_cod_arr by auto also have "... = ECC.tuple (f1 \ \' J.Zero) (\' J.Zero)" using par ECC.prod_tuple [of "\' J.Zero" "\' J.Zero" f1 "dom f1"] comp_cod_arr by (metis (no_types, lifting) 4 5 g1 in_homE seqI) also have "... = ECC.tuple (f0 \ \' J.Zero) (\' J.Zero)" using par D.is_equalized_by_cone \'.cone_axioms by auto also have "... = ECC.prod f0 (dom f0) \ ECC.tuple (\' J.Zero) (\' J.Zero)" using par ECC.prod_tuple [of "\' J.Zero" "\' J.Zero" f0 "dom f0"] comp_cod_arr by (metis (no_types, lifting) 4 5 g1 in_homE seqI) also have "... = ECC.prod f0 (dom f0) \ ECC.dup (dom f0) \ \' J.Zero" using par D.map_def J.arr_char comp_cod_arr by auto also have "... = ?g0 \ \' J.Zero" using comp_assoc by simp finally show ?thesis by blast qed qed show "\!f. \f : a' \ dom e0\ \ D.cones_map f (D.mkCone e0) = \'" proof define f where "f = ECP.tuple (\' J.Zero) ?g1 ?g0 (\' J.Zero)" have 4: "e0 \ f = \' J.Zero" using ECP.universal by (simp add: "3" e1_def eq f_def) have f: "\f : a' \ dom e0\" proof - have "a' = dom (\' J.Zero)" by (simp add: J.arr_char) thus ?thesis using 3 f_def e0_def g0 g1 ECP.tuple_in_hom ECP.pbdom_def by simp qed moreover have 5: "D.cones_map f (D.mkCone e0) = \'" proof - have "\j. J.arr j \ D.mkCone e0 j \ f = \' j" proof - fix j assume j: "J.arr j" show "D.mkCone e0 j \ f = \' j" proof (cases "j = J.Zero") case True moreover have "e0 \ f = \' J.Zero" using 4 by simp ultimately show ?thesis unfolding f_def D.mkCone_def comp_assoc using J.arr_char by simp next case F: False hence 1: "(f0 \ e0) \ f = f0 \ \' J.Zero" using 4 comp_assoc by simp also have "... = \' j" by (metis (no_types, lifting) F D.mkCone_cone D.mkCone_def \'.cone_axioms j) finally show ?thesis by (simp add: F D.mkCone_def j) qed qed thus ?thesis using f e0 \.cone_axioms \'.is_extensional by auto qed ultimately show "\f : a' \ dom e0\ \ D.cones_map f (D.mkCone e0) = \'" by simp fix f' assume f': "\f' : a' \ dom e0\ \ D.cones_map f' (D.mkCone e0) = \'" show "f' = f" proof - have "e0 \ f' = \' J.Zero" using f' D.mkCone_cone D.mkCone_def \'.cone_axioms comp_assoc J.arr_char \.cone_axioms by auto thus ?thesis using f' 3 4 eq ECP.universal [of ?g1 ?g0 "e1 \ f'" "e0 \ f'"] e0_def e1_def by (metis (no_types, lifting)) qed qed qed qed show "has_as_equalizer f0 f1 e0" proof - have "par f0 f1" by fact moreover have "D.has_as_equalizer e0" .. ultimately show ?thesis using has_as_equalizer_def by blast qed qed qed interpretation category_with_finite_products C by (simp add: ECC.is_cartesian_category cartesian_category.is_category_with_finite_products) lemma has_finite_products: shows "category_with_finite_products C" .. lemma has_finite_limits: shows "category_with_finite_limits C" proof fix J :: "nat comp" assume J: "category J" interpret J: category J using J by simp assume finite: "finite (Collect J.arr)" show "has_limits_of_shape J" proof - have "Collect (partial_magma.ide J) \ Collect J.arr" by auto hence 1: "finite (Collect J.ide)" using finite finite_subset by blast have "has_products (Collect (partial_magma.ide J))" using 1 J.ideD(1) J.not_arr_null category_with_finite_products.has_finite_products is_category_with_finite_products by simp moreover have "Collect (partial_magma.ide J) \ UNIV" using J.not_arr_null by blast moreover have "Collect (partial_magma.arr J) \ UNIV" using J.not_arr_null by blast ultimately show ?thesis using finite 1 J.category_axioms has_limits_if_has_products has_finite_products' [of "Collect J.ide"] has_finite_products' [of "Collect J.arr"] by simp qed qed sublocale category_with_finite_limits C using has_finite_limits by simp end end diff --git a/thys/Category3/CategoryWithPullbacks.thy b/thys/Category3/CategoryWithPullbacks.thy --- a/thys/Category3/CategoryWithPullbacks.thy +++ b/thys/Category3/CategoryWithPullbacks.thy @@ -1,1114 +1,1114 @@ (* Title: CategoryWithPullbacks Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) chapter "Category with Pullbacks" theory CategoryWithPullbacks -imports Category3.Limit +imports Limit begin text \ \sloppypar In this chapter, we give a traditional definition of pullbacks in a category as limits of cospan diagrams and we define a locale \category_with_pullbacks\ that is satisfied by categories in which every cospan diagram has a limit. These definitions build on the general definition of limit that we gave in @{theory Category3.Limit}. We then define a locale \elementary_category_with_pullbacks\ that axiomatizes categories equipped with chosen functions that assign to each cospan a corresponding span of ``projections'', which enjoy the familiar universal property of a pullback. After developing consequences of the axioms, we prove that the two locales are in agreement, in the sense that every interpretation of \category_with_pullbacks\ extends to an interpretation of \elementary_category_with_pullbacks\, and conversely, the underlying category of an interpretation of \elementary_category_with_pullbacks\ always yields an interpretation of \category_with_pullbacks\. \ section "Commutative Squares" context category begin text \ The following provides some useful technology for working with commutative squares. \ definition commutative_square where "commutative_square f g h k \ cospan f g \ span h k \ dom f = cod h \ f \ h = g \ k" lemma commutative_squareI [intro, simp]: assumes "cospan f g" and "span h k" and "dom f = cod h" and "f \ h = g \ k" shows "commutative_square f g h k" using assms commutative_square_def by auto lemma commutative_squareE [elim]: assumes "commutative_square f g h k" and "\ arr f; arr g; arr h; arr k; cod f = cod g; dom h = dom k; dom f = cod h; dom g = cod k; f \ h = g \ k \ \ T" shows T using assms commutative_square_def by (metis (mono_tags, lifting) seqE seqI) lemma commutative_square_comp_arr: assumes "commutative_square f g h k" and "seq h l" shows "commutative_square f g (h \ l) (k \ l)" using assms apply (elim commutative_squareE, intro commutative_squareI, auto) using comp_assoc by metis lemma arr_comp_commutative_square: assumes "commutative_square f g h k" and "seq l f" shows "commutative_square (l \ f) (l \ g) h k" using assms comp_assoc by (elim commutative_squareE, intro commutative_squareI, auto) end section "Cospan Diagrams" (* TODO: Rework the ugly development of equalizers into this form. *) text \ The ``shape'' of a cospan diagram is a category having two non-identity arrows with distinct domains and a common codomain. \ locale cospan_shape begin datatype Arr = Null | AA | BB | TT | AT | BT fun comp where "comp AA AA = AA" | "comp AT AA = AT" | "comp TT AT = AT" | "comp BB BB = BB" | "comp BT BB = BT" | "comp TT BT = BT" | "comp TT TT = TT" | "comp _ _ = Null" interpretation partial_magma comp proof show "\!n. \f. comp n f = n \ comp f n = n" proof show "\f. comp Null f = Null \ comp f Null = Null" by simp show "\n. \f. comp n f = n \ comp f n = n \ n = Null" by (metis comp.simps(8)) qed qed lemma null_char: shows "null = Null" proof - have "\f. comp Null f = Null \ comp f Null = Null" by simp thus ?thesis using null_def ex_un_null theI [of "\n. \f. comp n f = n \ comp f n = n"] by (metis partial_magma.comp_null(2) partial_magma_axioms) qed lemma ide_char: shows "ide f \ f = AA \ f = BB \ f = TT" proof show "ide f \ f = AA \ f = BB \ f = TT" using ide_def null_char by (cases f, simp_all) show "f = AA \ f = BB \ f = TT \ ide f" proof - have 1: "\f g. f = AA \ f = BB \ f = TT \ comp f f \ Null \ (comp g f \ Null \ comp g f = g) \ (comp f g \ Null \ comp f g = g)" proof - fix f g show "f = AA \ f = BB \ f = TT \ comp f f \ Null \ (comp g f \ Null \ comp g f = g) \ (comp f g \ Null \ comp f g = g)" by (cases f; cases g, auto) qed assume f: "f = AA \ f = BB \ f = TT" show "ide f" using f 1 ide_def null_char by simp qed qed fun Dom where "Dom AA = AA" | "Dom BB = BB" | "Dom TT = TT" | "Dom AT = AA" | "Dom BT = BB" | "Dom _ = Null" fun Cod where "Cod AA = AA" | "Cod BB = BB" | "Cod TT = TT" | "Cod AT = TT" | "Cod BT = TT" | "Cod _ = Null" lemma domains_char': shows "domains f = (if f = Null then {} else {Dom f})" using domains_def ide_char null_char by (cases f, auto) lemma codomains_char': shows "codomains f = (if f = Null then {} else {Cod f})" using codomains_def ide_char null_char by (cases f, auto) lemma arr_char: shows "arr f \ f \ Null" using arr_def domains_char' codomains_char' by simp lemma seq_char: shows "seq g f \ (f = AA \ (g = AA \ g = AT)) \ (f = BB \ (g = BB \ g = BT)) \ (f = AT \ g = TT) \ (f = BT \ g = TT) \ (f = TT \ g = TT)" using arr_char null_char by (cases f; cases g, simp_all) interpretation category comp proof fix f g h show "comp g f \ null \ seq g f" using null_char arr_char seq_char by simp show "domains f \ {} \ codomains f \ {}" using domains_char' codomains_char' by auto show "seq h g \ seq (comp h g) f \ seq g f" using seq_char arr_char by (cases g; cases h; simp_all) show "seq h (comp g f) \ seq g f \ seq h g" using seq_char arr_char by (cases f; cases g; simp_all) show "seq g f \ seq h g \ seq (comp h g) f" using seq_char arr_char by (cases f; simp_all; cases g; simp_all; cases h; auto) show "seq g f \ seq h g \ comp (comp h g) f = comp h (comp g f)" using seq_char by (cases f; simp_all; cases g; simp_all; cases h; auto) qed lemma is_category: shows "category comp" .. lemma dom_char: shows "dom = Dom" using dom_def domains_char domains_char' null_char by fastforce lemma cod_char: shows "cod = Cod" using cod_def codomains_char codomains_char' null_char by fastforce end sublocale cospan_shape \ category comp using is_category by auto locale cospan_diagram = J: cospan_shape + C: category C for C :: "'c comp" (infixr "\" 55) and f0 :: 'c and f1 :: 'c + assumes is_cospan: "C.cospan f0 f1" begin no_notation J.comp (infixr "\" 55) notation J.comp (infixr "\\<^sub>J" 55) fun map where "map J.AA = C.dom f0" | "map J.BB = C.dom f1" | "map J.TT = C.cod f0" | "map J.AT = f0" | "map J.BT = f1" | "map _ = C.null" end sublocale cospan_diagram \ diagram J.comp C map proof show "\f. \ J.arr f \ map f = C.null" using J.arr_char by simp fix f assume f: "J.arr f" show "C.arr (map f)" using f J.arr_char is_cospan by (cases f, simp_all) show "C.dom (map f) = map (J.dom f)" using f J.arr_char J.dom_char is_cospan by (cases f, simp_all) show "C.cod (map f) = map (J.cod f)" using f J.arr_char J.cod_char is_cospan by (cases f, simp_all) next fix f g assume fg: "J.seq g f" show "map (g \\<^sub>J f) = map g \ map f" using fg J.seq_char J.null_char J.not_arr_null is_cospan apply (cases f; cases g, simp_all) using C.comp_arr_dom C.comp_cod_arr by auto qed section "Category with Pullbacks" text \ A \emph{pullback} in a category @{term C} is a limit of a cospan diagram in @{term C}. \ context cospan_diagram begin definition mkCone where "mkCone p0 p1 \ \j. if j = J.AA then p0 else if j = J.BB then p1 else if j = J.AT then f0 \ p0 else if j = J.BT then f1 \ p1 else if j = J.TT then f0 \ p0 else C.null" abbreviation is_rendered_commutative_by where "is_rendered_commutative_by p0 p1 \ C.seq f0 p0 \ f0 \ p0 = f1 \ p1" abbreviation has_as_pullback where "has_as_pullback p0 p1 \ limit_cone (C.dom p0) (mkCone p0 p1)" lemma cone_mkCone: assumes "is_rendered_commutative_by p0 p1" shows "cone (C.dom p0) (mkCone p0 p1)" proof - interpret E: constant_functor J.comp C \C.dom p0\ apply unfold_locales using assms by auto show "cone (C.dom p0) (mkCone p0 p1)" proof fix f show "\ J.arr f \ mkCone p0 p1 f = C.null" using mkCone_def J.arr_char by simp assume f: "J.arr f" show "C.dom (mkCone p0 p1 f) = E.map (J.dom f)" using assms f mkCone_def J.arr_char J.dom_char apply (cases f, simp_all) by (metis C.dom_comp)+ show "C.cod (mkCone p0 p1 f) = map (J.cod f)" using assms f mkCone_def J.arr_char J.cod_char is_cospan by (cases f, auto) show "map f \ mkCone p0 p1 (J.dom f) = mkCone p0 p1 f" using assms f mkCone_def J.arr_char J.dom_char C.comp_ide_arr is_cospan by (cases f, auto) show "mkCone p0 p1 (J.cod f) \ E.map f = mkCone p0 p1 f" using assms f mkCone_def J.arr_char J.cod_char C.comp_arr_dom apply (cases f, auto) apply (metis C.dom_comp C.seqE) by (metis C.dom_comp)+ qed qed lemma is_rendered_commutative_by_cone: assumes "cone a \" shows "is_rendered_commutative_by (\ J.AA) (\ J.BB)" proof - interpret \: cone J.comp C map a \ using assms by auto show ?thesis proof show "C.seq f0 (\ J.AA)" by (metis C.seqI J.cod_char J.seq_char \.preserves_cod \.preserves_reflects_arr J.seqE is_cospan J.Cod.simps(1) map.simps(1)) show "f0 \ \ J.AA = f1 \ \ J.BB" by (metis J.cod_char J.dom_char \.A.map_simp \.naturality J.Cod.simps(4-5) J.Dom.simps(4-5) J.comp.simps(2,5) J.seq_char map.simps(4-5)) qed qed lemma mkCone_cone: assumes "cone a \" shows "mkCone (\ J.AA) (\ J.BB) = \" proof - interpret \: cone J.comp C map a \ using assms by auto have 1: "is_rendered_commutative_by (\ J.AA) (\ J.BB)" using assms is_rendered_commutative_by_cone by blast interpret mkCone_\: cone J.comp C map \C.dom (\ J.AA)\ \mkCone (\ J.AA) (\ J.BB)\ using assms cone_mkCone 1 by auto show ?thesis proof - have "\j. j = J.AA \ mkCone (\ J.AA) (\ J.BB) j = \ j" using mkCone_def \.is_extensional by simp moreover have "\j. j = J.BB \ mkCone (\ J.AA) (\ J.BB) j = \ j" using mkCone_def \.is_extensional by simp moreover have "\j. j = J.TT \ mkCone (\ J.AA) (\ J.BB) j = \ j" using 1 mkCone_def \.is_extensional \.A.map_simp \.preserves_comp_1 cospan_shape.seq_char \.is_natural_2 apply simp by (metis J.seqE J.comp.simps(5) map.simps(5)) ultimately have "\j. J.ide j \ mkCone (\ J.AA) (\ J.BB) j = \ j" using J.ide_char by auto thus "mkCone (\ J.AA) (\ J.BB) = \" using mkCone_def NaturalTransformation.eqI [of J.comp C] \.natural_transformation_axioms mkCone_\.natural_transformation_axioms J.ide_char by simp qed qed end locale pullback_cone = J: cospan_shape + C: category C + D: cospan_diagram C f0 f1 + limit_cone J.comp C D.map \C.dom p0\ \D.mkCone p0 p1\ for C :: "'c comp" (infixr "\" 55) and f0 :: 'c and f1 :: 'c and p0 :: 'c and p1 :: 'c begin (* TODO: Equalizer should be simplifiable in the same way. *) lemma renders_commutative: shows "D.is_rendered_commutative_by p0 p1" using D.mkCone_def D.cospan_diagram_axioms cone_axioms cospan_diagram.is_rendered_commutative_by_cone by fastforce lemma is_universal': assumes "D.is_rendered_commutative_by p0' p1'" shows "\!h. \h : C.dom p0' \ C.dom p0\ \ p0 \ h = p0' \ p1 \ h = p1'" proof - have "D.cone (C.dom p0') (D.mkCone p0' p1')" using assms D.cone_mkCone by blast hence 1: "\!h. \h : C.dom p0' \ C.dom p0\ \ D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1'" using is_universal by simp have 2: "\h. \h : C.dom p0' \ C.dom p0\ \ D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1' \ p0 \ h = p0' \ p1 \ h = p1'" proof - fix h assume h: "\h : C.dom p0' \ C.dom p0\" show "D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1' \ p0 \ h = p0' \ p1 \ h = p1'" proof assume 3: "D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1'" show "p0 \ h = p0' \ p1 \ h = p1'" proof show "p0 \ h = p0'" proof - have "p0' = D.mkCone p0' p1' J.AA" using D.mkCone_def J.arr_char by simp also have "... = D.cones_map h (D.mkCone p0 p1) J.AA" using 3 by simp also have "... = p0 \ h" using h D.mkCone_def J.arr_char cone_\ by auto finally show ?thesis by auto qed show "p1 \ h = p1'" proof - have "p1' = D.mkCone p0' p1' J.BB" using D.mkCone_def J.arr_char by simp also have "... = D.cones_map h (D.mkCone p0 p1) J.BB" using 3 by simp also have "... = p1 \ h" using h D.mkCone_def J.arr_char cone_\ by auto finally show ?thesis by auto qed qed next assume 4: "p0 \ h = p0' \ p1 \ h = p1'" show "D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1'" proof fix j have "\ J.arr j \ D.cones_map h (D.mkCone p0 p1) j = D.mkCone p0' p1' j" using h cone_axioms D.mkCone_def J.arr_char by auto moreover have "J.arr j \ D.cones_map h (D.mkCone p0 p1) j = D.mkCone p0' p1' j" using assms h 4 cone_\ D.mkCone_def J.arr_char renders_commutative C.comp_assoc by fastforce ultimately show "D.cones_map h (D.mkCone p0 p1) j = D.mkCone p0' p1' j" using J.arr_char J.Dom.cases by blast qed qed qed thus ?thesis using 1 by blast qed lemma induced_arrowI': assumes "D.is_rendered_commutative_by p0' p1'" shows "\induced_arrow (C.dom p0') (D.mkCone p0' p1') : C.dom p0' \ C.dom p0\" and "p0 \ induced_arrow (C.dom p0') (D.mkCone p0' p1') = p0'" and "p1 \ induced_arrow (C.dom p1') (D.mkCone p0' p1') = p1'" proof - interpret A': constant_functor J.comp C \C.dom p0'\ using assms by (unfold_locales, auto) have cone: "D.cone (C.dom p0') (D.mkCone p0' p1')" using assms D.cone_mkCone [of p0' p1'] by blast show 1: "p0 \ induced_arrow (C.dom p0') (D.mkCone p0' p1') = p0'" proof - have "p0 \ induced_arrow (C.dom p0') (D.mkCone p0' p1') = D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1')) (D.mkCone p0 p1) J.AA" using cone induced_arrowI(1) D.mkCone_def J.arr_char cone_\ by force also have "... = p0'" proof - have "D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1')) (D.mkCone p0 p1) = D.mkCone p0' p1'" using cone induced_arrowI by blast thus ?thesis using J.arr_char D.mkCone_def by simp qed finally show ?thesis by auto qed show 2: "p1 \ induced_arrow (C.dom p1') (D.mkCone p0' p1') = p1'" proof - have "p1 \ induced_arrow (C.dom p1') (D.mkCone p0' p1') = D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1')) (D.mkCone p0 p1) J.BB" proof - have "C.dom p0' = C.dom p1'" using assms by (metis C.dom_comp) thus ?thesis using cone induced_arrowI(1) D.mkCone_def J.arr_char cone_\ by force qed also have "... = p1'" proof - have "D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1')) (D.mkCone p0 p1) = D.mkCone p0' p1'" using cone induced_arrowI by blast thus ?thesis using J.arr_char D.mkCone_def by simp qed finally show ?thesis by auto qed show "\induced_arrow (C.dom p0') (D.mkCone p0' p1') : C.dom p0' \ C.dom p0\" using 1 cone induced_arrowI by simp qed end context category begin definition has_as_pullback where "has_as_pullback f0 f1 p0 p1 \ cospan f0 f1 \ cospan_diagram.has_as_pullback C f0 f1 p0 p1" definition has_pullbacks where "has_pullbacks = (\f0 f1. cospan f0 f1 \ (\p0 p1. has_as_pullback f0 f1 p0 p1))" end locale category_with_pullbacks = category + assumes has_pullbacks: has_pullbacks section "Elementary Category with Pullbacks" text \ An \emph{elementary category with pullbacks} is a category equipped with a specific way of mapping each cospan to a span such that the resulting square commutes and such that the span is universal for that property. It is useful to assume that the functions, mapping a cospan to the two projections of the pullback, are extensional; that is, they yield @{term null} when applied to arguments that do not form a cospan. \ locale elementary_category_with_pullbacks = category C for C :: "'a comp" (infixr "\" 55) and prj0 :: "'a \ 'a \ 'a" ("\

\<^sub>0[_, _]") and prj1 :: "'a \ 'a \ 'a" ("\

\<^sub>1[_, _]") + assumes prj0_ext: "\ cospan f g \ \

\<^sub>0[f, g] = null" and prj1_ext: "\ cospan f g \ \

\<^sub>1[f, g] = null" and pullback_commutes [intro]: "cospan f g \ commutative_square f g \

\<^sub>1[f, g] \

\<^sub>0[f, g]" and universal: "commutative_square f g h k \ \!l. \

\<^sub>1[f, g] \ l = h \ \

\<^sub>0[f, g] \ l = k" begin lemma pullback_commutes': assumes "cospan f g" shows "f \ \

\<^sub>1[f, g] = g \ \

\<^sub>0[f, g]" using assms commutative_square_def by blast lemma prj0_in_hom': assumes "cospan f g" shows "\\

\<^sub>0[f, g] : dom \

\<^sub>0[f, g] \ dom g\" using assms pullback_commutes by (metis category.commutative_squareE category_axioms in_homI) lemma prj1_in_hom': assumes "cospan f g" shows "\\

\<^sub>1[f, g] : dom \

\<^sub>0[f, g] \ dom f\" using assms pullback_commutes by (metis category.commutative_squareE category_axioms in_homI) text \ The following gives us a notation for the common domain of the two projections of a pullback. \ definition pbdom (infix "\\" 51) where "f \\ g \ dom \

\<^sub>0[f, g]" lemma pbdom_in_hom [intro]: assumes "cospan f g" shows "\f \\ g : f \\ g \ f \\ g\" unfolding pbdom_def using assms prj0_in_hom' by (metis arr_dom_iff_arr arr_iff_in_hom cod_dom dom_dom in_homE) lemma ide_pbdom [simp]: assumes "cospan f g" shows "ide (f \\ g)" using assms ide_in_hom by auto[1] lemma prj0_in_hom [intro, simp]: assumes "cospan f g" and "a = f \\ g" and "b = dom g" shows "\\

\<^sub>0[f, g] : a \ b\" unfolding pbdom_def using assms prj0_in_hom' by (simp add: pbdom_def) lemma prj1_in_hom [intro, simp]: assumes "cospan f g" and "a = f \\ g" and "b = dom f" shows "\\

\<^sub>1[f, g] : a \ b\" unfolding pbdom_def using assms prj1_in_hom' by (simp add: pbdom_def) lemma prj0_simps [simp]: assumes "cospan f g" shows "arr \

\<^sub>0[f, g]" and "dom \

\<^sub>0[f, g] = f \\ g" and "cod \

\<^sub>0[f, g] = dom g" using assms prj0_in_hom by (blast, blast, blast) lemma prj0_simps_arr [iff]: shows "arr \

\<^sub>0[f, g] \ cospan f g" proof show "cospan f g \ arr \

\<^sub>0[f, g]" using prj0_in_hom by auto show "arr \

\<^sub>0[f, g] \ cospan f g" using prj0_ext not_arr_null by metis qed lemma prj1_simps [simp]: assumes "cospan f g" shows "arr \

\<^sub>1[f, g]" and "dom \

\<^sub>1[f, g] = f \\ g" and "cod \

\<^sub>1[f, g] = dom f" using assms prj1_in_hom by (blast, blast, blast) lemma prj1_simps_arr [iff]: shows "arr \

\<^sub>1[f, g] \ cospan f g" proof show "cospan f g \ arr \

\<^sub>1[f, g]" using prj1_in_hom by auto show "arr \

\<^sub>1[f, g] \ cospan f g" using prj1_ext not_arr_null by metis qed lemma span_prj: assumes "cospan f g" shows "span \

\<^sub>0[f, g] \

\<^sub>1[f, g]" using assms by simp text \ We introduce a notation for tupling, which produces the induced arrow into a pullback. In our notation, the ``$0$-side'', which we regard as the input, occurs on the right, and the ``$1$-side'', which we regard as the output, occurs on the left. \ definition tuple ("\_ \_, _\ _\") where "\h \f, g\ k\ \ if commutative_square f g h k then THE l. \

\<^sub>0[f, g] \ l = k \ \

\<^sub>1[f, g] \ l = h else null" lemma tuple_in_hom [intro]: assumes "commutative_square f g h k" shows "\\h \f, g\ k\ : dom h \ f \\ g\" proof have 1: "\

\<^sub>0[f, g] \ \h \f, g\ k\ = k \ \

\<^sub>1[f, g] \ \h \f, g\ k\ = h" unfolding tuple_def using assms universal theI [of "\l. \

\<^sub>0[f, g] \ l = k \ \

\<^sub>1[f, g] \ l = h"] apply simp by meson show "arr \h \f, g\ k\" using assms 1 apply (elim commutative_squareE) by (metis (no_types, lifting) seqE) show "dom \h \f, g\ k\ = dom h" using assms 1 apply (elim commutative_squareE) by (metis (no_types, lifting) dom_comp) show "cod \h \f, g\ k\ = f \\ g" unfolding pbdom_def using assms 1 apply (elim commutative_squareE) by (metis seqE) qed lemma tuple_is_extensional: assumes "\ commutative_square f g h k" shows "\h \f, g\ k\ = null" unfolding tuple_def using assms by simp lemma tuple_simps [simp]: assumes "commutative_square f g h k" shows "arr \h \f, g\ k\" and "dom \h \f, g\ k\ = dom h" and "cod \h \f, g\ k\ = f \\ g" using assms tuple_in_hom apply blast using assms tuple_in_hom apply blast using assms tuple_in_hom by blast lemma prj_tuple [simp]: assumes "commutative_square f g h k" shows "\

\<^sub>0[f, g] \ \h \f, g\ k\ = k" and "\

\<^sub>1[f, g] \ \h \f, g\ k\ = h" proof - have 1: "\

\<^sub>0[f, g] \ \h \f, g\ k\ = k \ \

\<^sub>1[f, g] \ \h \f, g\ k\ = h" unfolding tuple_def using assms universal theI [of "\l. \

\<^sub>0[f, g] \ l = k \ \

\<^sub>1[f, g] \ l = h"] apply simp by meson show "\

\<^sub>0[f, g] \ \h \f, g\ k\ = k" using 1 by simp show "\

\<^sub>1[f, g] \ \h \f, g\ k\ = h" using 1 by simp qed lemma tuple_prj: assumes "cospan f g" and "seq \

\<^sub>1[f, g] h" shows "\\

\<^sub>1[f, g] \ h \f, g\ \

\<^sub>0[f, g] \ h\ = h" proof - have 1: "commutative_square f g (\

\<^sub>1[f, g] \ h) (\

\<^sub>0[f, g] \ h)" using assms pullback_commutes by (simp add: commutative_square_comp_arr) have "\

\<^sub>0[f, g] \ \\

\<^sub>1[f, g] \ h \f, g\ \

\<^sub>0[f, g] \ h\ = \

\<^sub>0[f, g] \ h" using assms 1 by simp moreover have "\

\<^sub>1[f, g] \ \\

\<^sub>1[f, g] \ h \f, g\ \

\<^sub>0[f, g] \ h\ = \

\<^sub>1[f, g] \ h" using assms 1 by simp ultimately show ?thesis unfolding tuple_def using assms 1 universal [of f g "\

\<^sub>1[f, g] \ h" "\

\<^sub>0[f, g] \ h"] theI_unique [of "\l. \

\<^sub>0[f, g] \ l = \

\<^sub>0[f, g] \ h \ \

\<^sub>1[f, g] \ l = \

\<^sub>1[f, g] \ h" h] by auto qed lemma tuple_prj_spc [simp]: assumes "cospan f g" shows "\\

\<^sub>1[f, g] \f, g\ \

\<^sub>0[f, g]\ = f \\ g" proof - have "\\

\<^sub>1[f, g] \f, g\ \

\<^sub>0[f, g]\ = \\

\<^sub>1[f, g] \ (f \\ g) \f, g\ \

\<^sub>0[f, g] \ (f \\ g)\" using assms comp_arr_dom by simp thus ?thesis using assms tuple_prj by simp qed lemma prj_joint_monic: assumes "cospan f g" and "seq \

\<^sub>1[f, g] h" and "seq \

\<^sub>1[f, g] h'" and "\

\<^sub>0[f, g] \ h = \

\<^sub>0[f, g] \ h'" and "\

\<^sub>1[f, g] \ h = \

\<^sub>1[f, g] \ h'" shows "h = h'" proof - have "h = \\

\<^sub>1[f, g] \ h \f, g\ \

\<^sub>0[f, g] \ h\" using assms tuple_prj [of f g h] by simp also have "... = \\

\<^sub>1[f, g] \ h' \f, g\ \

\<^sub>0[f, g] \ h'\" using assms by simp also have "... = h'" using assms tuple_prj [of f g h'] by simp finally show ?thesis by blast qed text \ The pullback of an identity along an arbitrary arrow is an isomorphism. \ lemma iso_pullback_ide: assumes "cospan \ \" and "ide \" shows "iso \

\<^sub>0[\, \]" proof - have "inverse_arrows \

\<^sub>0[\, \] \\ \\, \\ dom \\" proof show 1: "ide (\

\<^sub>0[\, \] \ \\ \\, \\ dom \\)" using assms comp_arr_dom comp_cod_arr prj_tuple(1) by simp show "ide (\\ \\, \\ dom \\ \ \

\<^sub>0[\, \])" proof - have "\\ \\, \\ dom \\ \ \

\<^sub>0[\, \] = (\ \\ \)" proof - have "\

\<^sub>0[\, \] \ \\ \\, \\ dom \\ \ \

\<^sub>0[\, \] = \

\<^sub>0[\, \] \ (\ \\ \)" proof - have "\

\<^sub>0[\, \] \ \\ \\, \\ dom \\ \ \

\<^sub>0[\, \] = (\

\<^sub>0[\, \] \ \\ \\, \\ dom \\) \ \

\<^sub>0[\, \]" using assms 1 comp_reduce by blast also have "... = \

\<^sub>0[\, \] \ (\ \\ \)" using assms prj_tuple(1) pullback_commutes comp_arr_dom comp_cod_arr by simp finally show ?thesis by blast qed moreover have "\

\<^sub>1[\, \] \ \\ \\, \\ dom \\ \ \

\<^sub>0[\, \] = \

\<^sub>1[\, \] \ (\ \\ \)" proof - have "\

\<^sub>1[\, \] \ \\ \\, \\ dom \\ \ \

\<^sub>0[\, \] = (\

\<^sub>1[\, \] \ \\ \\, \\ dom \\) \ \

\<^sub>0[\, \]" using assms(2) comp_assoc by simp also have "... = \ \ \

\<^sub>0[\, \]" using assms comp_arr_dom comp_cod_arr prj_tuple(2) by fastforce also have "... = \ \ \

\<^sub>1[\, \]" using assms pullback_commutes commutative_square_def by simp also have "... = \

\<^sub>1[\, \] \ (\ \\ \)" using assms comp_arr_dom comp_cod_arr pullback_commutes commutative_square_def by simp finally show ?thesis by simp qed ultimately show ?thesis using assms prj0_in_hom prj1_in_hom comp_arr_dom prj1_simps(1-2) prj_joint_monic by metis qed thus ?thesis using assms by auto qed qed thus ?thesis by auto qed lemma comp_tuple_arr: assumes "commutative_square f g h k" and "seq h l" shows "\h \f, g\ k\ \ l = \h \ l \f, g\ k \ l\" proof - have "\

\<^sub>0[f, g] \ \h \f, g\ k\ \ l = \

\<^sub>0[f, g] \ \h \ l \f, g\ k \ l\" proof - have "\

\<^sub>0[f, g] \ \h \f, g\ k\ \ l = (\

\<^sub>0[f, g] \ \h \f, g\ k\) \ l" using comp_assoc by simp also have "... = \

\<^sub>0[f, g] \ \h \ l \f, g\ k \ l\" using assms commutative_square_comp_arr by auto finally show ?thesis by blast qed moreover have "\

\<^sub>1[f, g] \ \h \f, g\ k\ \ l = \

\<^sub>1[f, g] \ \h \ l \f, g\ k \ l\" proof - have "\

\<^sub>1[f, g] \ \h \f, g\ k\ \ l = (\

\<^sub>1[f, g] \ \h \f, g\ k\) \ l" using comp_assoc by simp also have "... = \

\<^sub>1[f, g] \ \h \ l \f, g\ k \ l\" using assms commutative_square_comp_arr by auto finally show ?thesis by blast qed moreover have "seq \

\<^sub>1[f, g] (\h \f, g\ k\ \ l)" using assms tuple_in_hom prj1_in_hom by fastforce ultimately show ?thesis using assms prj_joint_monic [of f g "\h \f, g\ k\ \ l" "\h \ l \f, g\ k \ l\"] by auto qed lemma pullback_arr_cod: assumes "arr f" shows "inverse_arrows \

\<^sub>1[f, cod f] \dom f \f, cod f\ f\" and "inverse_arrows \

\<^sub>0[cod f, f] \f \cod f, f\ dom f\" proof - show "inverse_arrows \

\<^sub>1[f, cod f] \dom f \f, cod f\ f\" proof show "ide (\dom f \f, cod f\ f\ \ \

\<^sub>1[f, cod f])" proof - have "\dom f \f, cod f\ f\ \ \

\<^sub>1[f, cod f] = f \\ cod f" proof - have "\

\<^sub>0[f, cod f] \ \dom f \f, cod f\ f\ \ \

\<^sub>1[f, cod f] = \

\<^sub>0[f, cod f] \ (f \\ cod f)" proof - have "\

\<^sub>0[f, cod f] \ \dom f \f, cod f\ f\ \ \

\<^sub>1[f, cod f] = (\

\<^sub>0[f, cod f] \ \dom f \f, cod f\ f\) \ \

\<^sub>1[f, cod f]" using comp_assoc by simp also have "... = \

\<^sub>0[f, cod f] \ (f \\ cod f)" using assms pullback_commutes [of f "cod f"] comp_arr_dom comp_cod_arr by auto finally show ?thesis by blast qed moreover have "\

\<^sub>1[f, cod f] \ \dom f \f, cod f\ f\ \ \

\<^sub>1[f, cod f] = \

\<^sub>1[f, cod f] \ (f \\ cod f)" proof - have "\

\<^sub>1[f, cod f] \ \dom f \f, cod f\ f\ \ \

\<^sub>1[f, cod f] = (\

\<^sub>1[f, cod f] \ \dom f \f, cod f\ f\) \ \

\<^sub>1[f, cod f]" using assms comp_assoc by presburger also have "... = \

\<^sub>1[f, cod f] \ (f \\ cod f)" using assms comp_arr_dom comp_cod_arr by simp finally show ?thesis by blast qed ultimately show ?thesis using assms prj_joint_monic [of f "cod f" "\dom f \f, cod f\ f\ \ \

\<^sub>1[f, cod f]" "f \\ cod f"] by simp qed thus ?thesis using assms arr_cod cod_cod prj1_simps_arr by simp qed show "ide (\

\<^sub>1[f, cod f] \ \dom f \f, cod f\ f\)" using assms comp_arr_dom comp_cod_arr by fastforce qed show "inverse_arrows \

\<^sub>0[cod f, f] \f \cod f, f\ dom f\" proof show "ide (\

\<^sub>0[cod f, f] \ \f \cod f, f\ dom f\)" using assms comp_arr_dom comp_cod_arr by simp show "ide (\f \cod f, f\ dom f\ \ \

\<^sub>0[cod f, f])" proof - have "\f \cod f, f\ dom f\ \ \

\<^sub>0[cod f, f] = cod f \\ f" proof - have "\

\<^sub>0[cod f, f] \ \f \cod f, f\ dom f\ \ \

\<^sub>0[cod f, f] = \

\<^sub>0[cod f, f] \ (cod f \\ f)" proof - have "\

\<^sub>0[cod f, f] \ \f \cod f, f\ dom f\ \ \

\<^sub>0[cod f, f] = (\

\<^sub>0[cod f, f] \ \f \cod f, f\ dom f\) \ \

\<^sub>0[cod f, f]" using comp_assoc by simp also have "... = dom f \ \

\<^sub>0[cod f, f]" using assms comp_arr_dom comp_cod_arr by simp also have "... = \

\<^sub>0[cod f, f] \ (cod f \\ f)" using assms comp_arr_dom comp_cod_arr by simp finally show ?thesis using prj0_in_hom by blast qed moreover have "\

\<^sub>1[cod f, f] \ \f \cod f, f\ dom f\ \ \

\<^sub>0[cod f, f] = \

\<^sub>1[cod f, f] \ (cod f \\ f)" proof - have "\

\<^sub>1[cod f, f] \ \f \cod f, f\ dom f\ \ \

\<^sub>0[cod f, f] = (\

\<^sub>1[cod f, f] \ \f \cod f, f\ dom f\) \ \

\<^sub>0[cod f, f]" using comp_assoc by simp also have "... = \

\<^sub>1[cod f, f] \ (cod f \\ f)" using assms pullback_commutes [of "cod f" f] comp_arr_dom comp_cod_arr by auto finally show ?thesis by blast qed ultimately show ?thesis using assms prj_joint_monic [of "cod f" f "\f \cod f, f\ dom f\ \ \

\<^sub>0[cod f, f]"] by simp qed thus ?thesis using assms by simp qed qed qed text \ The pullback of a monomorphism along itself is automatically symmetric: the left and right projections are equal. \ lemma pullback_mono_self: assumes "mono f" shows "\

\<^sub>0[f, f] = \

\<^sub>1[f, f]" proof - have "f \ \

\<^sub>0[f, f] = f \ \

\<^sub>1[f, f]" using assms pullback_commutes [of f f] by (metis commutative_squareE mono_implies_arr) thus ?thesis using assms monoE [of f "\

\<^sub>1[f, f]" "\

\<^sub>0[f, f]"] by (metis mono_implies_arr prj0_simps(1,3) seqI) qed lemma pullback_iso_self: assumes "iso f" shows "\

\<^sub>0[f, f] = \

\<^sub>1[f, f]" using assms pullback_mono_self iso_is_section section_is_mono by simp lemma pullback_ide_self [simp]: assumes "ide a" shows "\

\<^sub>0[a, a] = \

\<^sub>1[a, a]" using assms pullback_iso_self ide_is_iso by blast end section "Agreement between the Definitions" text \ It is very easy to write locale assumptions that have unintended consequences or that are even inconsistent. So, to keep ourselves honest, we don't just accept the definition of ``elementary category with pullbacks'', but in fact we formally establish the sense in which it agrees with our standard definition of ``category with pullbacks'', which is given in terms of limit cones. This is extra work, but it ensures that we didn't make a mistake. \ context category_with_pullbacks begin definition prj1 where "prj1 f g \ if cospan f g then fst (SOME x. cospan_diagram.has_as_pullback C f g (fst x) (snd x)) else null" definition prj0 where "prj0 f g \ if cospan f g then snd (SOME x. cospan_diagram.has_as_pullback C f g (fst x) (snd x)) else null" lemma prj_yields_pullback: assumes "cospan f g" shows "cospan_diagram.has_as_pullback C f g (prj1 f g) (prj0 f g)" proof - have "\x. cospan_diagram.has_as_pullback C f g (fst x) (snd x)" using assms has_pullbacks has_pullbacks_def has_as_pullback_def by simp thus ?thesis using assms has_pullbacks has_pullbacks_def prj0_def prj1_def someI_ex [of "\x. cospan_diagram.has_as_pullback C f g (fst x) (snd x)"] by simp qed interpretation elementary_category_with_pullbacks C prj0 prj1 proof show "\f g. \ cospan f g \ prj0 f g = null" using prj0_def by auto show "\f g. \ cospan f g \ prj1 f g = null" using prj1_def by auto show "\f g. cospan f g \ commutative_square f g (prj1 f g) (prj0 f g)" proof fix f g assume fg: "cospan f g" show "cospan f g" by fact interpret J: cospan_shape . interpret D: cospan_diagram C f g using fg by (unfold_locales, auto) let ?\ = "D.mkCone (prj1 f g) (prj0 f g)" interpret \: limit_cone J.comp C D.map \dom (prj1 f g)\ ?\ using fg prj_yields_pullback by auto have 1: "prj1 f g = ?\ J.AA \ prj0 f g = ?\ J.BB" using D.mkCone_def by simp show "span (prj1 f g) (prj0 f g)" proof - have "arr (prj1 f g) \ arr (prj0 f g)" using 1 J.arr_char J.seq_char by (metis J.seqE \.preserves_reflects_arr) moreover have "dom (prj1 f g) = dom (prj0 f g)" using 1 D.is_rendered_commutative_by_cone \.cone_axioms J.seq_char by (metis J.cod_eqI J.seqE \.A.map_simp \.preserves_dom J.ide_char) ultimately show ?thesis by simp qed show "dom f = cod (prj1 f g)" using 1 fg \.preserves_cod [of J.BB] J.cod_char D.mkCone_def by (metis D.map.simps(1) D.preserves_cod J.seqE \.preserves_cod cod_dom J.seq_char) show "f \ prj1 f g = g \ prj0 f g" using 1 fg D.is_rendered_commutative_by_cone \.cone_axioms by force qed show "\f g h k. commutative_square f g h k \ \!l. prj1 f g \ l = h \ prj0 f g \ l = k" proof - fix f g h k assume fghk: "commutative_square f g h k" interpret J: cospan_shape . interpret D: cospan_diagram C f g using fghk by (unfold_locales, auto) let ?\ = "D.mkCone (prj1 f g) (prj0 f g)" interpret \: limit_cone J.comp C D.map \dom (prj1 f g)\ ?\ using fghk prj_yields_pullback by auto interpret \: pullback_cone C f g \prj1 f g\ \prj0 f g\ .. have 1: "prj1 f g = ?\ J.AA \ prj0 f g = ?\ J.BB" using D.mkCone_def by simp show "\!l. prj1 f g \ l = h \ prj0 f g \ l = k" proof let ?l = "SOME l. prj1 f g \ l = h \ prj0 f g \ l = k" show "prj1 f g \ ?l = h \ prj0 f g \ ?l = k" using fghk \.is_universal' \.renders_commutative someI_ex [of "\l. prj1 f g \ l = h \ prj0 f g \ l = k"] by blast thus "\l. prj1 f g \ l = h \ prj0 f g \ l = k \ l = ?l" using fghk \.is_universal' \.renders_commutative limit_cone_def by (metis (no_types, lifting) in_homI seqE commutative_squareE dom_comp seqI) qed qed qed proposition extends_to_elementary_category_with_pullbacks: shows "elementary_category_with_pullbacks C prj0 prj1" .. end context elementary_category_with_pullbacks begin interpretation category_with_pullbacks C proof show "has_pullbacks" proof (unfold has_pullbacks_def) have "\f g. cospan f g \ \p0 p1. has_as_pullback f g p0 p1" proof - fix f g assume fg: "cospan f g" interpret J: cospan_shape . interpret D: cospan_diagram C f g using fg by (unfold_locales, auto) let ?\ = "D.mkCone \

\<^sub>1[f, g] \

\<^sub>0[f, g]" interpret \: cone J.comp C D.map \dom \

\<^sub>1[f, g]\ ?\ proof - have "D.is_rendered_commutative_by \

\<^sub>1[f, g] \

\<^sub>0[f, g]" using fg pullback_commutes' by simp thus "cone J.comp C D.map (dom \

\<^sub>1[f, g]) ?\" using D.cone_mkCone by auto qed interpret \: limit_cone J.comp C D.map \dom \

\<^sub>1[f, g]\ ?\ proof fix a' \' assume \': "D.cone a' \'" interpret \': cone J.comp C D.map a' \' using \' by simp have 1: "commutative_square f g (\' J.AA) (\' J.BB)" using fg J.ide_char J.cod_char D.is_rendered_commutative_by_cone \'.cone_axioms by auto show "\!h. \h : a' \ dom \

\<^sub>1[f, g]\ \ D.cones_map h (D.mkCone \

\<^sub>1[f, g] \

\<^sub>0[f, g]) = \'" proof let ?h = "\\' J.AA \f, g\ \' J.BB\" show h': "\?h : a' \ dom \

\<^sub>1[f, g]\ \ D.cones_map ?h (D.mkCone \

\<^sub>1[f, g] \

\<^sub>0[f, g]) = \'" proof show h: "\?h : a' \ dom \

\<^sub>1[f, g]\" using fg 1 by fastforce show "D.cones_map ?h (D.mkCone \

\<^sub>1[f, g] \

\<^sub>0[f, g]) = \'" proof - have "D.mkCone \

\<^sub>1[f, g] \

\<^sub>0[f, g] \ D.cones (cod \\' J.AA \f, g\ \' J.BB\)" using fg h D.cone_mkCone D.is_rendered_commutative_by_cone \.cone_axioms by auto hence 2: "D.cones_map ?h (D.mkCone \

\<^sub>1[f, g] \

\<^sub>0[f, g]) \ D.cones a'" using fg h D.cones_map_mapsto by blast interpret \'h: cone J.comp C D.map a' \D.cones_map ?h (D.mkCone \

\<^sub>1[f, g] \

\<^sub>0[f, g])\ using 2 by simp show ?thesis proof - have "\j. J.ide j \ D.cones_map ?h (D.mkCone \

\<^sub>1[f, g] \

\<^sub>0[f, g]) j = \' j" proof - fix j show "J.ide j \ D.cones_map ?h (D.mkCone \

\<^sub>1[f, g] \

\<^sub>0[f, g]) j = \' j" using fg h 1 J.ide_char D.mkCone_def \.cone_axioms apply (cases j, simp_all) by (metis D.map.simps(4) J.dom_eqI \'.is_natural_1 \'.naturality J.seqE \'.A.map_simp J.comp.simps(3,7) J.seq_char prj_tuple(2) comp_assoc) qed thus ?thesis using NaturalTransformation.eqI \'.natural_transformation_axioms \'h.natural_transformation_axioms by blast qed qed qed show "\h. \h : a' \ dom \

\<^sub>1[f, g]\ \ D.cones_map h (D.mkCone \

\<^sub>1[f, g] \

\<^sub>0[f, g]) = \' \ h = ?h" proof - fix h assume 2: "\h : a' \ dom \

\<^sub>1[f, g]\ \ D.cones_map h (D.mkCone \

\<^sub>1[f, g] \

\<^sub>0[f, g]) = \'" show "h = ?h" proof - have "\

\<^sub>0[f, g] \ h = \

\<^sub>0[f, g] \ ?h \ \

\<^sub>1[f, g] \ h = \

\<^sub>1[f, g] \ ?h" using 1 2 fg J.arr_char \.cone_axioms D.mkCone_def by auto thus ?thesis using fg 2 h' prj_joint_monic by blast qed qed qed qed show "\p0 p1. has_as_pullback f g p0 p1" using fg has_as_pullback_def \.limit_cone_axioms by blast qed thus "\f g. cospan f g \ (\p0 p1. has_as_pullback f g p0 p1)" by simp qed qed proposition is_category_with_pullbacks: shows "category_with_pullbacks C" .. end sublocale elementary_category_with_pullbacks \ category_with_pullbacks using is_category_with_pullbacks by auto end diff --git a/thys/MonoidalCategory/CartesianMonoidalCategory.thy b/thys/MonoidalCategory/CartesianMonoidalCategory.thy --- a/thys/MonoidalCategory/CartesianMonoidalCategory.thy +++ b/thys/MonoidalCategory/CartesianMonoidalCategory.thy @@ -1,654 +1,654 @@ (* Title: CartesianMonoidalCategory Author: Eugene W. Stark , 2020 Maintainer: Eugene W. Stark *) chapter "Cartesian Monoidal Category" theory CartesianMonoidalCategory -imports MonoidalCategory.MonoidalCategory Category3.CartesianCategory +imports MonoidalCategory Category3.CartesianCategory begin locale symmetric_monoidal_category = monoidal_category C T \ \ + S: symmetry_functor C C + ToS: composite_functor CC.comp CC.comp C S.map T + \: natural_isomorphism CC.comp C T ToS.map \ for C :: "'a comp" (infixr "\" 55) and T :: "'a * 'a \ 'a" and \ :: "'a * 'a * 'a \ 'a" and \ :: 'a and \ :: "'a * 'a \ 'a" + assumes sym_inverse: "\ ide a; ide b \ \ inverse_arrows (\ (a, b)) (\ (b, a))" and unitor_coherence: "ide a \ \[a] \ \ (a, \) = \[a]" and assoc_coherence: "\ ide a; ide b; ide c \ \ \ (b, c, a) \ \ (a, b \ c) \ \ (a, b, c) = (b \ \ (a, c)) \ \ (b, a, c) \ (\ (a, b) \ c)" begin abbreviation sym ("\[_, _]") where "sym a b \ \ (a, b)" end locale elementary_symmetric_monoidal_category = elementary_monoidal_category C tensor unity lunit runit assoc for C :: "'a comp" (infixr "\" 55) and tensor :: "'a \ 'a \ 'a" (infixr "\" 53) and unity :: 'a ("\") and lunit :: "'a \ 'a" ("\[_]") and runit :: "'a \ 'a" ("\[_]") and assoc :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and sym :: "'a \ 'a \ 'a" ("\[_, _]") + assumes sym_in_hom: "\ ide a; ide b \ \ \\[a, b] : a \ b \ b \ a\" and sym_naturality: "\ arr f; arr g \ \ \[cod f, cod g] \ (f \ g) = (g \ f) \ \[dom f, dom g]" and sym_inverse: "\ ide a; ide b \ \ inverse_arrows \[a, b] \[b, a]" and unitor_coherence: "ide a \ \[a] \ \[a, \] = \[a]" and assoc_coherence: "\ ide a; ide b; ide c \ \ \[b, c, a] \ \[a, b \ c] \ \[a, b, c] = (b \ \[a, c]) \ \[b, a, c] \ (\[a, b] \ c)" begin lemma sym_simps [simp]: assumes "ide a" and "ide b" shows "arr \[a, b]" and "dom \[a, b] = a \ b" and "cod \[a, b] = b \ a" using assms sym_in_hom by auto interpretation monoidal_category C T \ \ using induces_monoidal_category by simp interpretation CC: product_category C C .. interpretation S: symmetry_functor C C .. interpretation ToS: composite_functor CC.comp CC.comp C S.map T .. definition \ :: "'a * 'a \ 'a" where "\ f \ if CC.arr f then \[cod (fst f), cod (snd f)] \ (fst f \ snd f) else null" interpretation \: natural_isomorphism CC.comp C T ToS.map \ proof - interpret \: transformation_by_components CC.comp C T ToS.map "\a. \[fst a, snd a]" apply unfold_locales using sym_in_hom sym_naturality by auto interpret \: natural_isomorphism CC.comp C T ToS.map \.map apply unfold_locales using sym_inverse \.map_simp_ide by auto have "\ = \.map" using \_def \.map_def sym_naturality by fastforce thus "natural_isomorphism CC.comp C T ToS.map \" using \.natural_isomorphism_axioms by presburger qed interpretation symmetric_monoidal_category C T \ \ \ proof show "\a b. \ ide a; ide b \ \ inverse_arrows (\ (a, b)) (\ (b, a))" proof - fix a b assume a: "ide a" and b: "ide b" show "inverse_arrows (\ (a, b)) (\ (b, a))" using a b sym_inverse comp_arr_dom \_def by auto qed (* * TODO: Here just using "lunit" refers to the locale parameter, rather than * the constant introduced by the interpretation above of monoidal_category. * This is slightly mysterious. *) show "\a. ide a \ local.lunit a \ \ (a, local.unity) = local.runit a" proof - fix a assume a: "ide a" show "local.lunit a \ \ (a, local.unity) = local.runit a" using a lunit_agreement \_agreement sym_in_hom comp_arr_dom [of "\[a, \]"] unitor_coherence runit_agreement \_def by simp qed show "\a b c. \ ide a; ide b; ide c \ \ local.assoc b c a \ \ (a, local.tensor b c) \ local.assoc a b c = local.tensor b (\ (a, c)) \ local.assoc b a c \ local.tensor (\ (a, b)) c" proof - fix a b c assume a: "ide a" and b: "ide b" and c: "ide c" show "local.assoc b c a \ \ (a, local.tensor b c) \ local.assoc a b c = local.tensor b (\ (a, c)) \ local.assoc b a c \ local.tensor (\ (a, b)) c" using a b c sym_in_hom tensor_preserves_ide \_def assoc_coherence comp_arr_dom comp_cod_arr by simp qed qed lemma induces_symmetric_monoidal_category: shows "symmetric_monoidal_category C T \ \ \" .. end context symmetric_monoidal_category begin interpretation elementary_monoidal_category C tensor unity lunit runit assoc using induces_elementary_monoidal_category by auto lemma induces_elementary_symmetric_monoidal_category: shows "elementary_symmetric_monoidal_category C tensor unity lunit runit assoc (\a b. \ (a, b))" using \.naturality unitor_coherence assoc_coherence sym_inverse by unfold_locales auto end (* TODO: This definition of "diagonal_functor" conflicts with the one in Category3.Limit. *) locale diagonal_functor = C: category C + CC: product_category C C for C :: "'a comp" begin abbreviation map where "map f \ if C.arr f then (f, f) else CC.null" lemma is_functor: shows "functor C CC.comp map" using map_def by unfold_locales auto sublocale "functor" C CC.comp map using is_functor by simp end locale cartesian_monoidal_category = monoidal_category C T \ \ + \: constant_functor C C \ + \: diagonal_functor C + \: natural_transformation C C map \.map \ + \: natural_transformation C C map \T o \.map\ \ for C :: "'a comp" (infixr "\" 55) and T :: "'a * 'a \ 'a" and \ :: "'a * 'a * 'a \ 'a" and \ :: 'a and \ :: "'a \ 'a" ("\[_]") and \ :: "'a \ 'a" ("\[_]") + assumes trm_unity: "\[\] = \" and pr0_dup: "ide a \ \[a] \ (a \ \[a]) \ \ a = a" and pr1_dup: "ide a \ \[a] \ (\[a] \ a) \ \ a = a" and tuple_pr: "\ ide a; ide b \ \ (\[a] \ (a \ \[b]) \ \[b] \ (\[a] \ b)) \ \[a \ b] = a \ b" locale elementary_cartesian_monoidal_category = elementary_monoidal_category C tensor unity lunit runit assoc for C :: "'a comp" (infixr "\" 55) and tensor :: "'a \ 'a \ 'a" (infixr "\" 53) and unity :: 'a ("\") and lunit :: "'a \ 'a" ("\[_]") and runit :: "'a \ 'a" ("\[_]") and assoc :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and trm :: "'a \ 'a" ("\[_]") and dup :: "'a \ 'a" ("\[_]") + assumes trm_in_hom [intro]: "ide a \ \\[a] : a \ \\" and trm_unity: "\[\] = \" and trm_naturality: "arr f \ \[cod f] \ f = \[dom f]" and dup_in_hom [intro]: "ide a \ \\[a] : a \ a \ a\" and dup_naturality: "arr f \ \[cod f] \ f = (f \ f) \ \[dom f]" and prj0_dup: "ide a \ \[a] \ (a \ \[a]) \ \[a] = a" and prj1_dup: "ide a \ \[a] \ (\[a] \ a) \ \[a] = a" and tuple_prj: "\ ide a; ide b \ \ (\[a] \ (a \ \[b]) \ \[b] \ (\[a] \ b)) \ \[a \ b] = a \ b" context cartesian_monoidal_category begin lemma terminal_unity: shows "terminal \" proof show "ide \" by simp show "\a. ide a \ \!f. \f : a \ \\" proof fix a assume a: "ide a" show "\\ a : a \ \\" using a by auto show "\f. \f : a \ \\ \ f = \ a" using trm_unity \.naturality comp_cod_arr by fastforce qed qed lemma trm_is_terminal_arr: assumes "ide a" shows "terminal_arr \[a]" using assms terminal_unity by simp interpretation elementary_monoidal_category C tensor unity lunit runit assoc using induces_elementary_monoidal_category by simp interpretation elementary_cartesian_monoidal_category C tensor unity lunit runit assoc \ \ proof show "\a. ide a \ \\[a] : a \ \\" using \_in_hom by force show "\[\] = \" using \.preserves_hom \_in_hom ide_unity trm_is_terminal_arr terminal_unity by (intro terminal_arr_unique) auto show "\f. arr f \ \[cod f] \ f = \[dom f]" using \.naturality comp_cod_arr by simp show "\a. ide a \ \\[a] : a \ a \ a\" by auto show "\f. arr f \ \[cod f] \ f = (f \ f) \ \[dom f]" using \.naturality by simp show "\a. ide a \ \[a] \ (\[a] \ a) \ \[a] = a" using pr1_dup lunit_in_hom by simp show "\a. ide a \ \[a] \ (a \ \[a]) \ \[a] = a" using pr0_dup runit_in_hom by simp show "\a0 a1. \ ide a0; ide a1 \ \ (\[a0] \ (a0 \ \[a1]) \ \[a1] \ (\[a0] \ a1)) \ \[a0 \ a1] = a0 \ a1" using tuple_pr by simp qed lemma induces_elementary_cartesian_monoidal_category: shows "elementary_cartesian_monoidal_category C tensor \ lunit runit assoc \ \" .. end context elementary_cartesian_monoidal_category begin lemma trm_simps [simp]: assumes "ide a" shows "arr \[a]" and "dom \[a] = a" and "cod \[a] = \" using assms trm_in_hom by auto lemma dup_simps [simp]: assumes "ide a" shows "arr \[a]" and "dom \[a] = a" and "cod \[a] = a \ a" using assms dup_in_hom by auto definition \ :: "'a \ 'a" where "\ f \ if arr f then \[dom f] else null" definition \ :: "'a \ 'a" where "\ f \ if arr f then \[cod f] \ f else null" interpretation CC: product_category C C .. interpretation MC: monoidal_category C T \ \ using induces_monoidal_category by auto interpretation I: constant_functor C C MC.unity by unfold_locales auto interpretation \: diagonal_functor C .. interpretation D: composite_functor C CC.comp C \.map T .. interpretation \: natural_transformation C C map I.map \ using trm_naturality I.map_def \_def \_agreement comp_cod_arr by unfold_locales auto interpretation \: natural_transformation C C map D.map \ using dup_naturality \_def comp_arr_dom by unfold_locales auto interpretation MC: cartesian_monoidal_category C T \ \ \ \ proof show "\ MC.unity = MC.unity" using \_agreement trm_unity \_def by simp show "\a. ide a \ MC.runit a \ MC.tensor a (\ a) \ \ a = a" using runit_agreement \_def \_def prj0_dup comp_arr_dom by auto show "\a. ide a \ MC.lunit a \ MC.tensor (\ a) a \ \ a = a" using lunit_agreement \_def \_def prj1_dup comp_arr_dom by auto show "\a b. \ ide a; ide b \ \ MC.tensor (MC.runit a \ MC.tensor a (\ b)) (MC.lunit b \ MC.tensor (\ a) b) \ \ (MC.tensor a b) = MC.tensor a b" proof - fix a b assume a: "ide a" and b: "ide b" have "seq \[a] (a \ \[b])" by (metis a b arr_tensor cod_tensor ide_char in_homE runit_in_hom seqI trm_simps(1,3)) moreover have "seq \[b] (\[a] \ b)" by (metis a b arr_tensor cod_tensor ide_char in_homE lunit_in_hom seqI trm_simps(1,3)) ultimately show "MC.tensor (MC.runit a \ MC.tensor a (\ b)) (MC.lunit b \ MC.tensor (\ a) b) \ \ (MC.tensor a b) = MC.tensor a b" using a b lunit_agreement runit_agreement unitor_coincidence \_def \_def comp_arr_dom tensor_preserves_ide tuple_prj T_def by auto qed qed lemma induces_cartesian_monoidal_category: shows "cartesian_monoidal_category C T \ \ \ \" .. end text \ A cartesian category extends to a a cartesian monoidal category by using the product structure to obtain the various canonical maps. \ context cartesian_category begin interpretation C: elementary_cartesian_category C pr0 pr1 \ trm using extends_to_elementary_cartesian_category by simp interpretation CC: product_category C C .. interpretation CCC: product_category C CC.comp .. interpretation T: binary_functor C C C Prod using binary_functor_Prod by simp interpretation T: binary_endofunctor C Prod .. interpretation ToTC: "functor" CCC.comp C T.ToTC using T.functor_ToTC by auto interpretation ToCT: "functor" CCC.comp C T.ToCT using T.functor_ToCT by auto interpretation \: transformation_by_components CCC.comp C T.ToTC T.ToCT \\abc. assoc (fst abc) (fst (snd abc)) (snd (snd abc))\ proof show "\abc. CCC.ide abc \ \assoc (fst abc) (fst (snd abc)) (snd (snd abc)) : T.ToTC abc \ T.ToCT abc\" using CCC.ide_char CC.ide_char CCC.arr_char CC.arr_char T.ToTC_def T.ToCT_def by auto show "\f. CCC.arr f \ assoc (fst (CCC.cod f)) (fst (snd (CCC.cod f))) (snd (snd (CCC.cod f))) \ T.ToTC f = T.ToCT f \ assoc (fst (CCC.dom f)) (fst (snd (CCC.dom f))) (snd (snd (CCC.dom f)))" using CCC.arr_char CC.arr_char CCC.dom_char CCC.cod_char T.ToTC_def T.ToCT_def assoc_naturality by simp blast qed interpretation \: natural_isomorphism CCC.comp C T.ToTC T.ToCT \.map proof show "\a. CCC.ide a \ iso (\.map a)" using CCC.ide_char CC.ide_char \.map_simp_ide inverse_arrows_assoc by auto qed interpretation L: "functor" C C \\f. Prod (cod \, f)\ using \_is_terminal_arr T.fixing_ide_gives_functor_1 by simp interpretation L: endofunctor C \\f. Prod (cod \, f)\ .. interpretation \: transformation_by_components C C \\f. Prod (cod \, f)\ map \\a. pr0 (cod \) a\ using \_is_terminal_arr by unfold_locales auto interpretation \: natural_isomorphism C C \\f. Prod (cod \, f)\ map \.map using \.map_simp_ide inverse_arrows_lunit ide_some_terminal by unfold_locales auto interpretation L: equivalence_functor C C \\f. Prod (cod \, f)\ using \.natural_isomorphism_axioms naturally_isomorphic_def L.isomorphic_to_identity_is_equivalence by blast interpretation R: "functor" C C \\f. Prod (f, cod \)\ using \_is_terminal_arr T.fixing_ide_gives_functor_2 by simp interpretation R: endofunctor C\\f. Prod (f, cod \)\ .. interpretation \: transformation_by_components C C \\f. Prod (f, cod \)\ map \\a. pr1 a (cod \)\ using \_is_terminal_arr by unfold_locales auto interpretation \: natural_isomorphism C C \\f. Prod (f, cod \)\ map \.map using \.map_simp_ide inverse_arrows_runit ide_some_terminal by unfold_locales auto interpretation R: equivalence_functor C C \\f. Prod (f, cod \)\ using \.natural_isomorphism_axioms naturally_isomorphic_def R.isomorphic_to_identity_is_equivalence by blast interpretation C: monoidal_category C Prod \.map \ using ide_some_terminal \_is_iso pentagon comp_assoc by unfold_locales auto interpretation \: constant_functor C C C.unity using C.ide_unity by unfold_locales auto interpretation \: natural_transformation C C map \.map trm using C.unity_def \.map_def ide_some_terminal trm_naturality comp_cod_arr trm_in_hom apply unfold_locales using trm_def apply auto[1] apply fastforce apply fastforce apply (metis in_homE trm_eqI trm_in_hom cod_pr0 dom_dom) by (metis trm_eqI trm_in_hom dom_dom map_simp) interpretation \: "functor" C CC.comp Diag using functor_Diag by simp interpretation \o\: composite_functor C CC.comp C Diag Prod .. interpretation natural_transformation C C map \Prod o Diag\ dup using dup_is_natural_transformation by simp lemma unity_agreement: shows "C.unity = \" using C.unity_def ide_some_terminal by simp lemma assoc_agreement: assumes "ide a" and "ide b" and "ide c" shows "C.assoc a b c = assoc a b c" using assms assoc_def \.map_simp_ide by simp lemma assoc'_agreement: assumes "ide a" and "ide b" and "ide c" shows "C.assoc' a b c = assoc' a b c" using assms inverse_arrows_assoc inverse_unique by auto lemma runit_char_eqn: assumes "ide a" shows "prod (runit a) \ = prod a \ \ assoc a \ \" using assms ide_one assoc_def comp_assoc prod_tuple comp_cod_arr by (intro pr_joint_monic [of a \ "prod (runit a) \" "prod a \ \ assoc a \ \"]) auto lemma runit_agreement: assumes "ide a" shows "runit a = C.runit a" using assms unity_agreement assoc_agreement C.runit_char(2) runit_char_eqn ide_some_terminal by (intro C.runit_eqI) auto lemma lunit_char_eqn: assumes "ide a" shows "prod \ (lunit a) = prod \ a \ assoc' \ \ a" proof (intro pr_joint_monic [of \ a "prod \ (lunit a)" "prod \ a \ assoc' \ \ a"]) show "seq (lunit a) (local.prod \ (lunit a))" using assms ide_one by simp show "lunit a \ prod \ (lunit a) = lunit a \ prod \ a \ assoc' \ \ a" using assms ide_one assoc'_def comp_assoc prod_tuple comp_cod_arr by simp show "pr1 \ a \ prod \ (lunit a) = pr1 \ a \ prod \ a \ assoc' \ \ a" using assms ide_one assoc'_def comp_cod_arr prod_tuple pr_naturality apply simp by (metis trm_eqI cod_pr0 cod_pr1 comp_in_homI' ide_prod pr_simps(1,3-6) pr1_in_hom') qed lemma lunit_agreement: assumes "ide a" shows "lunit a = C.lunit a" using assms unity_agreement assoc'_agreement C.lunit_char(2) lunit_char_eqn ide_some_terminal by (intro C.lunit_eqI) auto interpretation C: cartesian_monoidal_category C Prod \.map \ dup trm proof show "trm C.unity = C.unity" by (simp add: C.unity_def ide_some_terminal trm_one) show "\a. ide a \ C.runit a \ C.tensor a \[a] \ dup a = a" using comp_runit_term_dup runit_agreement by simp show "\a. ide a \ C.lunit a \ C.tensor \[a] a \ dup a = a" using comp_lunit_term_dup lunit_agreement by auto show "\a b. \ide a; ide b\ \ C.tensor (C.runit a \ C.tensor a \[b]) (C.lunit b \ C.tensor \[a] b) \ dup (C.tensor a b) = C.tensor a b" proof - fix a b assume a: "ide a" and b: "ide b" have "C.tensor (C.runit a \ C.tensor a \[b]) (C.lunit b \ C.tensor \[a] b) \ dup (C.tensor a b) = prod (C.runit a \ prod a \[b]) (C.lunit b \ prod \[a] b) \ dup (prod a b)" using a b by simp also have "... = tuple ((C.runit a \ prod a \[b]) \ prod a b) ((C.lunit b \ prod \[a] b) \ prod a b)" using a b ide_one trm_in_hom [of a] trm_in_hom [of b] unity_agreement prod_tuple by fastforce also have "... = tuple (C.runit a \ prod a \[b] \ prod a b) (C.lunit b \ prod \[a] b \ prod a b)" using comp_assoc by simp also have "... = tuple (C.runit a \ prod a \[b]) (C.lunit b \ prod \[a] b)" using a b comp_arr_dom by simp also have "... = tuple (runit a \ prod a \[b]) (lunit b \ prod \[a] b)" using a b lunit_agreement runit_agreement by simp also have "... = tuple (pr1 a b) (pr0 a b)" proof - have "runit a \ prod a \[b] = pr1 a b" using a b pr_naturality(2) trm_in_hom [of b] by (metis cod_pr1 comp_ide_arr ide_char in_homE pr_simps(4,6) seqI) moreover have "lunit b \ prod \[a] b = pr0 a b" using a b pr_naturality(1) [of b b b "\[a]" a \] trm_in_hom [of a] comp_cod_arr by (metis cod_pr0 ide_char in_homE pr_simps(1)) ultimately show ?thesis by simp qed also have "... = prod a b" using a b by simp finally show "C.tensor (C.runit a \ C.tensor a \[b]) (C.lunit b \ C.tensor \[a] b) \ dup (C.tensor a b) = C.tensor a b" by auto qed qed lemma extends_to_cartesian_monoidal_category: shows "cartesian_monoidal_category C Prod \.map \ dup trm" .. end text \ In a \cartesian_monoidal_category\, the monoidal structure is given by a categorical product and terminal object, so that the underlying category is cartesian. \ context cartesian_monoidal_category begin definition pr0 ("\

\<^sub>0[_, _]") where "\

\<^sub>0[a, b] \ if ide a \ ide b then \[a] \ (a \ \[b]) else null" definition pr1 ("\

\<^sub>1[_, _]") where "\

\<^sub>1[a, b] \ if ide a \ ide b then \[b] \ (\[a] \ b) else null" lemma pr_in_hom [intro]: assumes "ide a0" and "ide a1" shows "\\

\<^sub>0[a0, a1] : a0 \ a1 \ a0\" and "\\

\<^sub>1[a0, a1] : a0 \ a1 \ a1\" proof - show "\\

\<^sub>0[a0, a1] : a0 \ a1 \ a0\" unfolding pr0_def using assms runit_in_hom by fastforce show "\\

\<^sub>1[a0, a1] : a0 \ a1 \ a1\" unfolding pr1_def using assms lunit_in_hom by fastforce qed lemma pr_simps [simp]: assumes "ide a0" and "ide a1" shows "arr \

\<^sub>0[a0, a1]" and "dom \

\<^sub>0[a0, a1] = a0 \ a1" and "cod \

\<^sub>0[a0, a1] = a0" and "arr \

\<^sub>1[a0, a1]" and "dom \

\<^sub>1[a0, a1] = a0 \ a1" and "cod \

\<^sub>1[a0, a1] = a1" using assms pr_in_hom(1-2) by blast+ interpretation P: composite_functor CC.comp C CC.comp T \.map .. interpretation ECC: elementary_cartesian_category C pr1 pr0 \ \ proof show "\a b. \ide a; ide b\ \ span \

\<^sub>0[a, b] \

\<^sub>1[a, b]" by simp show "\a b. \ide a; ide b\ \ cod \

\<^sub>1[a, b] = b" by simp show "\a b. \ide a; ide b\ \ cod \

\<^sub>0[a, b] = a" by simp show "ide \" by simp show "\a. ide a \ \\ a : a \ \\" by auto show "\a f. \ide a; \f : a \ \\\ \ f = \[a]" using \\a. ide a \ \\[a] : a \ \\\ terminalE terminal_unity by blast show "\a b. \ (ide a \ ide b) \ \

\<^sub>1[a, b] = null" using pr1_def by auto show "\a b. \ (ide a \ ide b) \ \

\<^sub>0[a, b] = null" using pr0_def by auto show "\f g. span f g \ \!l. \

\<^sub>0[cod f, cod g] \ l = f \ \

\<^sub>1[cod f, cod g] \ l = g" proof - fix f g assume fg: "span f g" let ?l = "(f \ g) \ \[dom f]" have "\

\<^sub>0[cod f, cod g] \ ?l = f" proof - have "\

\<^sub>0[cod f, cod g] \ ?l = (\[cod f] \ (cod f \ \[cod g])) \ (f \ g) \ \[dom f]" using fg pr0_def by simp also have "... = \[cod f] \ ((cod f \ \[cod g]) \ (f \ g)) \ \[dom f]" using comp_assoc by simp also have "... = \[cod f] \ (f \ \[dom f]) \ \[dom f]" using fg interchange comp_cod_arr \.naturality by simp also have "... = \[cod f] \ ((cod f \ \[cod f]) \ (f \ f)) \ \[dom f]" using fg interchange comp_cod_arr \.naturality by simp also have "... = \[cod f] \ (cod f \ \[cod f]) \ (f \ f) \ \[dom f]" using comp_assoc by simp also have "... = (\[cod f] \ (cod f \ \[cod f]) \ \[cod f]) \ f" using fg \.naturality comp_assoc by simp also have "... = f" using fg pr0_dup comp_cod_arr by simp finally show ?thesis by blast qed moreover have "\

\<^sub>1[cod f, cod g] \ ?l = g" proof - have "\

\<^sub>1[cod f, cod g] \ ?l = \[cod g] \ ((\[cod f] \ cod g) \ (f \ g)) \ \[dom g]" using fg pr1_def comp_assoc by simp also have "... = \[cod g] \ (\[dom f] \ g) \ \[dom g]" using fg interchange comp_cod_arr \.naturality by simp also have "... = \[cod g] \ ((\[cod g] \ cod g) \ (g \ g)) \ \[dom g]" using fg interchange comp_cod_arr \.naturality by simp also have "... = \[cod g] \ (\[cod g] \ cod g) \ (g \ g) \ \[dom g]" using comp_assoc by simp also have "... = (\[cod g] \ (\[cod g] \ cod g) \ \[cod g]) \ g" using fg \.naturality comp_assoc by simp also have "... = g" using fg pr1_dup comp_cod_arr by simp finally show ?thesis by blast qed moreover have "\l. \ \

\<^sub>0[cod f, cod g] \ l = f; \

\<^sub>1[cod f, cod g] \ l = g \ \ l = ?l" proof - fix l assume f: "\

\<^sub>0[cod f, cod g] \ l = f" and g: "\

\<^sub>1[cod f, cod g] \ l = g" have l: "\l : dom f \ cod f \ cod g\" using f g fg by (metis arr_iff_in_hom dom_comp ide_cod pr_simps(5) seqE) have "?l = (\

\<^sub>0[cod f, cod g] \ l \ \

\<^sub>1[cod f, cod g] \ l) \ \[dom f]" using f g by simp also have "... = ((\

\<^sub>0[cod f, cod g] \ \

\<^sub>1[cod f, cod g]) \ (l \ l)) \ \[dom f]" using fg f g interchange by simp also have "... = (\

\<^sub>0[cod f, cod g] \ \

\<^sub>1[cod f, cod g]) \ (l \ l) \ \[dom f]" using comp_assoc by simp also have "... = ((\

\<^sub>0[cod f, cod g] \ \

\<^sub>1[cod f, cod g]) \ \[cod f \ cod g]) \ l" using l \.naturality [of l] comp_assoc by auto also have "... = (cod f \ cod g) \ l" using f g fg tuple_pr pr0_def pr1_def by auto also have "... = l" using l comp_cod_arr by auto finally show "l = ?l" by simp qed ultimately show "\!l. \

\<^sub>0[cod f, cod g] \ l = f \ \

\<^sub>1[cod f, cod g] \ l = g" by auto qed qed lemma extends_to_elementary_cartesian_category: shows "elementary_cartesian_category C pr1 pr0 \ \" .. lemma is_cartesian_category: shows "cartesian_category C" using ECC.is_cartesian_category by simp end (* * TODO: I would like to have coherence theorems for symmetric monoidal and cartesian * monoidal categories here, but I haven't yet figured out a suitably economical way * to extend the existing result. *) end