diff --git a/thys/Bicategory/Bicategory.thy b/thys/Bicategory/Bicategory.thy --- a/thys/Bicategory/Bicategory.thy +++ b/thys/Bicategory/Bicategory.thy @@ -1,2433 +1,2434 @@ (* Title: Bicategory Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) theory Bicategory imports Prebicategory Category3.Subcategory Category3.DiscreteCategory MonoidalCategory.MonoidalCategory begin section "Bicategories" text \ A \emph{bicategory} is a (vertical) category that has been equipped with a horizontal composition, an associativity natural isomorphism, and for each object a ``unit isomorphism'', such that horizontal composition on the left by target and on the right by source are fully faithful endofunctors of the vertical category, and such that the usual pentagon coherence condition holds for the associativity. \ locale bicategory = horizontal_composition V H src trg + \: natural_isomorphism VVV.comp V HoHV HoVH \\\\\. \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))\ + L: fully_faithful_functor V V L + R: fully_faithful_functor V V R for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" + assumes unit_in_vhom: "obj a \ \\[a] : a \ a \ a\" and iso_unit: "obj a \ iso \[a]" and pentagon: "\ ide f; ide g; ide h; ide k; src f = trg g; src g = trg h; src h = trg k \ \ (f \ \[g, h, k]) \ \[f, g \ h, k] \ (\[f, g, h] \ k) = \[f, g, h \ k] \ \[f \ g, h, k]" begin (* * TODO: the mapping \ is not currently assumed to be extensional. * It might be best in the long run if it were. *) definition \ where "\ \\\ \ \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))" lemma assoc_in_hom': assumes "arr \" and "arr \" and "arr \" and "src \ = trg \" and "src \ = trg \" shows "in_hhom \[\, \, \] (src \) (trg \)" and "\\[\, \, \] : (dom \ \ dom \) \ dom \ \ cod \ \ cod \ \ cod \\" proof - show "\\[\, \, \] : (dom \ \ dom \) \ dom \ \ cod \ \ cod \ \ cod \\" proof - have 1: "VVV.in_hom (\, \, \) (dom \, dom \, dom \) (cod \, cod \, cod \)" using assms VVV.in_hom_char VVV.arr_char VV.arr_char by auto have "\\[\, \, \] : HoHV (dom \, dom \, dom \) \ HoVH (cod \, cod \, cod \)\" using 1 \.preserves_hom by auto moreover have "HoHV (dom \, dom \, dom \) = (dom \ \ dom \) \ dom \" using 1 HoHV_def by (simp add: VVV.in_hom_char) moreover have "HoVH (cod \, cod \, cod \) = cod \ \ cod \ \ cod \" using 1 HoVH_def by (simp add: VVV.in_hom_char) ultimately show ?thesis by simp qed thus "in_hhom \[\, \, \] (src \) (trg \)" using assms src_cod trg_cod vconn_implies_hpar(1) vconn_implies_hpar(2) by auto qed lemma assoc_is_natural_1: assumes "arr \" and "arr \" and "arr \" and "src \ = trg \" and "src \ = trg \" shows "\[\, \, \] = (\ \ \ \ \) \ \[dom \, dom \, dom \]" using assms \.is_natural_1 [of "(\, \, \)"] VVV.arr_char VV.arr_char VVV.dom_char HoVH_def src_dom trg_dom by simp lemma assoc_is_natural_2: assumes "arr \" and "arr \" and "arr \" and "src \ = trg \" and "src \ = trg \" shows "\[\, \, \] = \[cod \, cod \, cod \] \ ((\ \ \) \ \)" using assms \.is_natural_2 [of "(\, \, \)"] VVV.arr_char VV.arr_char VVV.cod_char HoHV_def src_dom trg_dom by simp lemma assoc_naturality: assumes "arr \" and "arr \" and "arr \" and "src \ = trg \" and "src \ = trg \" shows "\[cod \, cod \, cod \] \ ((\ \ \) \ \) = (\ \ \ \ \) \ \[dom \, dom \, dom \]" using assms \.naturality VVV.arr_char VV.arr_char HoVH_def HoHV_def VVV.dom_char VVV.cod_char by auto lemma assoc_in_hom [intro]: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "in_hhom \[f, g, h] (src h) (trg f)" and "\\[f, g, h] : (dom f \ dom g) \ dom h \ cod f \ cod g \ cod h\" using assms assoc_in_hom' apply auto[1] using assms assoc_in_hom' ideD(1) by metis lemma assoc_simps [simp]: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "arr \[f, g, h]" and "src \[f, g, h] = src h" and "trg \[f, g, h] = trg f" and "dom \[f, g, h] = (dom f \ dom g) \ dom h" and "cod \[f, g, h] = cod f \ cod g \ cod h" using assms assoc_in_hom apply auto using assoc_in_hom(1) by auto lemma iso_assoc [intro, simp]: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "iso \[f, g, h]" using assms \.components_are_iso [of "(f, g, h)"] VVV.ide_char VVV.arr_char VV.arr_char by simp end subsection "Categories Induce Bicategories" text \ In this section we show that a category becomes a bicategory if we take the vertical composition to be discrete, we take the composition of the category as the horizontal composition, and we take the vertical domain and codomain as \src\ and \trg\. \ (* * It is helpful to make a few local definitions here, but I don't want them to * clutter the category locale. Using a context and private definitions does not * work as expected. So we have to define a new locale just for the present purpose. *) locale category_as_bicategory = category begin interpretation V: discrete_category \Collect arr\ null using not_arr_null by (unfold_locales, blast) abbreviation V where "V \ V.comp" interpretation src: "functor" V V dom using V.null_char by (unfold_locales, simp add: has_domain_iff_arr dom_def, auto) interpretation trg: "functor" V V cod using V.null_char by (unfold_locales, simp add: has_codomain_iff_arr cod_def, auto) interpretation H: horizontal_homs V dom cod by (unfold_locales, auto) interpretation H: "functor" H.VV.comp V \\\\. fst \\ \ snd \\\ apply (unfold_locales) using H.VV.arr_char V.null_char ext apply force using H.VV.arr_char V.null_char H.VV.dom_char H.VV.cod_char apply auto[3] proof - show "\g f. H.VV.seq g f \ fst (H.VV.comp g f) \ snd (H.VV.comp g f) = V (fst g \ snd g) (fst f \ snd f)" proof - have 0: "\f. H.VV.arr f \ V.arr (fst f \ snd f)" using H.VV.arr_char by auto have 1: "\f g. V.seq g f \ V.ide f \ g = f" using V.arr_char V.dom_char V.cod_char V.not_arr_null by force have 2: "\f g. H.VxV.seq g f \ H.VxV.ide f \ g = f" using 1 H.VxV.seq_char by (metis H.VxV.dom_eqI H.VxV.ide_Ide) fix f g assume fg: "H.VV.seq g f" have 3: "H.VV.ide f \ f = g" using fg 2 H.VV.seq_char H.VV.ide_char by blast show "fst (H.VV.comp g f) \ snd (H.VV.comp g f) = V (fst g \ snd g) (fst f \ snd f)" using fg 0 1 3 H.VV.comp_char H.VV.arr_char H.VV.ide_char V.arr_char V.comp_char H.VV.comp_arr_ide by (metis (no_types, lifting)) qed qed interpretation H: horizontal_composition V C dom cod by (unfold_locales, auto) abbreviation \ where "\ f g h \ f \ g \ h" interpretation \: natural_isomorphism H.VVV.comp V H.HoHV H.HoVH \\\\\. \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))\ apply unfold_locales using V.null_char ext apply fastforce using H.HoHV_def H.HoVH_def H.VVV.arr_char H.VV.arr_char H.VVV.dom_char H.VV.dom_char H.VVV.cod_char H.VV.cod_char H.VVV.ide_char comp_assoc by auto interpretation fully_faithful_functor V V H.R using comp_arr_dom by (unfold_locales, auto) interpretation fully_faithful_functor V V H.L using comp_cod_arr by (unfold_locales, auto) abbreviation \ where "\ \ \x. x" proposition induces_bicategory: shows "bicategory V C \ \ dom cod" apply (unfold_locales, auto simp add: comp_assoc) using comp_arr_dom by fastforce end subsection "Monoidal Categories induce Bicategories" text \ In this section we show that our definition of bicategory directly generalizes our definition of monoidal category: a monoidal category becomes a bicategory when equipped with the constant-\\\ functors as src and trg and \\\ as the unit isomorphism from \\ \ \\ to \\\. There is a slight mismatch because the bicategory locale assumes that the associator is given in curried form, whereas for monoidal categories it is given in tupled form. Ultimately, the monoidal category locale should be revised to also use curried form, which ends up being more convenient in most situations. \ context monoidal_category begin interpretation I: constant_functor C C \ using \_in_hom by unfold_locales auto interpretation horizontal_homs C I.map I.map by unfold_locales auto lemma CC_eq_VV: shows "CC.comp = VV.comp" proof - have "\g f. CC.comp g f = VV.comp g f" proof - fix f g show "CC.comp g f = VV.comp g f" proof - have "CC.seq g f \ CC.comp g f = VV.comp g f" using VV.comp_char VV.arr_char CC.seq_char by (elim CC.seqE seqE, simp) moreover have "\ CC.seq g f \ CC.comp g f = VV.comp g f" using VV.seq_char VV.ext VV.null_char CC.ext by (metis (no_types, lifting)) ultimately show ?thesis by blast qed qed thus ?thesis by blast qed lemma CCC_eq_VVV: shows "CCC.comp = VVV.comp" proof - have "\g f. CCC.comp g f = VVV.comp g f" proof - fix f g show "CCC.comp g f = VVV.comp g f" proof - have "CCC.seq g f \ CCC.comp g f = VVV.comp g f" - using VVV.comp_char VVV.arr_char CCC.seq_char VV.arr_char - by (elim CCC.seqE CC.seqE seqE, simp) + by (metis (no_types, lifting) CC.arrE CCC.seqE CC_eq_VV I.map_simp + I.preserves_reflects_arr VV.seq_char VVV.arrI VVV.comp_simp VVV.seq_char + trg_vcomp vseq_implies_hpar(1)) moreover have "\ CCC.seq g f \ CCC.comp g f = VVV.comp g f" using VVV.seq_char VVV.ext VVV.null_char CCC.ext by (metis (no_types, lifting)) ultimately show ?thesis by blast qed qed thus ?thesis by blast qed interpretation H: "functor" VV.comp C \\\\. fst \\ \ snd \\\ using CC_eq_VV T.functor_axioms by simp interpretation H: horizontal_composition C tensor I.map I.map by (unfold_locales, simp_all) lemma HoHV_eq_ToTC: shows "H.HoHV = T.ToTC" using H.HoHV_def T.ToTC_def CCC_eq_VVV by presburger lemma HoVH_eq_ToCT: shows "H.HoVH = T.ToCT" using H.HoVH_def T.ToCT_def CCC_eq_VVV by presburger interpretation \: natural_isomorphism VVV.comp C H.HoHV H.HoVH \ using \.natural_isomorphism_axioms CCC_eq_VVV HoHV_eq_ToTC HoVH_eq_ToCT by simp lemma R'_eq_R: shows "H.R = R" using H.is_extensional CC_eq_VV CC.arr_char by force lemma L'_eq_L: shows "H.L = L" using H.is_extensional CC_eq_VV CC.arr_char by force interpretation R': fully_faithful_functor C C H.R using R'_eq_R R.fully_faithful_functor_axioms unity_def by auto interpretation L': fully_faithful_functor C C H.L using L'_eq_L L.fully_faithful_functor_axioms unity_def by auto lemma obj_char: shows "obj a \ a = \" using obj_def [of a] \_in_hom by fastforce proposition induces_bicategory: shows "bicategory C tensor (\\ \ \. \ (\, \, \)) (\_. \) I.map I.map" using obj_char \_in_hom \_is_iso pentagon \.is_extensional \.is_natural_1 \.is_natural_2 by unfold_locales simp_all end subsection "Prebicategories Extend to Bicategories" text \ In this section, we show that a prebicategory with homs and units extends to a bicategory. The main work is to show that the endofunctors \L\ and \R\ are fully faithful. We take the left and right unitor isomorphisms, which were obtained via local constructions in the left and right hom-subcategories defined by a specified weak unit, and show that in the presence of the chosen sources and targets they are the components of a global natural isomorphisms \\\ and \\\ from the endofunctors \L\ and \R\ to the identity functor. A consequence is that functors \L\ and \R\ are endo-equivalences, hence fully faithful. \ context prebicategory_with_homs begin text \ Once it is equipped with a particular choice of source and target for each arrow, a prebicategory determines a horizontal composition. \ lemma induces_horizontal_composition: shows "horizontal_composition V H src trg" proof - interpret H: "functor" VV.comp V \\\\. fst \\ \ snd \\\ proof - have "VV.comp = VoV.comp" using composable_char\<^sub>P\<^sub>B\<^sub>H by meson thus "functor VV.comp V (\\\. fst \\ \ snd \\)" using functor_axioms by argo qed show "horizontal_composition V H src trg" using src_hcomp' trg_hcomp' composable_char\<^sub>P\<^sub>B\<^sub>H not_arr_null by (unfold_locales; metis) qed end sublocale prebicategory_with_homs \ horizontal_composition V H src trg using induces_horizontal_composition by auto locale prebicategory_with_homs_and_units = prebicategory_with_units + prebicategory_with_homs begin no_notation in_hom ("\_ : _ \ _\") text \ The next definitions extend the left and right unitors that were defined locally with respect to a particular weak unit, to globally defined versions using the chosen source and target for each arrow. \ definition lunit ("\[_]") where "lunit f \ left_hom_with_unit.lunit V H \ \[trg f] (trg f) f" definition runit ("\[_]") where "runit f \ right_hom_with_unit.runit V H \ \[src f] (src f) f" lemma lunit_in_hom: assumes "ide f" shows "\\[f] : src f \\<^sub>W\<^sub>C trg f\" and "\\[f] : trg f \ f \ f\" proof - interpret Left: subcategory V \left (trg f)\ using assms left_hom_is_subcategory by simp interpret Left: left_hom_with_unit V H \ \\[trg f]\ \trg f\ using assms obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by (unfold_locales, auto) have 0: "Left.ide f" using assms Left.ide_char Left.arr_char left_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto show 1: "\\[f] : trg f \ f \ f\" unfolding lunit_def using assms 0 Left.lunit_char(1) Left.hom_char H\<^sub>L_def by auto show "\\[f] : src f \\<^sub>W\<^sub>C trg f\" using 1 src_cod trg_cod src_in_sources trg_in_targets by (metis arrI vconn_implies_hpar) qed lemma runit_in_hom: assumes "ide f" shows "\\[f] : src f \\<^sub>W\<^sub>C trg f\" and "\\[f] : f \ src f \ f\" proof - interpret Right: subcategory V \right (src f)\ using assms right_hom_is_subcategory weak_unit_self_composable by force interpret Right: right_hom_with_unit V H \ \\[src f]\ \src f\ using assms obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by (unfold_locales, auto) have 0: "Right.ide f" using assms Right.ide_char Right.arr_char right_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto show 1: "\\[f] : f \ src f \ f\" unfolding runit_def using assms 0 Right.runit_char(1) Right.hom_char H\<^sub>R_def by auto show "\\[f] : src f \\<^sub>W\<^sub>C trg f\" using 1 src_cod trg_cod src_in_sources trg_in_targets by (metis arrI vconn_implies_hpar) qed text \ The characterization of the locally defined unitors yields a corresponding characterization of the globally defined versions, by plugging in the chosen source or target for each arrow for the unspecified weak unit in the the local versions. \ lemma lunit_char: assumes "ide f" shows "\\[f] : src f \\<^sub>W\<^sub>C trg f\" and "\\[f] : trg f \ f \ f\" and "trg f \ \[f] = (\[trg f] \ f) \ inv \[trg f, trg f, f]" and "\!\. \\ : trg f \ f \ f\ \ trg f \ \ = (\[trg f] \ f) \ inv \[trg f, trg f, f]" proof - let ?a = "src f" and ?b = "trg f" interpret Left: subcategory V \left ?b\ using assms left_hom_is_subcategory weak_unit_self_composable by force interpret Left: left_hom_with_unit V H \ \\[?b]\ ?b using assms obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by (unfold_locales, auto) have 0: "Left.ide f" using assms Left.ide_char Left.arr_char left_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto show "\\[f] : src f \\<^sub>W\<^sub>C trg f\" using assms lunit_in_hom by simp show A: "\\[f] : trg f \ f \ f\" using assms lunit_in_hom by simp show B: "?b \ \[f] = (\[?b] \ f) \ inv \[?b, ?b, f]" unfolding lunit_def using 0 Left.lunit_char(2) H\<^sub>L_def by (metis Left.comp_simp Left.characteristic_iso(1-2) Left.seqI') show "\!\. \\ : trg f \ f \ f\ \ trg f \ \ = (\[?b] \ f) \ inv \[?b, ?b, f]" proof - have 1: "hom (trg f \ f) f = Left.hom (Left.L f) f" proof have 1: "Left.L f = ?b \ f" using 0 H\<^sub>L_def by simp show "Left.hom (Left.L f) f \ hom (?b \ f) f" using assms Left.hom_char [of "?b \ f" f] H\<^sub>L_def by simp show "hom (?b \ f) f \ Left.hom (Left.L f) f" using assms 1 ide_in_hom composable_char\<^sub>P\<^sub>B\<^sub>H hom_connected left_def Left.hom_char by auto qed let ?P = "\\. Left.in_hom \ (Left.L f) f" let ?P' = "\\. \\ : ?b \ f \ f\" let ?Q = "\\. Left.L \ = (\[?b] \ f) \ (inv \[?b, ?b, f])" let ?R = "\\. ?b \ \ = (\[?b] \ f) \ (inv \[?b, ?b, f])" have 2: "?P = ?P'" using 0 1 H\<^sub>L_def Left.hom_char by blast moreover have "\\. ?P \ \ (?Q \ \ ?R \)" using 2 Left.lunit_eqI H\<^sub>L_def by presburger moreover have "(\!\. ?P \ \ ?Q \)" using 0 2 A B Left.lunit_char(3) Left.ide_char Left.arr_char by (metis (no_types, lifting) Left.lunit_char(2) calculation(2) lunit_def) ultimately show ?thesis by metis qed qed lemma runit_char: assumes "ide f" shows "\\[f] : src f \\<^sub>W\<^sub>C trg f\" and "\\[f] : f \ src f \ f\" and "\[f] \ src f = (f \ \[src f]) \ \[f, src f, src f]" and "\!\. \\ : f \ src f \ f\ \ \ \ src f = (f \ \[src f]) \ \[f, src f, src f]" proof - let ?a = "src f" and ?b = "trg f" interpret Right: subcategory V \right ?a\ using assms right_hom_is_subcategory weak_unit_self_composable by force interpret Right: right_hom_with_unit V H \ \\[?a]\ ?a using assms obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by (unfold_locales, auto) have 0: "Right.ide f" using assms Right.ide_char Right.arr_char right_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto show "\\[f] : src f \\<^sub>W\<^sub>C trg f\" using assms runit_in_hom by simp show A: "\\[f] : f \ src f \ f\" using assms runit_in_hom by simp show B: "\[f] \ ?a = (f \ \[?a]) \ \[f, ?a, ?a]" unfolding runit_def using 0 Right.runit_char(2) H\<^sub>R_def using Right.comp_simp Right.characteristic_iso(4) Right.iso_is_arr by auto show "\!\. \\ : f \ src f \ f\ \ \ \ ?a = (f \ \[?a]) \ \[f, ?a, ?a]" proof - have 1: "hom (f \ ?a) f = Right.hom (Right.R f) f" proof have 1: "Right.R f = f \ ?a" using 0 H\<^sub>R_def by simp show "Right.hom (Right.R f) f \ hom (f \ ?a) f" using assms Right.hom_char [of "f \ ?a" f] H\<^sub>R_def by simp show "hom (f \ ?a) f \ Right.hom (Right.R f) f" using assms 1 ide_in_hom composable_char\<^sub>P\<^sub>B\<^sub>H hom_connected right_def Right.hom_char by auto qed let ?P = "\\. Right.in_hom \ (Right.R f) f" let ?P' = "\\. \\ : f \ ?a \ f\" let ?Q = "\\. Right.R \ = (f \ \[?a]) \ \[f, ?a, ?a]" let ?R = "\\. \ \ ?a = (f \ \[?a]) \ \[f, ?a, ?a]" have 2: "?P = ?P'" using 0 1 H\<^sub>R_def Right.hom_char by blast moreover have "\\. ?P \ \ (?Q \ \ ?R \)" using 2 Right.runit_eqI H\<^sub>R_def by presburger moreover have "(\!\. ?P \ \ ?Q \)" using 0 2 A B Right.runit_char(3) Right.ide_char Right.arr_char by (metis (no_types, lifting) Right.runit_char(2) calculation(2) runit_def) ultimately show ?thesis by metis qed qed lemma lunit_eqI: assumes "ide f" and "\\ : trg f \ f \ f\" and "trg f \ \ = (\[trg f] \ f) \ (inv \[trg f, trg f, f])" shows "\ = \[f]" using assms lunit_char(2-4) by blast lemma runit_eqI: assumes "ide f" and "\\ : f \ src f \ f\" and "\ \ src f = (f \ \[src f]) \ \[f, src f, src f]" shows "\ = \[f]" using assms runit_char(2-4) by blast lemma iso_lunit: assumes "ide f" shows "iso \[f]" proof - let ?b = "trg f" interpret Left: subcategory V \left ?b\ using assms left_hom_is_subcategory weak_unit_self_composable by force interpret Left: left_hom_with_unit V H \ \\[?b]\ ?b using assms obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by (unfold_locales, auto) show ?thesis proof - have 0: "Left.ide f" using assms Left.ide_char Left.arr_char left_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto thus ?thesis unfolding lunit_def using Left.iso_lunit Left.iso_char by blast qed qed lemma iso_runit: assumes "ide f" shows "iso \[f]" proof - let ?a = "src f" interpret Right: subcategory V \right ?a\ using assms right_hom_is_subcategory weak_unit_self_composable by force interpret Right: right_hom_with_unit V H \ \\[?a]\ ?a using assms obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by (unfold_locales, auto) show ?thesis proof - have 0: "Right.ide f" using assms Right.ide_char Right.arr_char right_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto thus ?thesis unfolding runit_def using Right.iso_runit Right.iso_char by blast qed qed lemma lunit_naturality: assumes "arr \" shows "\ \ \[dom \] = \[cod \] \ (trg \ \ \)" proof - let ?a = "src \" and ?b = "trg \" interpret Left: subcategory V \left ?b\ using assms obj_trg left_hom_is_subcategory weak_unit_self_composable by force interpret Left: left_hom_with_unit V H \ \\[?b]\ ?b using assms obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by (unfold_locales, auto) interpret Left.L: endofunctor \Left ?b\ Left.L using assms endofunctor_H\<^sub>L [of ?b] weak_unit_self_composable obj_trg obj_is_weak_unit by blast have 1: "Left.in_hom \ (dom \) (cod \)" using assms Left.hom_char Left.arr_char left_def composable_char\<^sub>P\<^sub>B\<^sub>H obj_trg by auto have 2: "Left.in_hom \[Left.dom \] (?b \ dom \) (dom \)" unfolding lunit_def using assms 1 Left.in_hom_char trg_dom Left.lunit_char(1) H\<^sub>L_def Left.arr_char Left.dom_char Left.ide_dom by force have 3: "Left.in_hom \[Left.cod \] (?b \ cod \) (cod \)" unfolding lunit_def using assms 1 Left.in_hom_char trg_cod Left.lunit_char(1) H\<^sub>L_def Left.cod_char Left.ide_cod by force have 4: "Left.in_hom (Left.L \) (?b \ dom \) (?b \ cod \)" using 1 Left.L.preserves_hom [of \ "dom \" "cod \"] H\<^sub>L_def by auto show ?thesis proof - have "\ \ \[dom \] = Left.comp \ \[Left.dom \]" using 1 2 Left.comp_simp by fastforce also have "... = Left.comp \ (Left.lunit (Left.dom \))" using assms 1 lunit_def by auto also have "... = Left.comp (Left.lunit (Left.cod \)) (Left.L \)" using 1 Left.lunit_naturality Left.cod_simp by auto also have "... = Left.comp (lunit (Left.cod \)) (Left.L \)" using assms 1 lunit_def by auto also have "... = \[cod \] \ Left.L \" using 1 3 4 Left.comp_char Left.cod_char Left.in_hom_char by auto also have "... = \[cod \] \ (trg \ \ \)" using 1 by (simp add: H\<^sub>L_def) finally show ?thesis by simp qed qed lemma runit_naturality: assumes "arr \" shows "\ \ \[dom \] = \[cod \] \ (\ \ src \)" proof - let ?a = "src \" and ?b = "trg \" interpret Right: subcategory V \right ?a\ using assms right_hom_is_subcategory weak_unit_self_composable by force interpret Right: right_hom_with_unit V H \ \\[?a]\ ?a using assms obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by (unfold_locales, auto) interpret Right.R: endofunctor \Right ?a\ Right.R using assms endofunctor_H\<^sub>R [of ?a] weak_unit_self_composable obj_src obj_is_weak_unit by blast have 1: "Right.in_hom \ (dom \) (cod \)" using assms Right.hom_char Right.arr_char right_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto have 2: "Right.in_hom \[Right.dom \] (dom \ \ ?a) (dom \)" unfolding runit_def using 1 Right.in_hom_char trg_dom Right.runit_char(1) [of "Right.dom \"] H\<^sub>R_def Right.arr_char Right.dom_char Right.ide_dom assms by force have 3: "\[Right.cod \] \ Right.hom (cod \ \ ?a) (cod \)" unfolding runit_def using 1 Right.in_hom_char trg_cod Right.runit_char(1) [of "Right.cod \"] H\<^sub>R_def Right.cod_char Right.ide_cod assms by force have 4: "Right.R \ \ Right.hom (dom \ \ ?a) (cod \ \ ?a)" using 1 Right.R.preserves_hom [of \ "dom \" "cod \"] H\<^sub>R_def by auto show ?thesis proof - have "\ \ \[dom \] = Right.comp \ \[Right.dom \]" by (metis 1 2 Right.comp_char Right.in_homE Right.seqI' Right.seq_char) also have "... = Right.comp \ (Right.runit (Right.dom \))" using assms 1 src_dom trg_dom Right.hom_char runit_def by auto also have "... = Right.comp (Right.runit (Right.cod \)) (Right.R \)" using 1 Right.runit_naturality Right.cod_simp by auto also have "... = Right.comp (runit (Right.cod \)) (Right.R \)" using assms 1 runit_def by auto also have "... = \[cod \] \ Right.R \" using 1 3 4 Right.comp_char Right.cod_char Right.in_hom_char by auto also have "... = \[cod \] \ (\ \ ?a)" using 1 by (simp add: H\<^sub>R_def) finally show ?thesis by simp qed qed interpretation L: endofunctor V L using endofunctor_L by auto interpretation \: transformation_by_components V V L map lunit using lunit_in_hom lunit_naturality by unfold_locales auto interpretation \: natural_isomorphism V V L map \.map using iso_lunit by unfold_locales auto lemma natural_isomorphism_\: shows "natural_isomorphism V V L map \.map" .. interpretation L: equivalence_functor V V L using L.isomorphic_to_identity_is_equivalence \.natural_isomorphism_axioms by simp lemma equivalence_functor_L: shows "equivalence_functor V V L" .. lemma lunit_commutes_with_L: assumes "ide f" shows "\[L f] = L \[f]" proof - have "seq \[f] (L \[f])" using assms lunit_char(2) L.preserves_hom by fastforce moreover have "seq \[f] \[L f]" using assms lunit_char(2) lunit_char(2) [of "L f"] L.preserves_ide by auto ultimately show ?thesis using assms lunit_char(2) [of f] lunit_naturality [of "\[f]"] iso_lunit iso_is_section section_is_mono monoE [of "\[f]" "L \[f]" "\[L f]"] by auto qed interpretation R: endofunctor V R using endofunctor_R by auto interpretation \: transformation_by_components V V R map runit using runit_in_hom runit_naturality by unfold_locales auto interpretation \: natural_isomorphism V V R map \.map using iso_runit by unfold_locales auto lemma natural_isomorphism_\: shows "natural_isomorphism V V R map \.map" .. interpretation R: equivalence_functor V V R using R.isomorphic_to_identity_is_equivalence \.natural_isomorphism_axioms by simp lemma equivalence_functor_R: shows "equivalence_functor V V R" .. lemma runit_commutes_with_R: assumes "ide f" shows "\[R f] = R \[f]" proof - have "seq \[f] (R \[f])" using assms runit_char(2) R.preserves_hom by fastforce moreover have "seq \[f] \[R f]" using assms runit_char(2) runit_char(2) [of "R f"] R.preserves_ide by auto ultimately show ?thesis using assms runit_char(2) [of f] runit_naturality [of "\[f]"] iso_runit iso_is_section section_is_mono monoE [of "\[f]" "R \[f]" "\[R f]"] by auto qed definition \ where "\ \ \ \ \ if VVV.arr (\, \, \) then (\ \ \ \ \) \ \[dom \, dom \, dom \] else null" lemma \_ide_simp [simp]: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "\ f g h = \[f, g, h]" proof - have "\ f g h = (f \ g \ h) \ \[dom f, dom g, dom h]" using assms \_def VVV.arr_char [of "(f, g, h)"] by auto also have "... = (f \ g \ h) \ \[f, g, h]" using assms by simp also have "... = \[f, g, h]" using assms \_def assoc_in_hom\<^sub>A\<^sub>W\<^sub>C hcomp_in_hom\<^sub>P\<^sub>B\<^sub>H VVV.arr_char VoV.arr_char comp_cod_arr composable_char\<^sub>P\<^sub>B\<^sub>H by auto finally show ?thesis by simp qed (* TODO: Figure out how this got reinstated. *) no_notation in_hom ("\_ : _ \ _\") lemma natural_isomorphism_\: shows "natural_isomorphism VVV.comp V HoHV HoVH (\\\\. \ (fst \\\) (fst (snd \\\)) (snd (snd \\\)))" proof - interpret \: transformation_by_components VVV.comp V HoHV HoVH \\f. \[fst f, fst (snd f), snd (snd f)]\ proof show 1: "\x. VVV.ide x \ \\[fst x, fst (snd x), snd (snd x)] : HoHV x \ HoVH x\" proof - fix x assume x: "VVV.ide x" show "\\[fst x, fst (snd x), snd (snd x)] : HoHV x \ HoVH x\" proof - have "ide (fst x) \ ide (fst (snd x)) \ ide (snd (snd x)) \ fst x \ fst (snd x) \ null \ fst (snd x) \ snd (snd x) \ null" using x VVV.ide_char VVV.arr_char VV.arr_char composable_char\<^sub>P\<^sub>B\<^sub>H by simp hence "\[fst x, fst (snd x), snd (snd x)] \ hom ((fst x \ fst (snd x)) \ snd (snd x)) (fst x \ fst (snd x) \ snd (snd x))" using x assoc_in_hom\<^sub>A\<^sub>W\<^sub>C by simp thus ?thesis unfolding HoHV_def HoVH_def using x VVV.ideD(1) by simp qed qed show "\f. VVV.arr f \ \[fst (VVV.cod f), fst (snd (VVV.cod f)), snd (snd (VVV.cod f))] \ HoHV f = HoVH f \ \[fst (VVV.dom f), fst (snd (VVV.dom f)), snd (snd (VVV.dom f))]" unfolding HoHV_def HoVH_def using assoc_naturality\<^sub>A\<^sub>W\<^sub>C VVV.arr_char VV.arr_char VVV.dom_char VVV.cod_char composable_char\<^sub>P\<^sub>B\<^sub>H by simp qed interpret \: natural_isomorphism VVV.comp V HoHV HoVH \.map proof fix f assume f: "VVV.ide f" show "iso (\.map f)" proof - have "fst f \ fst (snd f) \ null \ fst (snd f) \ snd (snd f) \ null" using f VVV.ideD(1) VVV.arr_char [of f] VV.arr_char composable_char\<^sub>P\<^sub>B\<^sub>H by auto thus ?thesis using f \.map_simp_ide iso_assoc\<^sub>A\<^sub>W\<^sub>C VVV.ide_char VVV.arr_char by simp qed qed have "(\\\\. \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))) = \.map" proof fix \\\ have "\ VVV.arr \\\ \ \ (fst \\\) (fst (snd \\\)) (snd (snd \\\)) = \.map \\\" using \_def \.map_def by simp moreover have "VVV.arr \\\ \ \ (fst \\\) (fst (snd \\\)) (snd (snd \\\)) = \.map \\\" proof - assume \\\: "VVV.arr \\\" have "\ (fst \\\) (fst (snd \\\)) (snd (snd \\\)) = (fst \\\ \ fst (snd \\\) \ snd (snd \\\)) \ \[dom (fst \\\), dom (fst (snd \\\)), dom (snd (snd \\\))]" using \\\ \_def by simp also have "... = \[cod (fst \\\), cod (fst (snd \\\)), cod (snd (snd \\\))] \ ((fst \\\ \ fst (snd \\\)) \ snd (snd \\\))" using \\\ HoHV_def HoVH_def VVV.arr_char VV.arr_char assoc_naturality\<^sub>A\<^sub>W\<^sub>C composable_char\<^sub>P\<^sub>B\<^sub>H by simp also have "... = \[fst (VVV.cod \\\), fst (snd (VVV.cod \\\)), snd (snd (VVV.cod \\\))] \ ((fst \\\ \ fst (snd \\\)) \ snd (snd \\\))" using \\\ VVV.arr_char VVV.cod_char VV.arr_char by simp also have "... = \.map \\\" using \\\ \.map_def HoHV_def composable_char\<^sub>P\<^sub>B\<^sub>H by auto finally show ?thesis by blast qed ultimately show "\ (fst \\\) (fst (snd \\\)) (snd (snd \\\)) = \.map \\\" by blast qed thus ?thesis using \.natural_isomorphism_axioms by simp qed proposition induces_bicategory: shows "bicategory V H \ \ src trg" proof - interpret VxVxV: product_category V VxV.comp .. interpret VoVoV: subcategory VxVxV.comp \\\\\. arr (fst \\\) \ VV.arr (snd \\\) \ src (fst \\\) = trg (fst (snd \\\))\ using subcategory_VVV by blast interpret HoHV: "functor" VVV.comp V HoHV using functor_HoHV by blast interpret HoVH: "functor" VVV.comp V HoVH using functor_HoVH by blast interpret \: natural_isomorphism VVV.comp V HoHV HoVH \\\\\. \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))\ using natural_isomorphism_\ by blast interpret L: equivalence_functor V V L using equivalence_functor_L by blast interpret R: equivalence_functor V V R using equivalence_functor_R by blast show "bicategory V H \ \ src trg" proof show "\a. obj a \ \\[a] : a \ a \ a\" using obj_is_weak_unit unit_in_vhom\<^sub>P\<^sub>B\<^sub>U by blast show "\a. obj a \ iso \[a]" using obj_is_weak_unit iso_unit\<^sub>P\<^sub>B\<^sub>U by blast show "\f g h k. \ ide f; ide g; ide h; ide k; src f = trg g; src g = trg h; src h = trg k \ \ (f \ \ g h k) \ \ f (g \ h) k \ (\ f g h \ k) = \ f g (h \ k) \ \ (f \ g) h k" proof - fix f g h k assume f: "ide f" and g: "ide g" and h: "ide h" and k: "ide k" and fg: "src f = trg g" and gh: "src g = trg h" and hk: "src h = trg k" have "sources f \ targets g \ {}" using f g fg src_in_sources [of f] trg_in_targets ideD(1) by auto moreover have "sources g \ targets h \ {}" using g h gh src_in_sources [of g] trg_in_targets ideD(1) by auto moreover have "sources h \ targets k \ {}" using h k hk src_in_sources [of h] trg_in_targets ideD(1) by auto moreover have "\ f g h = \[f, g, h] \ \ g h k = \[g, h, k]" using f g h k fg gh hk \_ide_simp by simp moreover have "\ f (g \ h) k = \[f, g \ h, k] \ \ f g (h \ k) = \[f, g, h \ k] \ \ (f \ g) h k = \[f \ g, h, k]" using f g h k fg gh hk \_ide_simp preserves_ide hcomp_in_hom\<^sub>P\<^sub>B\<^sub>H(1) by simp ultimately show "(f \ \ g h k) \ \ f (g \ h) k \ (\ f g h \ k) = \ f g (h \ k) \ \ (f \ g) h k" using f g h k fg gh hk pentagon\<^sub>A\<^sub>W\<^sub>C [of f g h k] \_ide_simp by presburger qed qed qed end text \ The following is the main result of this development: Every prebicategory extends to a bicategory, by making an arbitrary choice of representatives of each isomorphism class of weak units and using that to define the source and target mappings, and then choosing an arbitrary isomorphism in \hom (a \ a) a\ for each weak unit \a\. \ context prebicategory begin interpretation prebicategory_with_homs V H \ some_src some_trg using extends_to_prebicategory_with_homs by auto interpretation prebicategory_with_units V H \ some_unit using extends_to_prebicategory_with_units by auto interpretation prebicategory_with_homs_and_units V H \ some_unit some_src some_trg .. theorem extends_to_bicategory: shows "bicategory V H \ some_unit some_src some_trg" using induces_bicategory by simp end section "Bicategories as Prebicategories" subsection "Bicategories are Prebicategories" text \ In this section we show that a bicategory determines a prebicategory with homs, whose weak units are exactly those arrows that are isomorphic to their chosen source, or equivalently, to their chosen target. Moreover, the notion of horizontal composability, which in a bicategory is determined by the coincidence of chosen sources and targets, agrees with the version defined for the induced weak composition in terms of nonempty intersections of source and target sets, which is not dependent on any arbitrary choices. \ context bicategory begin (* TODO: Why does this get re-introduced? *) no_notation in_hom ("\_ : _ \ _\") interpretation \': inverse_transformation VVV.comp V HoHV HoVH \\\\\. \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))\ .. abbreviation \' where "\' \ \'.map" definition \' ("\\<^sup>-\<^sup>1[_, _, _]") where "\\<^sup>-\<^sup>1[\, \, \] \ \'.map (\, \, \)" lemma assoc'_in_hom': assumes "arr \" and "arr \" and "arr \" and "src \ = trg \" and "src \ = trg \" shows "in_hhom \\<^sup>-\<^sup>1[\, \, \] (src \) (trg \)" and "\\\<^sup>-\<^sup>1[\, \, \] : dom \ \ dom \ \ dom \ \ (cod \ \ cod \) \ cod \\" proof - show "\\\<^sup>-\<^sup>1[\, \, \] : dom \ \ dom \ \ dom \ \ (cod \ \ cod \) \ cod \\" proof - have 1: "VVV.in_hom (\, \, \) (dom \, dom \, dom \) (cod \, cod \, cod \)" using assms VVV.in_hom_char VVV.arr_char VV.arr_char by auto have "\\\<^sup>-\<^sup>1[\, \, \] : HoVH (dom \, dom \, dom \) \ HoHV (cod \, cod \, cod \)\" using 1 \'_def \'.preserves_hom by auto moreover have "HoVH (dom \, dom \, dom \) = dom \ \ dom \ \ dom \" using 1 HoVH_def by (simp add: VVV.in_hom_char) moreover have "HoHV (cod \, cod \, cod \) = (cod \ \ cod \) \ cod \" using 1 HoHV_def by (simp add: VVV.in_hom_char) ultimately show ?thesis by simp qed thus "in_hhom \\<^sup>-\<^sup>1[\, \, \] (src \) (trg \)" using assms vconn_implies_hpar(1) vconn_implies_hpar(2) by auto qed lemma assoc'_is_natural_1: assumes "arr \" and "arr \" and "arr \" and "src \ = trg \" and "src \ = trg \" shows "\\<^sup>-\<^sup>1[\, \, \] = ((\ \ \) \ \) \ \\<^sup>-\<^sup>1[dom \, dom \, dom \]" using assms \'.is_natural_1 [of "(\, \, \)"] VVV.arr_char VV.arr_char VVV.dom_char HoHV_def src_dom trg_dom \'_def by simp lemma assoc'_is_natural_2: assumes "arr \" and "arr \" and "arr \" and "src \ = trg \" and "src \ = trg \" shows "\\<^sup>-\<^sup>1[\, \, \] = \\<^sup>-\<^sup>1[cod \, cod \, cod \] \ (\ \ \ \ \)" using assms \'.is_natural_2 [of "(\, \, \)"] VVV.arr_char VV.arr_char VVV.cod_char HoVH_def src_dom trg_dom \'_def by simp lemma assoc'_naturality: assumes "arr \" and "arr \" and "arr \" and "src \ = trg \" and "src \ = trg \" shows "\\<^sup>-\<^sup>1[cod \, cod \, cod \] \ (\ \ \ \ \) = ((\ \ \) \ \) \ \\<^sup>-\<^sup>1[dom \, dom \, dom \]" using assms assoc'_is_natural_1 assoc'_is_natural_2 by metis lemma assoc'_in_hom [intro]: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "in_hhom \\<^sup>-\<^sup>1[f, g, h] (src h) (trg f)" and "\\\<^sup>-\<^sup>1[f, g, h] : dom f \ dom g \ dom h \ (cod f \ cod g) \ cod h\" using assms assoc'_in_hom'(1-2) ideD(1) by meson+ lemma assoc'_simps [simp]: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "arr \\<^sup>-\<^sup>1[f, g, h]" and "src \\<^sup>-\<^sup>1[f, g, h] = src h" and "trg \\<^sup>-\<^sup>1[f, g, h] = trg f" and "dom \\<^sup>-\<^sup>1[f, g, h] = dom f \ dom g \ dom h" and "cod \\<^sup>-\<^sup>1[f, g, h] = (cod f \ cod g) \ cod h" using assms assoc'_in_hom by blast+ lemma assoc'_eq_inv_assoc [simp]: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "\\<^sup>-\<^sup>1[f, g, h] = inv \[f, g, h]" using assms VVV.ide_char VVV.arr_char VV.ide_char VV.arr_char \'.map_ide_simp \'_def by auto lemma inverse_assoc_assoc' [intro]: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "inverse_arrows \[f, g, h] \\<^sup>-\<^sup>1[f, g, h]" using assms VVV.ide_char VVV.arr_char VV.ide_char VV.arr_char \'.map_ide_simp \'.inverts_components \'_def by auto lemma iso_assoc' [intro, simp]: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "iso \\<^sup>-\<^sup>1[f, g, h]" using assms by simp lemma comp_assoc_assoc' [simp]: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "\[f, g, h] \ \\<^sup>-\<^sup>1[f, g, h] = f \ g \ h" and "\\<^sup>-\<^sup>1[f, g, h] \ \[f, g, h] = (f \ g) \ h" using assms comp_arr_inv' comp_inv_arr' by auto lemma unit_in_hom [intro, simp]: assumes "obj a" shows "\\[a] : a \ a\" and "\\[a] : a \ a \ a\" proof - show "\\[a] : a \ a \ a\" using assms unit_in_vhom by simp thus "\\[a] : a \ a\" using assms by (metis arrI in_hhom_def obj_simps(2-3) vconn_implies_hpar(1-4)) qed interpretation weak_composition V H using is_weak_composition by auto lemma seq_if_composable: assumes "\ \ \ \ null" shows "src \ = trg \" using assms H.is_extensional [of "(\, \)"] VV.arr_char by auto lemma obj_self_composable: assumes "obj a" shows "a \ a \ null" and "isomorphic (a \ a) a" proof - show 1: "isomorphic (a \ a) a" using assms unit_in_hom iso_unit isomorphic_def by blast obtain \ where \: "iso \ \ \\ : a \ a \ a\" using 1 isomorphic_def by blast have "ide (a \ a)" using 1 \ ide_dom [of \] by fastforce thus "a \ a \ null" using ideD(1) not_arr_null by metis qed lemma obj_is_weak_unit: assumes "obj a" shows "weak_unit a" proof - interpret Left_a: subcategory V \left a\ using assms left_hom_is_subcategory by force interpret Right_a: subcategory V \right a\ using assms right_hom_is_subcategory by force text \ We know that \H\<^sub>L a\ is fully faithful as a global endofunctor, but the definition of weak unit involves its restriction to a subcategory. So we have to verify that the restriction is also a fully faithful functor. \ interpret La: endofunctor \Left a\ \H\<^sub>L a\ using assms obj_self_composable endofunctor_H\<^sub>L [of a] by force interpret La: fully_faithful_functor \Left a\ \Left a\ \H\<^sub>L a\ proof show "\f f'. Left_a.par f f' \ H\<^sub>L a f = H\<^sub>L a f' \ f = f'" proof - fix \ \' assume par: "Left_a.par \ \'" assume eq: "H\<^sub>L a \ = H\<^sub>L a \'" have 1: "par \ \'" using par Left_a.arr_char Left_a.dom_char Left_a.cod_char left_def composable_implies_arr null_agreement by metis moreover have "L \ = L \'" using par eq H\<^sub>L_def Left_a.arr_char left_def preserves_arr assms 1 seq_if_composable [of a \] not_arr_null seq_if_composable [of a \'] by auto ultimately show "\ = \'" using L.is_faithful by blast qed show "\f g \. \ Left_a.ide f; Left_a.ide g; Left_a.in_hom \ (H\<^sub>L a f) (H\<^sub>L a g) \ \ \\. Left_a.in_hom \ f g \ H\<^sub>L a \ = \" proof - fix f g \ assume f: "Left_a.ide f" and g: "Left_a.ide g" and \: "Left_a.in_hom \ (H\<^sub>L a f) (H\<^sub>L a g)" have 1: "a = trg f \ a = trg g" using assms f g Left_a.ide_char Left_a.arr_char left_def seq_if_composable [of a f] seq_if_composable [of a g] by auto show "\\. Left_a.in_hom \ f g \ H\<^sub>L a \ = \" proof - have 2: "\\. \\ : f \ g\ \ L \ = \" using f g \ 1 Left_a.ide_char H\<^sub>L_def L.preserves_reflects_arr Left_a.arr_char Left_a.in_hom_char L.is_full by force obtain \ where \: "\\ : f \ g\ \ L \ = \" using 2 by blast have "Left_a.arr \" using \ 1 trg_dom Left_a.arr_char left_def hseq_char' by fastforce moreover have "H\<^sub>L a \ = \" using \ 1 trg_dom H\<^sub>L_def by auto ultimately show ?thesis using \ Left_a.dom_simp Left_a.cod_simp by blast qed qed qed interpret Ra: endofunctor \Right a\ \H\<^sub>R a\ using assms obj_self_composable endofunctor_H\<^sub>R [of a] by force interpret Ra: fully_faithful_functor \Right a\ \Right a\ \H\<^sub>R a\ proof show "\f f'. Right_a.par f f' \ H\<^sub>R a f = H\<^sub>R a f' \ f = f'" proof - fix \ \' assume par: "Right_a.par \ \'" assume eq: "H\<^sub>R a \ = H\<^sub>R a \'" have 1: "par \ \'" using par Right_a.arr_char Right_a.dom_char Right_a.cod_char right_def composable_implies_arr null_agreement by metis moreover have "R \ = R \'" using par eq H\<^sub>R_def Right_a.arr_char right_def preserves_arr assms 1 seq_if_composable [of \ a] not_arr_null seq_if_composable [of \' a] by auto ultimately show "\ = \'" using R.is_faithful by blast qed show "\f g \. \ Right_a.ide f; Right_a.ide g; Right_a.in_hom \ (H\<^sub>R a f) (H\<^sub>R a g) \ \ \\. Right_a.in_hom \ f g \ H\<^sub>R a \ = \" proof - fix f g \ assume f: "Right_a.ide f" and g: "Right_a.ide g" and \: "Right_a.in_hom \ (H\<^sub>R a f) (H\<^sub>R a g)" have 1: "a = src f \ a = src g" using assms f g Right_a.ide_char Right_a.arr_char right_def seq_if_composable by auto show "\\. Right_a.in_hom \ f g \ H\<^sub>R a \ = \" proof - have 2: "\\. \\ : f \ g\ \ R \ = \" using f g \ 1 Right_a.ide_char H\<^sub>R_def R.preserves_reflects_arr Right_a.arr_char Right_a.in_hom_char R.is_full by force obtain \ where \: "\\ : f \ g\ \ R \ = \" using 2 by blast have "Right_a.arr \" using \ 1 src_dom Right_a.arr_char right_def hseq_char' by fastforce moreover have "H\<^sub>R a \ = \" using \ 1 src_dom H\<^sub>R_def by auto ultimately show ?thesis using \ Right_a.dom_simp Right_a.cod_simp by blast qed qed qed have "isomorphic (a \ a) a \ a \ a \ null" using assms obj_self_composable unit_in_hom iso_unit isomorphic_def by blast thus ?thesis using La.fully_faithful_functor_axioms Ra.fully_faithful_functor_axioms weak_unit_def by blast qed lemma src_in_sources: assumes "arr \" shows "src \ \ sources \" using assms obj_is_weak_unit R.preserves_arr hseq_char' by auto lemma trg_in_targets: assumes "arr \" shows "trg \ \ targets \" using assms obj_is_weak_unit L.preserves_arr hseq_char' by auto lemma weak_unit_cancel_left: assumes "weak_unit a" and "ide f" and "ide g" and "a \ f \ a \ g" shows "f \ g" proof - have 0: "ide a" using assms weak_unit_def by force interpret Left_a: subcategory V \left a\ using 0 left_hom_is_subcategory by simp interpret Left_a: left_hom V H a using assms weak_unit_self_composable by unfold_locales auto interpret La: fully_faithful_functor \Left a\ \Left a\ \H\<^sub>L a\ using assms weak_unit_def by fast obtain \ where \: "iso \ \ \\ : a \ f \ a \ g\" using assms by blast have 1: "Left_a.iso \ \ Left_a.in_hom \ (a \ f) (a \ g)" proof have "a \ \ \ null" proof - have "a \ dom \ \ null" using assms \ weak_unit_self_composable by (metis arr_dom_iff_arr hseq_char' in_homE match_4) thus ?thesis using hom_connected by simp qed thus "Left_a.in_hom \ (a \ f) (a \ g)" using \ Left_a.hom_char left_def by auto thus "Left_a.iso \" using \ Left_a.iso_char by auto qed hence 2: "Left_a.ide (a \ f) \ Left_a.ide (a \ g)" using Left_a.ide_dom [of \] Left_a.ide_cod [of \] Left_a.dom_simp Left_a.cod_simp by auto hence 3: "Left_a.ide f \ Left_a.ide g" by (metis Left_a.ideI Left_a.ide_def Left_a.null_char assms(2) assms(3) left_def) obtain \ where \: "\ \ Left_a.hom f g \ a \ \ = \" using assms 1 2 3 La.is_full [of g f \] H\<^sub>L_def by auto have "Left_a.iso \" using \ 1 H\<^sub>L_def La.reflects_iso by auto hence "iso \ \ \\ : f \ g\" using \ Left_a.iso_char Left_a.in_hom_char by auto thus ?thesis by auto qed lemma weak_unit_cancel_right: assumes "weak_unit a" and "ide f" and "ide g" and "f \ a \ g \ a" shows "f \ g" proof - have 0: "ide a" using assms weak_unit_def by force interpret Right_a: subcategory V \right a\ using 0 right_hom_is_subcategory by simp interpret Right_a: right_hom V H a using assms weak_unit_self_composable by unfold_locales auto interpret R: fully_faithful_functor \Right a\ \Right a\ \H\<^sub>R a\ using assms weak_unit_def by fast obtain \ where \: "iso \ \ in_hom \ (f \ a) (g \ a)" using assms by blast have 1: "Right_a.iso \ \ \ \ Right_a.hom (f \ a) (g \ a)" proof have "\ \ a \ null" proof - have "dom \ \ a \ null" using assms \ weak_unit_self_composable by (metis arr_dom_iff_arr hseq_char' in_homE match_3) thus ?thesis using hom_connected by simp qed thus "\ \ Right_a.hom (f \ a) (g \ a)" using \ Right_a.hom_char right_def by simp thus "Right_a.iso \" using \ Right_a.iso_char by auto qed hence 2: "Right_a.ide (f \ a) \ Right_a.ide (g \ a)" using Right_a.ide_dom [of \] Right_a.ide_cod [of \] Right_a.dom_simp Right_a.cod_simp by auto hence 3: "Right_a.ide f \ Right_a.ide g" using assms Right_a.ide_char Right_a.arr_char right_def Right_a.ide_def Right_a.null_char by metis obtain \ where \: "\ \ Right_a.hom f g \ \ \ a = \" using assms 1 2 3 R.is_full [of g f \] H\<^sub>R_def by auto have "Right_a.iso \" using \ 1 H\<^sub>R_def R.reflects_iso by auto hence "iso \ \ \\ : f \ g\" using \ Right_a.iso_char Right_a.in_hom_char by auto thus ?thesis by auto qed text \ All sources of an arrow ({\em i.e.}~weak units composable on the right with that arrow) are isomorphic to the chosen source, and similarly for targets. That these statements hold was somewhat surprising to me. \ lemma source_iso_src: assumes "arr \" and "a \ sources \" shows "a \ src \" proof - have 0: "ide a" using assms weak_unit_def by force have 1: "src a = trg a" using assms ide_dom sources_def weak_unit_iff_self_target seq_if_composable weak_unit_self_composable by simp obtain \ where \: "iso \ \ \\ : a \ a \ a\" using assms weak_unit_def by blast have "a \ src a \ src a \ src a" proof - have "src a \ src a \ src a" using 0 obj_is_weak_unit weak_unit_def isomorphic_symmetric by auto moreover have "a \ src a \ src a" proof - have "a \ a \ src a \ a \ src a" proof - have "iso (\ \ src a) \ \\ \ src a : (a \ a) \ src a \ a \ src a\" using 0 1 \ ide_in_hom(2) by auto moreover have "iso \\<^sup>-\<^sup>1[a, a, src a] \ \\\<^sup>-\<^sup>1[a, a, src a] : a \ a \ src a \ (a \ a) \ src a\" using 0 1 iso_assoc' by force ultimately show ?thesis using isos_compose isomorphic_def by auto qed thus ?thesis using assms 0 weak_unit_cancel_left by auto qed ultimately show ?thesis using isomorphic_transitive by meson qed hence "a \ src a" using 0 weak_unit_cancel_right [of "src a" a "src a"] obj_is_weak_unit by auto thus ?thesis using assms seq_if_composable 1 by auto qed lemma target_iso_trg: assumes "arr \" and "b \ targets \" shows "b \ trg \" proof - have 0: "ide b" using assms weak_unit_def by force have 1: "trg \ = src b" using assms seq_if_composable by auto obtain \ where \: "iso \ \ \\ : b \ b \ b\" using assms weak_unit_def by blast have "trg b \ b \ trg b \ trg b" proof - have "trg b \ trg b \ trg b" using 0 obj_is_weak_unit weak_unit_def isomorphic_symmetric by auto moreover have "trg b \ b \ trg b" proof - have "(trg b \ b) \ b \ trg b \ b" proof - have "iso (trg b \ \) \ \trg b \ \ : trg b \ b \ b \ trg b \ b\" using assms 0 1 \ ide_in_hom(2) targetsD(1) weak_unit_self_composable apply (intro conjI hcomp_in_vhom) by auto moreover have "iso \[trg b, b, b] \ \\[trg b, b, b] : (trg b \ b) \ b \ trg b \ b \ b\" using assms(2) 0 1 seq_if_composable targetsD(1-2) weak_unit_self_composable by auto ultimately show ?thesis using isos_compose isomorphic_def by auto qed thus ?thesis using assms 0 weak_unit_cancel_right by auto qed ultimately show ?thesis using isomorphic_transitive by meson qed hence "b \ trg b" using 0 weak_unit_cancel_left [of "trg b" b "trg b"] obj_is_weak_unit by simp thus ?thesis using assms 0 1 seq_if_composable weak_unit_iff_self_source targetsD(1-2) source_iso_src by simp qed lemma is_weak_composition_with_homs: shows "weak_composition_with_homs V H src trg" using src_in_sources trg_in_targets seq_if_composable composable_implies_arr by (unfold_locales, simp_all) interpretation weak_composition_with_homs V H src trg using is_weak_composition_with_homs by auto text \ In a bicategory, the notion of composability defined in terms of the chosen sources and targets coincides with the version defined for a weak composition, which does not involve particular choices. \ lemma connected_iff_seq: assumes "arr \" and "arr \" shows "sources \ \ targets \ \ {} \ src \ = trg \" proof show "src \ = trg \ \ sources \ \ targets \ \ {}" using assms src_in_sources [of \] trg_in_targets [of \] by auto show "sources \ \ targets \ \ {} \ src \ = trg \" proof - assume 1: "sources \ \ targets \ \ {}" obtain a where a: "a \ sources \ \ targets \" using assms 1 by blast have \: "arr \" using a composable_implies_arr by auto have \: "arr \" using a composable_implies_arr by auto have 1: "\a'. a' \ sources \ \ src a' = src a \ trg a' = trg a" proof fix a' assume a': "a' \ sources \" have 1: "a' \ a" using a a' \ src_dom sources_dom source_iso_src isomorphic_transitive isomorphic_symmetric by (meson IntD1) obtain \ where \: "iso \ \ \ \ hom a' a" using 1 by auto show "src a' = src a" using \ src_dom src_cod by auto show "trg a' = trg a" using \ trg_dom trg_cod by auto qed have 2: "\a'. a' \ targets \ \ src a' = src a \ trg a' = trg a" proof fix a' assume a': "a' \ targets \" have 1: "a' \ a" using a a' \ trg_dom targets_dom target_iso_trg isomorphic_transitive isomorphic_symmetric by (meson IntD2) obtain \ where \: "iso \ \ \ \ hom a' a" using 1 by auto show "src a' = src a" using \ src_dom src_cod by auto show "trg a' = trg a" using \ trg_dom trg_cod by auto qed have "src \ = src (src \)" using \ by simp also have "... = src (trg \)" using \ 1 [of "src \"] src_in_sources a weak_unit_self_composable seq_if_composable by auto also have "... = trg (trg \)" using \ by simp also have "... = trg \" using \ by simp finally show "src \ = trg \" by blast qed qed lemma is_associative_weak_composition: shows "associative_weak_composition V H \" proof - have 1: "\\ \. \ \ \ \ null \ src \ = trg \" using H.is_extensional VV.arr_char by force show "associative_weak_composition V H \" proof show "\f g h. ide f \ ide g \ ide h \ f \ g \ null \ g \ h \ null \ \\[f, g, h] : (f \ g) \ h \ f \ g \ h\" using 1 by auto show "\f g h. ide f \ ide g \ ide h \ f \ g \ null \ g \ h \ null \ iso \[f, g, h]" using 1 iso_assoc by presburger show "\\ \ \. \ \ \ \ null \ \ \ \ \ null \ \[cod \, cod \, cod \] \ ((\ \ \) \ \) = (\ \ \ \ \) \ \[dom \, dom \, dom \]" using 1 assoc_naturality hseq_char hseq_char' by metis show "\f g h k. ide f \ ide g \ ide h \ ide k \ sources f \ targets g \ {} \ sources g \ targets h \ {} \ sources h \ targets k \ {} \ (f \ \[g, h, k]) \ \[f, g \ h, k] \ (\[f, g, h] \ k) = \[f, g, h \ k] \ \[f \ g, h, k]" using 1 connected_iff_seq pentagon ideD(1) by auto qed qed interpretation associative_weak_composition V H \ using is_associative_weak_composition by auto theorem is_prebicategory: shows "prebicategory V H \" using src_in_sources trg_in_targets by (unfold_locales, auto) interpretation prebicategory V H \ using is_prebicategory by auto corollary is_prebicategory_with_homs: shows "prebicategory_with_homs V H \ src trg" .. interpretation prebicategory_with_homs V H \ src trg using is_prebicategory_with_homs by auto text \ In a bicategory, an arrow is a weak unit if and only if it is isomorphic to its chosen source (or to its chosen target). \ lemma weak_unit_char: shows "weak_unit a \ a \ src a" and "weak_unit a \ a \ trg a" proof - show "weak_unit a \ a \ src a" using isomorphism_respects_weak_units isomorphic_symmetric by (meson ideD(1) isomorphic_implies_ide(2) obj_is_weak_unit obj_src source_iso_src weak_unit_iff_self_source weak_unit_self_composable(1)) show "weak_unit a \ a \ trg a" using isomorphism_respects_weak_units isomorphic_symmetric by (metis \weak_unit a = isomorphic a (src a)\ ideD(1) isomorphic_implies_hpar(3) isomorphic_implies_ide(1) src_trg target_iso_trg weak_unit_iff_self_target) qed interpretation H: partial_magma H using is_partial_magma by auto text \ Every arrow with respect to horizontal composition is also an arrow with respect to vertical composition. The converse is not necessarily true. \ lemma harr_is_varr: assumes "H.arr \" shows "arr \" proof - have "H.domains \ \ {} \ arr \" proof - assume 1: "H.domains \ \ {}" obtain a where a: "H.ide a \ \ \ a \ null" using 1 H.domains_def by auto show "arr \" using a hseq_char' H.ide_def by blast qed moreover have "H.codomains \ \ {} \ arr \" proof - assume 1: "H.codomains \ \ {}" obtain a where a: "H.ide a \ a \ \ \ null" using 1 H.codomains_def by auto show "arr \" using a hseq_char' ide_def by blast qed ultimately show ?thesis using assms H.arr_def by auto qed text \ An identity for horizontal composition is also an identity for vertical composition. \ lemma horizontal_identity_is_ide: assumes "H.ide \" shows "ide \" proof - have \: "arr \" using assms H.ide_def composable_implies_arr(2) by auto hence 1: "\ \ dom \ \ null" using assms hom_connected H.ide_def by auto have "\ \ dom \ = dom \" using assms 1 H.ide_def by simp moreover have "\ \ dom \ = \" using assms 1 H.ide_def [of \] null_agreement by (metis \ cod_cod cod_dom hcomp_simps\<^sub>W\<^sub>C(3) ideD(2) ide_char' paste_1) ultimately have "dom \ = \" by simp thus ?thesis using \ by (metis ide_dom) qed text \ Every identity for horizontal composition is a weak unit. \ lemma horizontal_identity_is_weak_unit: assumes "H.ide \" shows "weak_unit \" using assms weak_unit_char by (metis H.ide_def comp_target_ide horizontal_identity_is_ide ideD(1) isomorphism_respects_weak_units null_agreement targetsD(2-3) trg_in_targets) end subsection "Vertically Discrete Bicategories are Categories" text \ In this section we show that if a bicategory is discrete with respect to vertical composition, then it is a category with respect to horizontal composition. To obtain this result, we need to establish that the set of arrows for the horizontal composition coincides with the set of arrows for the vertical composition. This is not true for a general bicategory, and even with the assumption that the vertical category is discrete it is not immediately obvious from the definitions. The issue is that the notion ``arrow'' for the horizontal composition is defined in terms of the existence of ``domains'' and ``codomains'' with respect to that composition, whereas the axioms for a bicategory only relate the notion ``arrow'' for the vertical category to the existence of sources and targets with respect to the horizontal composition. So we have to establish that, under the assumption of vertical discreteness, sources coincide with domains and targets coincide with codomains. We also need the fact that horizontal identities are weak units, which previously required some effort to show. \ locale vertically_discrete_bicategory = bicategory + assumes vertically_discrete: "ide = arr" begin interpretation prebicategory_with_homs V H \ src trg using is_prebicategory_with_homs by auto interpretation H: partial_magma H using is_partial_magma(1) by auto lemma weak_unit_is_horizontal_identity: assumes "weak_unit a" shows "H.ide a" proof - have "a \ a \ H.null" using assms weak_unit_self_composable by simp moreover have "\f. f \ a \ H.null \ f \ a = f" proof - fix f assume "f \ a \ H.null" hence "f \ a \ f" using assms comp_ide_source composable_implies_arr(2) sourcesI vertically_discrete by auto thus "f \ a = f" using vertically_discrete isomorphic_def by auto qed moreover have "\f. a \ f \ H.null \ a \ f = f" proof - fix f assume "a \ f \ H.null" hence "a \ f \ f" using assms comp_target_ide composable_implies_arr(1) targetsI vertically_discrete by auto thus "a \ f = f" using vertically_discrete isomorphic_def by auto qed ultimately show "H.ide a" using H.ide_def by simp qed lemma sources_eq_domains: shows "sources \ = H.domains \" using weak_unit_is_horizontal_identity H.domains_def sources_def horizontal_identity_is_weak_unit by auto lemma targets_eq_codomains: shows "targets \ = H.codomains \" using weak_unit_is_horizontal_identity H.codomains_def targets_def horizontal_identity_is_weak_unit by auto lemma arr_agreement: shows "arr = H.arr" using arr_def H.arr_def arr_iff_has_src arr_iff_has_trg sources_eq_domains targets_eq_codomains by auto interpretation H: category H proof show "\g f. g \ f \ H.null \ H.seq g f" using arr_agreement hcomp_simps\<^sub>W\<^sub>C(1) by auto show "\f. (H.domains f \ {}) = (H.codomains f \ {})" using sources_eq_domains targets_eq_codomains arr_iff_has_src arr_iff_has_trg by simp fix f g h show "H.seq h g \ H.seq (h \ g) f \ H.seq g f" using null_agreement arr_agreement H.not_arr_null preserves_arr VoV.arr_char by (metis hseq_char' match_1) show "H.seq h (g \ f) \ H.seq g f \ H.seq h g" using null_agreement arr_agreement H.not_arr_null preserves_arr VoV.arr_char by (metis hseq_char' match_2) show "H.seq g f \ H.seq h g \ H.seq (h \ g) f" using arr_agreement match_3 hseq_char(1) by auto show "H.seq g f \ H.seq h g \ (h \ g) \ f = h \ g \ f" proof - assume hg: "H.seq h g" assume gf: "H.seq g f" have "iso \[h, g, f] \ \\[h, g, f] : (h \ g) \ f \ h \ g \ f\" using hg gf vertically_discrete arr_agreement hseq_char assoc_in_hom iso_assoc by auto thus ?thesis using arr_agreement vertically_discrete by auto qed qed proposition is_category: shows "category H" .. end subsection "Obtaining the Unitors" text \ We now want to exploit the construction of unitors in a prebicategory with units, to obtain left and right unitors in a bicategory. However, a bicategory is not \emph{a priori} a prebicategory with units, because a bicategory only assigns unit isomorphisms to each \emph{object}, not to each weak unit. In order to apply the results about prebicategories with units to a bicategory, we first need to extend the bicategory to a prebicategory with units, by extending the mapping \\\, which provides a unit isomorphism for each object, to a mapping that assigns a unit isomorphism to all weak units. This extension can be made in an arbitrary way, as the values chosen for non-objects ultimately do not affect the components of the unitors at objects. \ context bicategory begin interpretation prebicategory V H \ using is_prebicategory by auto definition \' where "\' a \ SOME \. iso \ \ \ \ hom (a \ a) a \ (obj a \ \ = \[a])" lemma \'_extends_\: assumes "weak_unit a" shows "iso (\' a)" and "\\' a : a \ a \ a\" and "obj a \ \' a = \[a]" proof - let ?P = "\a \. iso \ \ \\ : a \ a \ a\ \ (obj a \ \ = \[a])" have "\\. ?P a \" - using assms unit_in_hom iso_unit weak_unit_def isomorphic_def by blast + by (metis assms iso_some_unit(1) iso_some_unit(2) iso_unit unit_in_vhom) hence 1: "?P a (\' a)" using \'_def someI_ex [of "?P a"] by simp show "iso (\' a)" using 1 by simp show "\\' a : a \ a \ a\" using 1 by simp show "obj a \ \' a = \[a]" using 1 by simp qed proposition extends_to_prebicategory_with_units: shows "prebicategory_with_units V H \ \'" using \'_extends_\ by unfold_locales auto interpretation PB: prebicategory_with_units V H \ \' using extends_to_prebicategory_with_units by auto interpretation PB: prebicategory_with_homs V H \ src trg using is_prebicategory_with_homs by auto interpretation PB: prebicategory_with_homs_and_units V H \ \' src trg .. proposition extends_to_prebicategory_with_homs_and_units: shows "prebicategory_with_homs_and_units V H \ \' src trg" .. definition lunit ("\[_]") where "\[a] \ PB.lunit a" definition runit ("\[_]") where "\[a] \ PB.runit a" abbreviation lunit' ("\\<^sup>-\<^sup>1[_]") where "\\<^sup>-\<^sup>1[a] \ inv \[a]" abbreviation runit' ("\\<^sup>-\<^sup>1[_]") where "\\<^sup>-\<^sup>1[a] \ inv \[a]" text \ \sloppypar The characterizations of the left and right unitors that we obtain from locale @{locale prebicategory_with_homs_and_units} mention the arbitarily chosen extension \\'\, rather than the given \\\. We want ``native versions'' for the present context. \ lemma lunit_char: assumes "ide f" shows "\\[f] : L f \ f\" and "L \[f] = (\[trg f] \ f) \ \\<^sup>-\<^sup>1[trg f, trg f, f]" and "\!\. \\ : L f \ f\ \ L \ = (\[trg f] \ f) \ \\<^sup>-\<^sup>1[trg f, trg f, f]" proof - have 1: "trg (PB.lunit f) = trg f" using assms PB.lunit_char [of f] vconn_implies_hpar(2) vconn_implies_hpar(4) by metis show "\\[f] : L f \ f\" unfolding lunit_def using assms PB.lunit_char by simp show "L \[f] = (\[trg f] \ f) \ \\<^sup>-\<^sup>1[trg f, trg f, f]" unfolding lunit_def using assms 1 PB.lunit_char obj_is_weak_unit \'_extends_\ by simp let ?P = "\\. \\ : L f \ f\ \ L \ = (\[trg f] \ f) \ \\<^sup>-\<^sup>1[trg f, trg f, f]" have "?P = (\\. \\ : trg f \ f \ f\ \ trg f \ \ = (\' (trg f) \ f) \ inv \[trg f, trg f, f])" proof - have "\\. \\ : L f \ f\ \ \\ : trg f \ f \ f\" using assms by simp moreover have "\\. \\ : L f \ f\ \ L \ = (\[trg f] \ f) \ \\<^sup>-\<^sup>1[trg f, trg f, f] \ trg f \ \ = (\' (trg f) \ f) \ inv \[trg f, trg f, f]" using calculation obj_is_weak_unit \'_extends_\ by auto ultimately show ?thesis by blast qed thus "\!\. \\ : L f \ f\ \ L \ = (\[trg f] \ f) \ \\<^sup>-\<^sup>1[trg f, trg f, f]" using assms PB.lunit_char by simp qed lemma lunit_in_hom [intro]: assumes "ide f" shows "\\[f] : src f \ trg f\" and "\\[f] : trg f \ f \ f\" proof - show "\\[f] : trg f \ f \ f\" using assms lunit_char by auto thus "\\[f] : src f \ trg f\" by (metis arrI in_hhomI vconn_implies_hpar(1-4)) qed lemma lunit_in_vhom [simp]: assumes "ide f" and "trg f = b" shows "\\[f] : b \ f \ f\" using assms by auto lemma lunit_simps [simp]: assumes "ide f" shows "arr \[f]" and "src \[f] = src f" and "trg \[f] = trg f" and "dom \[f] = trg f \ f" and "cod \[f] = f" using assms lunit_in_hom apply auto using assms lunit_in_hom apply blast using assms lunit_in_hom by blast lemma runit_char: assumes "ide f" shows "\\[f] : R f \ f\" and "R \[f] = (f \ \[src f]) \ \[f, src f, src f]" and "\!\. \\ : R f \ f\ \ R \ = (f \ \[src f]) \ \[f, src f, src f]" proof - have 1: "src (PB.runit f) = src f" using assms PB.runit_char [of f] vconn_implies_hpar(1) vconn_implies_hpar(3) by metis show "\\[f] : R f \ f\" unfolding runit_def using assms PB.runit_char by simp show "R \[f] = (f \ \[src f]) \ \[f, src f, src f]" unfolding runit_def using assms 1 PB.runit_char obj_is_weak_unit \'_extends_\ by simp let ?P = "\\. \\ : R f \ f\ \ R \ = (f \ \[src f]) \ \[f, src f, src f]" have "?P = (\\. \\ : f \ src f \ f\ \ \ \ src f = (f \ \' (src f)) \ \[f, src f, src f])" proof - have "\\. \\ : R f \ f\ \ \\ : f \ src f \ f\" using assms by simp moreover have "\\. \\ : R f \ f\ \ R \ = (f \ \[src f]) \ \[f, src f, src f] \ \ \ src f = (f \ \' (src f)) \ \[f, src f, src f]" using calculation obj_is_weak_unit \'_extends_\ by auto ultimately show ?thesis by blast qed thus "\!\. \\ : R f \ f\ \ R \ = (f \ \[src f]) \ \[f, src f, src f]" using assms PB.runit_char by simp qed lemma runit_in_hom [intro]: assumes "ide f" shows "\\[f] : src f \ trg f\" and "\\[f] : f \ src f \ f\" proof - show "\\[f] : f \ src f \ f\" using assms runit_char by auto thus "\\[f] : src f \ trg f\" by (metis arrI in_hhom_def vconn_implies_hpar(1-4)) qed lemma runit_in_vhom [simp]: assumes "ide f" and "src f = a" shows "\\[f] : f \ a \ f\" using assms by auto lemma runit_simps [simp]: assumes "ide f" shows "arr \[f]" and "src \[f] = src f" and "trg \[f] = trg f" and "dom \[f] = f \ src f" and "cod \[f] = f" using assms runit_in_hom apply auto using assms runit_in_hom apply blast using assms runit_in_hom by blast lemma lunit_eqI: assumes "ide f" and "\\ : trg f \ f \ f\" and "trg f \ \ = (\[trg f] \ f) \ \\<^sup>-\<^sup>1[trg f, trg f, f]" shows "\ = \[f]" unfolding lunit_def using assms PB.lunit_eqI \'_extends_\ trg.preserves_ide obj_is_weak_unit by simp lemma runit_eqI: assumes "ide f" and "\\ : f \ src f \ f\" and "\ \ src f = (f \ \[src f]) \ \[f, src f, src f]" shows "\ = \[f]" unfolding runit_def using assms PB.runit_eqI \'_extends_\ src.preserves_ide obj_is_weak_unit by simp lemma lunit_naturality: assumes "arr \" shows "\ \ \[dom \] = \[cod \] \ (trg \ \ \)" unfolding lunit_def using assms PB.lunit_naturality by auto lemma runit_naturality: assumes "arr \" shows "\ \ \[dom \] = \[cod \] \ (\ \ src \)" unfolding runit_def using assms PB.runit_naturality by auto lemma iso_lunit [simp]: assumes "ide f" shows "iso \[f]" unfolding lunit_def using assms PB.iso_lunit by blast lemma iso_runit [simp]: assumes "ide f" shows "iso \[f]" unfolding runit_def using assms PB.iso_runit by blast lemma iso_lunit' [simp]: assumes "ide f" shows "iso \\<^sup>-\<^sup>1[f]" using assms iso_lunit by blast lemma iso_runit' [simp]: assumes "ide f" shows "iso \\<^sup>-\<^sup>1[f]" using assms iso_runit by blast lemma lunit'_in_hom [intro]: assumes "ide f" shows "\\\<^sup>-\<^sup>1[f] : src f \ trg f\" and "\\\<^sup>-\<^sup>1[f] : f \ trg f \ f\" proof - show "\\\<^sup>-\<^sup>1[f] : f \ trg f \ f\" using assms lunit_char iso_lunit by simp thus "\\\<^sup>-\<^sup>1[f] : src f \ trg f\" using assms src_dom trg_dom by simp qed lemma lunit'_in_vhom [simp]: assumes "ide f" and "trg f = b" shows "\\\<^sup>-\<^sup>1[f] : f \ b \ f\" using assms by auto lemma lunit'_simps [simp]: assumes "ide f" shows "arr \\<^sup>-\<^sup>1[f]" and "src \\<^sup>-\<^sup>1[f] = src f" and "trg \\<^sup>-\<^sup>1[f] = trg f" and "dom \\<^sup>-\<^sup>1[f] = f" and "cod \\<^sup>-\<^sup>1[f] = trg f \ f" using assms lunit'_in_hom by auto lemma runit'_in_hom [intro]: assumes "ide f" shows "\\\<^sup>-\<^sup>1[f] : src f \ trg f\" and "\\\<^sup>-\<^sup>1[f] : f \ f \ src f\" proof - show "\\\<^sup>-\<^sup>1[f] : f \ f \ src f\" using assms runit_char iso_runit by simp thus "\\\<^sup>-\<^sup>1[f] : src f \ trg f\" using src_dom trg_dom by (simp add: assms) qed lemma runit'_in_vhom [simp]: assumes "ide f" and "src f = a" shows "\\\<^sup>-\<^sup>1[f] : f \ f \ a\" using assms by auto lemma runit'_simps [simp]: assumes "ide f" shows "arr \\<^sup>-\<^sup>1[f]" and "src \\<^sup>-\<^sup>1[f] = src f" and "trg \\<^sup>-\<^sup>1[f] = trg f" and "dom \\<^sup>-\<^sup>1[f] = f" and "cod \\<^sup>-\<^sup>1[f] = f \ src f" using assms runit'_in_hom by auto interpretation L: endofunctor V L .. interpretation \: transformation_by_components V V L map lunit using lunit_in_hom lunit_naturality by unfold_locales auto interpretation \: natural_isomorphism V V L map \.map using iso_lunit by (unfold_locales, auto) lemma natural_isomorphism_\: shows "natural_isomorphism V V L map \.map" .. abbreviation \ where "\ \ \.map" lemma \_ide_simp: assumes "ide f" shows "\ f = \[f]" using assms by simp interpretation L: equivalence_functor V V L using L.isomorphic_to_identity_is_equivalence \.natural_isomorphism_axioms by simp lemma equivalence_functor_L: shows "equivalence_functor V V L" .. lemma lunit_commutes_with_L: assumes "ide f" shows "\[L f] = L \[f]" unfolding lunit_def using assms PB.lunit_commutes_with_L by blast interpretation R: endofunctor V R .. interpretation \: transformation_by_components V V R map runit using runit_in_hom runit_naturality by unfold_locales auto interpretation \: natural_isomorphism V V R map \.map using iso_runit by (unfold_locales, auto) lemma natural_isomorphism_\: shows "natural_isomorphism V V R map \.map" .. abbreviation \ where "\ \ \.map" lemma \_ide_simp: assumes "ide f" shows "\ f = \[f]" using assms by simp interpretation R: equivalence_functor V V R using R.isomorphic_to_identity_is_equivalence \.natural_isomorphism_axioms by simp lemma equivalence_functor_R: shows "equivalence_functor V V R" .. lemma runit_commutes_with_R: assumes "ide f" shows "\[R f] = R \[f]" unfolding runit_def using assms PB.runit_commutes_with_R by blast lemma lunit'_naturality: assumes "arr \" shows "(trg \ \ \) \ \\<^sup>-\<^sup>1[dom \] = \\<^sup>-\<^sup>1[cod \] \ \" using assms iso_lunit lunit_naturality invert_opposite_sides_of_square L.preserves_arr L.preserves_cod arr_cod ide_cod ide_dom lunit_simps(1) lunit_simps(4) seqI by presburger lemma runit'_naturality: assumes "arr \" shows "(\ \ src \) \ \\<^sup>-\<^sup>1[dom \] = \\<^sup>-\<^sup>1[cod \] \ \" using assms iso_runit runit_naturality invert_opposite_sides_of_square R.preserves_arr R.preserves_cod arr_cod ide_cod ide_dom runit_simps(1) runit_simps(4) seqI by presburger lemma isomorphic_unit_right: assumes "ide f" shows "f \ src f \ f" using assms runit'_in_hom iso_runit' isomorphic_def isomorphic_symmetric by blast lemma isomorphic_unit_left: assumes "ide f" shows "trg f \ f \ f" using assms lunit'_in_hom iso_lunit' isomorphic_def isomorphic_symmetric by blast end subsection "Further Properties of Bicategories" text \ Here we derive further properties of bicategories, now that we have the unitors at our disposal. This section generalizes the corresponding development in theory @{theory MonoidalCategory.MonoidalCategory}, which has some diagrams to illustrate the longer calculations. The present section also includes some additional facts that are now nontrivial due to the partiality of horizontal composition. \ context bicategory begin lemma unit_simps [simp]: assumes "obj a" shows "arr \[a]" and "src \[a] = a" and "trg \[a] = a" and "dom \[a] = a \ a" and "cod \[a] = a" using assms unit_in_hom by blast+ lemma triangle: assumes "ide f" and "ide g" and "src g = trg f" shows "(g \ \[f]) \ \[g, src g, f] = \[g] \ f" proof - let ?b = "src g" have *: "(g \ \[?b \ f]) \ \[g, ?b, ?b \ f] = \[g] \ ?b \ f" proof - have 1: "((g \ \[?b \ f]) \ \[g, ?b, ?b \ f]) \ \[g \ ?b, ?b, f] = (\[g] \ ?b \ f) \ \[g \ ?b, ?b, f]" proof - have "((g \ \[?b \ f]) \ \[g, ?b, ?b \ f]) \ \[g \ ?b, ?b, f] = (g \ \[?b \ f]) \ \[g, ?b, ?b \ f] \ \[g \ ?b, ?b, f]" using HoVH_def HoHV_def comp_assoc by auto also have "... = (g \ \[?b \ f]) \ (g \ \[?b, ?b, f]) \ \[g, ?b \ ?b, f] \ (\[g, ?b, ?b] \ f)" using assms pentagon by force also have "... = ((g \ \[?b \ f]) \ (g \ \[?b, ?b, f])) \ \[g, ?b \ ?b, f] \ (\[g, ?b, ?b] \ f)" using assms assoc_in_hom HoVH_def HoHV_def comp_assoc by auto also have "... = ((g \ ?b \ \[f]) \ (g \ \[?b, ?b, f])) \ \[g, ?b \ ?b, f] \ (\[g, ?b, ?b] \ f)" using assms lunit_commutes_with_L lunit_in_hom by force also have "... = ((g \ (\[?b] \ f) \ \\<^sup>-\<^sup>1[?b, ?b, f]) \ (g \ \[?b, ?b, f])) \ \[g, ?b \ ?b, f] \ (\[g, ?b, ?b] \ f)" using assms lunit_char(2) by force also have "... = (g \ ((\[?b] \ f) \ \\<^sup>-\<^sup>1[?b, ?b, f]) \ \[?b, ?b, f]) \ \[g, ?b \ ?b, f] \ (\[g, ?b, ?b] \ f)" using assms interchange [of g g "(\[?b] \ f) \ \\<^sup>-\<^sup>1[?b, ?b, f]" "\[?b, ?b, f]"] by auto also have "... = ((g \ \[?b] \ f) \ \[g, ?b \ ?b, f]) \ (\[g, ?b, ?b] \ f)" using assms comp_arr_dom comp_assoc_assoc' comp_assoc by auto also have "... = (\[g, ?b, f] \ ((g \ \[?b]) \ f)) \ (\[g, ?b, ?b] \ f)" using assms assoc_naturality [of g "\[?b]" f] by simp also have "... = \[g, ?b, f] \ ((g \ \[?b]) \ \[g, ?b, ?b] \ f)" using assms interchange [of "g \ \[?b]" "\[g, ?b, ?b]" f f] comp_assoc by simp also have "... = \[g, ?b, f] \ ((\[g] \ ?b) \ f)" using assms runit_char(2) by force also have "... = (\[g] \ ?b \ f) \ \[g \ ?b, ?b, f]" using assms assoc_naturality [of "\[g]" ?b f] by auto finally show ?thesis by blast qed show "(g \ \[?b \ f]) \ \[g, ?b, ?b \ f] = \[g] \ ?b \ f" proof - have "epi \[g \ ?b, ?b, f]" using assms preserves_ide iso_assoc iso_is_retraction retraction_is_epi by force thus ?thesis using assms 1 by auto qed qed have "(g \ \[f]) \ \[g, ?b, f] = ((g \ \[f]) \ (g \ \[?b \ f]) \ (g \ ?b \ \\<^sup>-\<^sup>1[f])) \ (g \ ?b \ \[f]) \ \[g, ?b, ?b \ f] \ ((g \ ?b) \ \\<^sup>-\<^sup>1[f])" proof - have "\[g, ?b, f] = (g \ ?b \ \[f]) \ \[g, ?b, ?b \ f] \ ((g \ ?b) \ \\<^sup>-\<^sup>1[f])" proof - have "\[g, ?b, f] = (g \ ?b \ f) \ \[g, ?b, f]" using assms comp_cod_arr by simp have "\[g, ?b, f] = ((g \ ?b \ \[f]) \ (g \ ?b \ \\<^sup>-\<^sup>1[f])) \ \[g, ?b, f]" using assms comp_cod_arr comp_arr_inv' whisker_left [of g] whisker_left [of ?b "\[f]" "\\<^sup>-\<^sup>1[f]"] by simp also have "... = (g \ ?b \ \[f]) \ \[g, ?b, ?b \ f] \ ((g \ ?b) \ \\<^sup>-\<^sup>1[f])" using assms iso_lunit assoc_naturality [of g ?b "\\<^sup>-\<^sup>1[f]"] comp_assoc by force finally show ?thesis by blast qed moreover have "g \ \[f] = (g \ \[f]) \ (g \ \[?b \ f]) \ (g \ ?b \ \\<^sup>-\<^sup>1[f])" proof - have "(g \ \[?b \ f]) \ (g \ ?b \ \\<^sup>-\<^sup>1[f]) = g \ ?b \ f" proof - have "(g \ \[?b \ f]) \ (g \ ?b \ \\<^sup>-\<^sup>1[f]) = (g \ ?b \ \[f]) \ (g \ ?b \ \\<^sup>-\<^sup>1[f])" using assms lunit_in_hom lunit_commutes_with_L by simp also have "... = g \ ?b \ f" using assms comp_arr_inv' whisker_left [of g] whisker_left [of ?b "\[f]" "\\<^sup>-\<^sup>1[f]"] by simp finally show ?thesis by blast qed thus ?thesis using assms comp_arr_dom by auto qed ultimately show ?thesis by simp qed also have "... = (g \ \[f]) \ (g \ \[?b \ f]) \ ((g \ ?b \ \\<^sup>-\<^sup>1[f]) \ (g \ ?b \ \[f])) \ \[g, ?b, ?b \ f] \ ((g \ ?b) \ \\<^sup>-\<^sup>1[f])" using comp_assoc by simp also have "... = (g \ \[f]) \ (g \ \[?b \ f]) \ ((g \ ?b \ (?b \ f)) \ \[g, ?b, ?b \ f]) \ ((g \ ?b) \ \\<^sup>-\<^sup>1[f])" using assms iso_lunit comp_inv_arr' interchange [of g g "?b \ \\<^sup>-\<^sup>1[f]" "?b \ \[f]"] interchange [of ?b ?b "\\<^sup>-\<^sup>1[f]" "\[f]"] comp_assoc by auto also have "... = (g \ \[f]) \ ((g \ \[?b \ f]) \ \[g, ?b, ?b \ f]) \ ((g \ ?b) \ \\<^sup>-\<^sup>1[f])" using assms comp_cod_arr comp_assoc by auto also have "... = \[g] \ f" proof - have "\[g] \ f = (g \ \[f]) \ (\[g] \ ?b \ f) \ ((g \ ?b) \ \\<^sup>-\<^sup>1[f])" proof - have "(g \ \[f]) \ (\[g] \ ?b \ f) \ ((g \ ?b) \ \\<^sup>-\<^sup>1[f]) = (g \ \[f]) \ (\[g] \ (g \ ?b) \ (?b \ f) \ \\<^sup>-\<^sup>1[f])" using assms iso_lunit interchange [of "\[g]" "g \ ?b" "?b \ f" "\\<^sup>-\<^sup>1[f]"] by force also have "... = (g \ \[f]) \ (\[g] \ \\<^sup>-\<^sup>1[f])" using assms comp_arr_dom comp_cod_arr by simp also have "... = \[g] \ \[f] \ \\<^sup>-\<^sup>1[f]" using assms interchange [of g "\[g]" "\[f]" "\\<^sup>-\<^sup>1[f]"] comp_cod_arr by simp also have "... = \[g] \ f" using assms iso_lunit comp_arr_inv' by simp finally show ?thesis by argo qed thus ?thesis using assms * by argo qed finally show ?thesis by blast qed lemma lunit_hcomp_gen: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "(f \ \[g \ h]) \ (f \ \[trg g, g, h]) = f \ \[g] \ h" proof - have "((f \ \[g \ h]) \ (f \ \[trg g, g, h])) \ \[f, trg g \ g, h] \ (\[f, trg g, g] \ h) = (f \ (\[g] \ h)) \ \[f, trg g \ g, h] \ (\[f, trg g, g] \ h)" proof - have "((f \ \[g \ h]) \ (f \ \[trg g, g, h])) \ (\[f, trg g \ g, h] \ (\[f, trg g, g] \ h)) = ((f \ \[g \ h]) \ \[f, trg g, g \ h]) \ \[f \ trg g, g, h]" using assms pentagon comp_assoc by simp also have "... = (\[f] \ (g \ h)) \ \[f \ trg g, g, h]" using assms triangle [of "g \ h" f] by auto also have "... = \[f, g, h] \ ((\[f] \ g) \ h)" using assms assoc_naturality [of "\[f]" g h] by simp also have "... = (\[f, g, h] \ ((f \ \[g]) \ h)) \ (\[f, trg g, g] \ h)" using assms triangle interchange [of "f \ \[g]" "\[f, trg g, g]" h h] comp_assoc by auto also have "... = (f \ (\[g] \ h)) \ (\[f, trg g \ g, h] \ (\[f, trg g, g] \ h))" using assms assoc_naturality [of f "\[g]" h] comp_assoc by simp finally show ?thesis by blast qed moreover have "iso (\[f, trg g \ g, h] \ (\[f, trg g, g] \ h))" using assms iso_assoc isos_compose by simp ultimately show ?thesis using assms iso_is_retraction retraction_is_epi epiE [of "\[f, trg g \ g, h] \ (\[f, trg g, g] \ h)" "(f \ \[g \ h]) \ (f \ \[trg g, g, h])" "f \ \[g] \ h"] by auto qed lemma lunit_hcomp: assumes "ide f" and "ide g" and "src f = trg g" shows "\[f \ g] \ \[trg f, f, g] = \[f] \ g" and "\\<^sup>-\<^sup>1[trg f, f, g] \ \\<^sup>-\<^sup>1[f \ g] = \\<^sup>-\<^sup>1[f] \ g" and "\[f \ g] = (\[f] \ g) \ \\<^sup>-\<^sup>1[trg f, f, g]" and "\\<^sup>-\<^sup>1[f \ g] = \[trg f, f, g] \ (\\<^sup>-\<^sup>1[f] \ g)" proof - show 1: "\[f \ g] \ \[trg f, f, g] = \[f] \ g" proof - have "L (\[f \ g] \ \[trg f, f, g]) = L (\[f] \ g)" using assms interchange [of "trg f" "trg f" "\[f \ g]" "\[trg f, f, g]"] lunit_hcomp_gen by fastforce thus ?thesis using assms L.is_faithful [of "\[f \ g] \ \[trg f, f, g]" "\[f] \ g"] by force qed show "\\<^sup>-\<^sup>1[trg f, f, g] \ \\<^sup>-\<^sup>1[f \ g] = \\<^sup>-\<^sup>1[f] \ g" proof - have "\\<^sup>-\<^sup>1[trg f, f, g] \ \\<^sup>-\<^sup>1[f \ g] = inv (\[f \ g] \ \[trg f, f, g])" using assms by (simp add: inv_comp) also have "... = inv (\[f] \ g)" using 1 by simp also have "... = \\<^sup>-\<^sup>1[f] \ g" using assms by simp finally show ?thesis by simp qed show 2: "\[f \ g] = (\[f] \ g) \ \\<^sup>-\<^sup>1[trg f, f, g]" using assms 1 invert_side_of_triangle(2) by auto show "\\<^sup>-\<^sup>1[f \ g] = \[trg f, f, g] \ (\\<^sup>-\<^sup>1[f] \ g)" proof - have "\\<^sup>-\<^sup>1[f \ g] = inv ((\[f] \ g) \ \\<^sup>-\<^sup>1[trg f, f, g])" using 2 by simp also have "... = \[trg f, f, g] \ inv (\[f] \ g)" using assms inv_comp by simp also have "... = \[trg f, f, g] \ (\\<^sup>-\<^sup>1[f] \ g)" using assms by simp finally show ?thesis by simp qed qed lemma runit_hcomp_gen: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "\[f \ g] \ h = ((f \ \[g]) \ h) \ (\[f, g, src g] \ h)" proof - have "\[f \ g] \ h = ((f \ g) \ \[h]) \ \[f \ g, src g, h]" using assms triangle by simp also have "... = (\\<^sup>-\<^sup>1[f, g, h] \ (f \ g \ \[h]) \ \[f, g, src g \ h]) \ \[f \ g, src g, h]" using assms assoc_naturality [of f g "\[h]"] invert_side_of_triangle(1) by simp also have "... = \\<^sup>-\<^sup>1[f, g, h] \ (f \ g \ \[h]) \ \[f, g, src g \ h] \ \[f \ g, src g, h]" using comp_assoc by simp also have "... = (\\<^sup>-\<^sup>1[f, g, h] \ (f \ (\[g] \ h))) \ (f \ \\<^sup>-\<^sup>1[g, src g, h]) \ \[f, g, src g \ h] \ \[f \ g, src g, h]" using assms interchange [of f f] triangle comp_assoc invert_side_of_triangle(2) [of "\[g] \ h" "g \ \[h]" "\[g, src g, h]"] by simp also have "... = ((f \ \[g]) \ h) \ \\<^sup>-\<^sup>1[f, g \ src g, h] \ (f \ \\<^sup>-\<^sup>1[g, src g, h]) \ \[f, g, src g \ h] \ \[f \ g, src g, h]" using assms assoc'_naturality [of f "\[g]" h] comp_assoc by simp also have "... = ((f \ \[g]) \ h) \ (\[f, g, src g] \ h)" using assms pentagon [of f g "src g" h] iso_assoc inv_hcomp invert_side_of_triangle(1) [of "\[f, g, src g \ h] \ \[f \ g, src g, h]" "f \ \[g, src g, h]" "\[f, g \ src g, h] \ (\[f, g, src g] \ h)"] invert_side_of_triangle(1) [of "(f \ \\<^sup>-\<^sup>1[g, src g, h]) \ \[f, g, src g \ h] \ \[f \ g, src g, h]" "\[f, g \ src g, h]" "\[f, g, src g] \ h"] by auto finally show ?thesis by blast qed lemma runit_hcomp: assumes "ide f" and "ide g" and "src f = trg g" shows "\[f \ g] = (f \ \[g]) \ \[f, g, src g]" and "\\<^sup>-\<^sup>1[f \ g] = \\<^sup>-\<^sup>1[f, g, src g] \ (f \ \\<^sup>-\<^sup>1[g])" and "\[f \ g] \ \\<^sup>-\<^sup>1[f, g, src g] = f \ \[g]" and "\[f, g, src g] \ \\<^sup>-\<^sup>1[f \ g] = f \ \\<^sup>-\<^sup>1[g]" proof - show 1: "\[f \ g] = (f \ \[g]) \ \[f, g, src g]" using assms interchange [of "f \ \[g]" "\[f, g, src g]" "src g" "src g"] runit_hcomp_gen [of f g "src g"] R.is_faithful [of "(f \ \[g]) \ (\[f, g, src g])" "\[f \ g]"] by simp show "\\<^sup>-\<^sup>1[f \ g] = \\<^sup>-\<^sup>1[f, g, src g] \ (f \ \\<^sup>-\<^sup>1[g])" using assms 1 inv_comp inv_hcomp by auto show 2: "\[f \ g] \ \\<^sup>-\<^sup>1[f, g, src g] = f \ \[g]" using assms 1 comp_arr_dom comp_cod_arr comp_assoc hseqI' comp_assoc_assoc' by auto show "\[f, g, src g] \ \\<^sup>-\<^sup>1[f \ g] = f \ \\<^sup>-\<^sup>1[g]" proof - have "\[f, g, src g] \ \\<^sup>-\<^sup>1[f \ g] = inv (\[f \ g] \ \\<^sup>-\<^sup>1[f, g, src g])" using assms inv_comp by simp also have "... = inv (f \ \[g])" using 2 by simp also have "... = f \ \\<^sup>-\<^sup>1[g]" using assms inv_hcomp [of f "\[g]"] by simp finally show ?thesis by simp qed qed lemma unitor_coincidence: assumes "obj a" shows "\[a] = \[a]" and "\[a] = \[a]" proof - have "R \[a] = R \[a] \ R \[a] = R \[a]" proof - have "R \[a] = (a \ \[a]) \ \[a, a, a]" using assms lunit_hcomp [of a a] lunit_commutes_with_L [of a] by auto moreover have "(a \ \[a]) \ \[a, a, a] = R \[a]" using assms triangle [of a a] by auto moreover have "(a \ \[a]) \ \[a, a, a] = R \[a]" proof - have "(a \ \[a]) \ \[a, a, a] = ((\[a] \ a) \ \\<^sup>-\<^sup>1[a, a, a]) \ \[a, a, a]" using assms lunit_char(2) by force also have "... = R \[a]" using assms comp_arr_dom comp_assoc comp_assoc_assoc' apply (elim objE) by (simp add: assms) finally show ?thesis by blast qed ultimately show ?thesis by argo qed moreover have "par \[a] \[a] \ par \[a] \[a]" using assms by auto ultimately have 1: "\[a] = \[a] \ \[a] = \[a]" using R.is_faithful by blast show "\[a] = \[a]" using 1 by auto show "\[a] = \[a]" using 1 by auto qed lemma unit_triangle: assumes "obj a" shows "\[a] \ a = (a \ \[a]) \ \[a, a, a]" and "(\[a] \ a) \ \\<^sup>-\<^sup>1[a, a, a] = a \ \[a]" proof - show 1: "\[a] \ a = (a \ \[a]) \ \[a, a, a]" using assms triangle [of a a] unitor_coincidence by auto show "(\[a] \ a) \ \\<^sup>-\<^sup>1[a, a, a] = a \ \[a]" using assms 1 invert_side_of_triangle(2) [of "\[a] \ a" "a \ \[a]" "\[a, a, a]"] assoc'_eq_inv_assoc by (metis hseqI' iso_assoc objE obj_def' unit_simps(1) unit_simps(2)) qed lemma hcomp_assoc_isomorphic: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "(f \ g) \ h \ f \ g \ h" using assms assoc_in_hom [of f g h] iso_assoc isomorphic_def by auto lemma hcomp_arr_obj: assumes "arr \" and "obj a" and "src \ = a" shows "\ \ a = \\<^sup>-\<^sup>1[cod \] \ \ \ \[dom \]" and "\[cod \] \ (\ \ a) \ \\<^sup>-\<^sup>1[dom \] = \" proof - show "\ \ a = \\<^sup>-\<^sup>1[cod \] \ \ \ \[dom \]" using assms iso_runit runit_naturality comp_cod_arr by (metis ide_cod ide_dom invert_side_of_triangle(1) runit_simps(1) runit_simps(5) seqI) show "\[cod \] \ (\ \ a) \ \\<^sup>-\<^sup>1[dom \] = \" using assms iso_runit runit_naturality [of \] comp_cod_arr by (metis ide_dom invert_side_of_triangle(2) comp_assoc runit_simps(1) runit_simps(5) seqI) qed lemma hcomp_obj_arr: assumes "arr \" and "obj b" and "b = trg \" shows "b \ \ = \\<^sup>-\<^sup>1[cod \] \ \ \ \[dom \]" and "\[cod \] \ (b \ \) \ \\<^sup>-\<^sup>1[dom \] = \" proof - show "b \ \ = \\<^sup>-\<^sup>1[cod \] \ \ \ \[dom \]" using assms iso_lunit lunit_naturality comp_cod_arr by (metis ide_cod ide_dom invert_side_of_triangle(1) lunit_simps(1) lunit_simps(5) seqI) show "\[cod \] \ (b \ \) \ \\<^sup>-\<^sup>1[dom \] = \" using assms iso_lunit lunit_naturality [of \] comp_cod_arr by (metis ide_dom invert_side_of_triangle(2) comp_assoc lunit_simps(1) lunit_simps(5) seqI) qed lemma hcomp_reassoc: assumes "arr \" and "arr \" and "arr \" and "src \ = trg \" and "src \ = trg \" shows "(\ \ \) \ \ = \\<^sup>-\<^sup>1[cod \, cod \, cod \] \ (\ \ \ \ \) \ \[dom \, dom \, dom \]" and "\ \ \ \ \ = \[cod \, cod \, cod \] \ ((\ \ \) \ \) \ \\<^sup>-\<^sup>1[dom \, dom \, dom \]" proof - show "(\ \ \) \ \ = \\<^sup>-\<^sup>1[cod \, cod \, cod \] \ (\ \ \ \ \) \ \[dom \, dom \, dom \]" proof - have "(\ \ \) \ \ = (\\<^sup>-\<^sup>1[cod \, cod \, cod \] \ \[cod \, cod \, cod \]) \ ((\ \ \) \ \)" using assms comp_assoc_assoc'(2) comp_cod_arr by simp also have "... = \\<^sup>-\<^sup>1[cod \, cod \, cod \] \ \[cod \, cod \, cod \] \ ((\ \ \) \ \)" using comp_assoc by simp also have "... = \\<^sup>-\<^sup>1[cod \, cod \, cod \] \ (\ \ \ \ \) \ \[dom \, dom \, dom \]" using assms assoc_naturality by simp finally show ?thesis by simp qed show "\ \ \ \ \ = \[cod \, cod \, cod \] \ ((\ \ \) \ \) \ \\<^sup>-\<^sup>1[dom \, dom \, dom \]" proof - have "\ \ \ \ \ = (\ \ \ \ \) \ \[dom \, dom \, dom \] \ \\<^sup>-\<^sup>1[dom \, dom \, dom \]" using assms comp_assoc_assoc'(1) comp_arr_dom by simp also have "... = ((\ \ \ \ \) \ \[dom \, dom \, dom \]) \ \\<^sup>-\<^sup>1[dom \, dom \, dom \]" using comp_assoc by simp also have "... = (\[cod \, cod \, cod \] \ ((\ \ \) \ \)) \ \\<^sup>-\<^sup>1[dom \, dom \, dom \]" using assms assoc_naturality by simp also have "... = \[cod \, cod \, cod \] \ ((\ \ \) \ \) \ \\<^sup>-\<^sup>1[dom \, dom \, dom \]" using comp_assoc by simp finally show ?thesis by simp qed qed lemma triangle': assumes "ide f" and "ide g" and "src f = trg g" shows "(f \ \[g]) = (\[f] \ g) \ \\<^sup>-\<^sup>1[f, src f, g]" proof - have "(\[f] \ g) \ \\<^sup>-\<^sup>1[f, src f, g] = ((f \ \[g]) \ \[f, src f, g]) \ \\<^sup>-\<^sup>1[f, src f, g]" using assms triangle by auto also have "... = (f \ \[g])" using assms comp_arr_dom comp_assoc comp_assoc_assoc' by auto finally show ?thesis by auto qed lemma pentagon': assumes "ide f" and "ide g" and "ide h" and "ide k" and "src f = trg g" and "src g = trg h" and "src h = trg k" shows "((\\<^sup>-\<^sup>1[f, g, h] \ k) \ \\<^sup>-\<^sup>1[f, g \ h, k]) \ (f \ \\<^sup>-\<^sup>1[g, h, k]) = \\<^sup>-\<^sup>1[f \ g, h, k] \ \\<^sup>-\<^sup>1[f, g, h \ k]" proof - have "((\\<^sup>-\<^sup>1[f, g, h] \ k) \ \\<^sup>-\<^sup>1[f, g \ h, k]) \ (f \ \\<^sup>-\<^sup>1[g, h, k]) = inv ((f \ \[g, h, k]) \ (\[f, g \ h, k] \ (\[f, g, h] \ k)))" proof - have "inv ((f \ \[g, h, k]) \ (\[f, g \ h, k] \ (\[f, g, h] \ k))) = inv (\[f, g \ h, k] \ (\[f, g, h] \ k)) \ inv (f \ \[g, h, k])" using assms inv_comp [of "\[f, g \ h, k] \ (\[f, g, h] \ k)" "f \ \[g, h, k]"] by force also have "... = (inv (\[f, g, h] \ k) \ inv \[f, g \ h, k]) \ inv (f \ \[g, h, k])" using assms iso_assoc inv_comp by simp also have "... = ((\\<^sup>-\<^sup>1[f, g, h] \ k) \ \\<^sup>-\<^sup>1[f, g \ h, k]) \ (f \ \\<^sup>-\<^sup>1[g, h, k])" using assms inv_hcomp by simp finally show ?thesis by simp qed also have "... = inv (\[f, g, h \ k] \ \[f \ g, h, k])" using assms pentagon by simp also have "... = \\<^sup>-\<^sup>1[f \ g, h, k] \ \\<^sup>-\<^sup>1[f, g, h \ k]" using assms inv_comp by simp finally show ?thesis by auto qed end text \ The following convenience locale extends @{locale bicategory} by pre-interpreting the various functors and natural transformations. \ locale extended_bicategory = bicategory + L: equivalence_functor V V L + R: equivalence_functor V V R + \: natural_isomorphism VVV.comp V HoHV HoVH \\\\\. \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))\ + \': inverse_transformation VVV.comp V HoHV HoVH \\\\\. \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))\ + \: natural_isomorphism V V L map \ + \': inverse_transformation V V L map \ + \: natural_isomorphism V V R map \ + \': inverse_transformation V V R map \ sublocale bicategory \ extended_bicategory V H \ \ src trg proof - interpret L: equivalence_functor V V L using equivalence_functor_L by auto interpret R: equivalence_functor V V R using equivalence_functor_R by auto interpret \': inverse_transformation VVV.comp V HoHV HoVH \\\\\. \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))\ .. interpret \: natural_isomorphism V V L map \ using natural_isomorphism_\ by auto interpret \': inverse_transformation V V L map \ .. interpret \: natural_isomorphism V V R map \ using natural_isomorphism_\ by auto interpret \': inverse_transformation V V R map \ .. interpret extended_bicategory V H \ \ src trg .. show "extended_bicategory V H \ \ src trg" .. qed end diff --git a/thys/Bicategory/BicategoryOfSpans.thy b/thys/Bicategory/BicategoryOfSpans.thy --- a/thys/Bicategory/BicategoryOfSpans.thy +++ b/thys/Bicategory/BicategoryOfSpans.thy @@ -1,14697 +1,14697 @@ (* Title: BicategoryOfSpans Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) section "Bicategories of Spans" theory BicategoryOfSpans imports Category3.ConcreteCategory IsomorphismClass CanonicalIsos EquivalenceOfBicategories SpanBicategory Tabulation begin text \ In this section, we prove CKS Theorem 4, which characterizes up to equivalence the bicategories of spans in a category with pullbacks. The characterization consists of three conditions: BS1: ``Every 1-cell is isomorphic to a composition \g \ f\<^sup>*\, where f and g are maps''; BS2: ``For every span of maps \(f, g)\ there is a 2-cell \\\ such that \(f, \, g)\ is a tabulation''; and BS3: ``Any two 2-cells between the same pair of maps are equal and invertible.'' One direction of the proof, which is the easier direction once it is established that BS1 and BS3 are respected by equivalence of bicategories, shows that if a bicategory \B\ is biequivalent to the bicategory of spans in some category \C\ with pullbacks, then it satisfies BS1 -- BS3. The other direction, which is harder, shows that a bicategory \B\ satisfying BS1 -- BS3 is biequivalent to the bicategory of spans in a certain category with pullbacks that is constructed from the sub-bicategory of maps of \B\. \ subsection "Definition" text \ We define a \emph{bicategory of spans} to be a bicategory that satisfies the conditions \BS1\ -- \BS3\ stated informally above. \ locale bicategory_of_spans = bicategory + chosen_right_adjoints + - assumes BS1: "\r. ide r \ \f g. is_left_adjoint f \ is_left_adjoint g \ isomorphic r (g \ f\<^sup>*)" - and BS2: "\f g. \ is_left_adjoint f; is_left_adjoint g; src f = src g \ + assumes BS1: "ide r \ \f g. is_left_adjoint f \ is_left_adjoint g \ isomorphic r (g \ f\<^sup>*)" + and BS2: "\ is_left_adjoint f; is_left_adjoint g; src f = src g \ \ \r \. tabulation V H \ \ src trg r \ f g" - and BS3: "\f f' \ \'. \ is_left_adjoint f; is_left_adjoint f'; \\ : f \ f'\; \\' : f \ f'\ \ + and BS3: "\ is_left_adjoint f; is_left_adjoint f'; \\ : f \ f'\; \\' : f \ f'\ \ \ iso \ \ iso \' \ \ = \'" text \ Using the already-established fact \equivalence_pseudofunctor.reflects_tabulation\ that tabulations are reflected by equivalence pseudofunctors, it is not difficult to prove that the notion `bicategory of spans' respects equivalence of bicategories. \ lemma bicategory_of_spans_respects_equivalence: assumes "equivalent_bicategories V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D" and "bicategory_of_spans V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C" shows "bicategory_of_spans V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D" proof - interpret C: bicategory_of_spans V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C using assms by simp interpret C: chosen_right_adjoints V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C .. interpret D: bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D using assms equivalent_bicategories_def equivalence_pseudofunctor.axioms(1) pseudofunctor.axioms(2) by fast interpret D: chosen_right_adjoints V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D .. obtain F \ where F: "equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \" using assms equivalent_bicategories_def by blast interpret F: equivalence_pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ using F by simp interpret E: equivalence_of_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C (* 17 sec *) F \ F.right_map F.right_cmp F.unit\<^sub>0 F.unit\<^sub>1 F.counit\<^sub>0 F.counit\<^sub>1 using F.extends_to_equivalence_of_bicategories by simp interpret E': converse_equivalence_of_bicategories V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F \ F.right_map F.right_cmp F.unit\<^sub>0 F.unit\<^sub>1 F.counit\<^sub>0 F.counit\<^sub>1 .. interpret G: equivalence_pseudofunctor V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C F.right_map F.right_cmp using E'.equivalence_pseudofunctor_left by simp write V\<^sub>C (infixr "\\<^sub>C" 55) write V\<^sub>D (infixr "\\<^sub>D" 55) write H\<^sub>C (infixr "\\<^sub>C" 53) write H\<^sub>D (infixr "\\<^sub>D" 53) write \\<^sub>C ("\\<^sub>C[_, _, _]") write \\<^sub>D ("\\<^sub>D[_, _, _]") write C.in_hhom ("\_ : _ \\<^sub>C _\") write C.in_hom ("\_ : _ \\<^sub>C _\") write D.in_hhom ("\_ : _ \\<^sub>D _\") write D.in_hom ("\_ : _ \\<^sub>D _\") write C.isomorphic (infix "\\<^sub>C" 50) write D.isomorphic (infix "\\<^sub>D" 50) write C.some_right_adjoint ("_\<^sup>*\<^sup>C" [1000] 1000) write D.some_right_adjoint ("_\<^sup>*\<^sup>D" [1000] 1000) show "bicategory_of_spans V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D" proof show "\r'. D.ide r' \ \f' g'. D.is_left_adjoint f' \ D.is_left_adjoint g' \ r' \\<^sub>D g' \\<^sub>D (f')\<^sup>*\<^sup>D" proof - fix r' assume r': "D.ide r'" obtain f g where fg: "C.is_left_adjoint f \ C.is_left_adjoint g \ F.right_map r' \\<^sub>C g \\<^sub>C f\<^sup>*\<^sup>C" using r' C.BS1 [of "F.right_map r'"] by auto have trg_g: "trg\<^sub>C g = E.G.map\<^sub>0 (trg\<^sub>D r')" using fg r' C.isomorphic_implies_ide C.isomorphic_implies_hpar by (metis C.ideD(1) C.trg_hcomp D.ideD(1) E.G.preserves_trg) have trg_f: "trg\<^sub>C f = E.G.map\<^sub>0 (src\<^sub>D r')" using fg r' C.isomorphic_implies_ide C.isomorphic_implies_hpar by (metis C.ideD(1) C.right_adjoint_simps(2) C.src_hcomp D.ideD(1) E.G.preserves_src) define d_src where "d_src \ F.counit\<^sub>0 (src\<^sub>D r')" define e_src where "e_src \ (F.counit\<^sub>0 (src\<^sub>D r'))\<^sup>~\<^sup>D" have d_src: "\d_src : F.map\<^sub>0 (E.G.map\<^sub>0 (src\<^sub>D r')) \\<^sub>D src\<^sub>D r'\ \ D.equivalence_map d_src" using d_src_def r' E.\.map\<^sub>0_in_hhom E.\.components_are_equivalences by simp have e_src: "\e_src : src\<^sub>D r' \\<^sub>D F.map\<^sub>0 (E.G.map\<^sub>0 (src\<^sub>D r'))\ \ D.equivalence_map e_src" using e_src_def r' E.\.map\<^sub>0_in_hhom E.\.components_are_equivalences by simp obtain \_src \_src where eq_src: "equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D d_src e_src \_src \_src" using d_src_def e_src_def d_src e_src D.quasi_inverses_some_quasi_inverse D.quasi_inverses_def by blast interpret eq_src: equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D d_src e_src \_src \_src using eq_src by simp define d_trg where "d_trg \ F.counit\<^sub>0 (trg\<^sub>D r')" define e_trg where "e_trg \ (F.counit\<^sub>0 (trg\<^sub>D r'))\<^sup>~\<^sup>D" have d_trg: "\d_trg : F.map\<^sub>0 (E.G.map\<^sub>0 (trg\<^sub>D r')) \\<^sub>D trg\<^sub>D r'\ \ D.equivalence_map d_trg" using d_trg_def r' E.\.map\<^sub>0_in_hhom E.\.components_are_equivalences by simp have e_trg: "\e_trg : trg\<^sub>D r' \\<^sub>D F.map\<^sub>0 (E.G.map\<^sub>0 (trg\<^sub>D r'))\ \ D.equivalence_map e_trg" using e_trg_def r' E.\.map\<^sub>0_in_hhom E.\.components_are_equivalences by simp obtain \_trg \_trg where eq_trg: "equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D d_trg e_trg \_trg \_trg" using d_trg_def e_trg_def d_trg e_trg D.quasi_inverses_some_quasi_inverse D.quasi_inverses_def by blast interpret eq_trg: equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D d_trg e_trg \_trg \_trg using eq_trg by simp interpret eqs: two_equivalences_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D d_src e_src \_src \_src d_trg e_trg \_trg \_trg .. interpret hom: subcategory V\<^sub>D \\\. \\ : trg\<^sub>D d_src \\<^sub>D trg\<^sub>D d_trg\\ using D.hhom_is_subcategory by simp interpret hom': subcategory V\<^sub>D \\\. \\ : src\<^sub>D d_src \\<^sub>D src\<^sub>D d_trg\\ using D.hhom_is_subcategory by simp interpret e: equivalence_of_categories hom.comp hom'.comp eqs.F eqs.G eqs.\ eqs.\ using eqs.induces_equivalence_of_hom_categories by simp have r'_in_hhom: "D.in_hhom r' (src\<^sub>D e_src) (src\<^sub>D e_trg)" using r' e_src e_trg by (simp add: D.in_hhom_def) define g' where "g' = d_trg \\<^sub>D F g" have g': "D.is_left_adjoint g'" unfolding g'_def using fg r' d_trg trg_g C.left_adjoint_is_ide D.equivalence_is_adjoint D.left_adjoints_compose F.preserves_left_adjoint C.ideD(1) D.in_hhom_def F.preserves_trg by metis have 1: "D.is_right_adjoint (F f\<^sup>*\<^sup>C \\<^sub>D e_src)" proof - have "D.is_right_adjoint e_src" using r' e_src D.equivalence_is_adjoint by simp moreover have "D.is_right_adjoint (F f\<^sup>*\<^sup>C)" using fg C.left_adjoint_extends_to_adjoint_pair F.preserves_adjoint_pair by blast moreover have "src\<^sub>D (F f\<^sup>*\<^sup>C) = trg\<^sub>D e_src" using fg r' trg_f C.right_adjoint_is_ide e_src by auto ultimately show ?thesis using fg r' D.right_adjoints_compose F.preserves_right_adjoint by blast qed obtain f' where f': "D.adjoint_pair f' (F f\<^sup>*\<^sup>C \\<^sub>D e_src)" using 1 by auto have f': "D.is_left_adjoint f' \ F f\<^sup>*\<^sup>C \\<^sub>D e_src \\<^sub>D (f')\<^sup>*\<^sup>D" using f' D.left_adjoint_determines_right_up_to_iso D.left_adjoint_extends_to_adjoint_pair by blast have "r' \\<^sub>D d_trg \\<^sub>D (e_trg \\<^sub>D r' \\<^sub>D d_src) \\<^sub>D e_src" using r' r'_in_hhom D.isomorphic_def eqs.\_in_hom eqs.\_components_are_iso D.isomorphic_symmetric D.ide_char eq_src.antipar(2) eq_trg.antipar(2) by metis also have 1: "... \\<^sub>D d_trg \\<^sub>D F (F.right_map r') \\<^sub>D e_src" proof - have "e_trg \\<^sub>D r' \\<^sub>D d_src \\<^sub>D F (F.right_map r')" proof - have "D.in_hom (F.counit\<^sub>1 r') (r' \\<^sub>D d_src) (F.counit\<^sub>0 (trg\<^sub>D r') \\<^sub>D F (F.right_map r'))" unfolding d_src_def using r' E.\.map\<^sub>1_in_hom(2) [of r'] by simp hence "r' \\<^sub>D d_src \\<^sub>D F.counit\<^sub>0 (trg\<^sub>D r') \\<^sub>D F (F.right_map r')" using r' D.isomorphic_def E.\.iso_map\<^sub>1_ide by auto thus ?thesis using r' e_trg_def E.\.components_are_equivalences D.isomorphic_symmetric D.quasi_inverse_transpose(2) by (metis D.isomorphic_implies_hpar(1) F.preserves_isomorphic d_trg d_trg_def eq_trg.ide_left fg) qed thus ?thesis using D.hcomp_ide_isomorphic D.hcomp_isomorphic_ide D.in_hhom_def D.isomorphic_implies_hpar(4) d_trg e_src eq_src.antipar(1-2) eq_trg.antipar(2) r' by force qed also have 2: "... \\<^sub>D d_trg \\<^sub>D (F g \\<^sub>D F f\<^sup>*\<^sup>C) \\<^sub>D e_src" proof - have "F (F.right_map r') \\<^sub>D F g \\<^sub>D F f\<^sup>*\<^sup>C" by (meson C.hseq_char C.ideD(1) C.isomorphic_implies_ide(2) C.left_adjoint_is_ide C.right_adjoint_simps(1) D.isomorphic_symmetric D.isomorphic_transitive F.preserves_isomorphic F.weakly_preserves_hcomp fg) thus ?thesis using D.hcomp_ide_isomorphic D.hcomp_isomorphic_ide by (metis 1 D.hseqE D.ideD(1) D.isomorphic_implies_hpar(2) eq_src.ide_right eq_trg.ide_left) qed also have 3: "... \\<^sub>D (d_trg \\<^sub>D F g) \\<^sub>D F f\<^sup>*\<^sup>C \\<^sub>D e_src" proof - have "... \\<^sub>D d_trg \\<^sub>D F g \\<^sub>D F f\<^sup>*\<^sup>C \\<^sub>D e_src" by (metis C.left_adjoint_is_ide C.right_adjoint_simps(1) D.hcomp_assoc_isomorphic D.hcomp_ide_isomorphic D.hcomp_simps(1) D.hseq_char D.ideD(1) D.isomorphic_implies_hpar(2) F.preserves_ide calculation eq_src.ide_right eq_trg.ide_left fg) also have "... \\<^sub>D (d_trg \\<^sub>D F g) \\<^sub>D F f\<^sup>*\<^sup>C \\<^sub>D e_src" by (metis C.left_adjoint_is_ide D.hcomp_assoc_isomorphic D.hcomp_simps(2) D.hseq_char D.ideD(1) D.isomorphic_implies_ide(1) D.isomorphic_symmetric F.preserves_ide calculation eq_trg.ide_left f' fg) finally show ?thesis by blast qed also have "... \\<^sub>D g' \\<^sub>D f'\<^sup>*\<^sup>D" using g'_def f' by (metis 3 D.adjoint_pair_antipar(1) D.hcomp_ide_isomorphic D.hseq_char D.ideD(1) D.isomorphic_implies_ide(2) g') finally have "D.isomorphic r' (g' \\<^sub>D f'\<^sup>*\<^sup>D)" by simp thus "\f' g'. D.is_left_adjoint f' \ D.is_left_adjoint g' \ r' \\<^sub>D g' \\<^sub>D f'\<^sup>*\<^sup>D" using f' g' by auto qed show "\f f' \ \'. \ D.is_left_adjoint f; D.is_left_adjoint f'; \\ : f \\<^sub>D f'\; \\' : f \\<^sub>D f'\ \ \ D.iso \ \ D.iso \' \ \ = \'" proof - fix f f' \ \' assume f: "D.is_left_adjoint f" and f': "D.is_left_adjoint f'" and \: "\\ : f \\<^sub>D f'\" and \': "\\' : f \\<^sub>D f'\" have "C.is_left_adjoint (F.right_map f) \ C.is_left_adjoint (F.right_map f')" using f f' E.G.preserves_left_adjoint by blast moreover have "\F.right_map \ : F.right_map f \\<^sub>C F.right_map f'\ \ \F.right_map \' : F.right_map f \\<^sub>C F.right_map f'\" using \ \' E.G.preserves_hom by simp ultimately have "C.iso (F.right_map \) \ C.iso (F.right_map \') \ F.right_map \ = F.right_map \'" using C.BS3 by blast thus "D.iso \ \ D.iso \' \ \ = \'" using \ \' G.reflects_iso G.is_faithful by blast qed show "\f g. \ D.is_left_adjoint f; D.is_left_adjoint g; src\<^sub>D f = src\<^sub>D g \ \ \r \. tabulation V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D r \ f g" proof - fix f g assume f: "D.is_left_adjoint f" assume g: "D.is_left_adjoint g" assume fg: "src\<^sub>D f = src\<^sub>D g" have "C.is_left_adjoint (F.right_map f)" using f E.G.preserves_left_adjoint by blast moreover have "C.is_left_adjoint (F.right_map g)" using g E.G.preserves_left_adjoint by blast moreover have "src\<^sub>C (F.right_map f) = src\<^sub>C (F.right_map g)" using f g D.left_adjoint_is_ide fg by simp ultimately have 1: "\r \. tabulation V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C r \ (F.right_map f) (F.right_map g)" using C.BS2 by simp obtain r \ where \: "tabulation V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C r \ (F.right_map f) (F.right_map g)" using 1 by auto interpret \: tabulation V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C r \ \F.right_map f\ \F.right_map g\ using \ by simp obtain r' where r': "D.ide r' \ D.in_hhom r' (trg\<^sub>D f) (trg\<^sub>D g) \ C.isomorphic (F.right_map r') r" using f g \.ide_base \.tab_in_hom G.locally_essentially_surjective by (metis D.obj_trg E.G.preserves_reflects_arr E.G.preserves_trg \.leg0_simps(2-3) \.leg1_simps(2,4) \.base_in_hom(1)) obtain \ where \: "\\ : r \\<^sub>C F.right_map r'\ \ C.iso \" using r' C.isomorphic_symmetric by blast have \: "tabulation V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C (F.right_map r') ((\ \\<^sub>C F.right_map f) \\<^sub>C \) (F.right_map f) (F.right_map g)" using \ \.is_preserved_by_base_iso by simp have 1: "\\'. \\' : g \\<^sub>D H\<^sub>D r' f\ \ F.right_map \' = F.right_cmp (r', f) \\<^sub>C (\ \\<^sub>C F.right_map f) \\<^sub>C \" proof - have "D.ide g" by (simp add: D.left_adjoint_is_ide g) moreover have "D.ide (H\<^sub>D r' f)" using f r' D.left_adjoint_is_ide by auto moreover have "src\<^sub>D g = src\<^sub>D (H\<^sub>D r' f)" using fg by (simp add: calculation(2)) moreover have "trg\<^sub>D g = trg\<^sub>D (H\<^sub>D r' f)" using calculation(2) r' by auto moreover have "\F.right_cmp (r', f) \\<^sub>C (\ \\<^sub>C F.right_map f) \\<^sub>C \ : F.right_map g \\<^sub>C F.right_map (r' \\<^sub>D f)\" using f g r' \ D.left_adjoint_is_ide \.ide_base by (intro C.comp_in_homI, auto) ultimately show ?thesis using G.locally_full by simp qed obtain \' where \': "\\' : g \\<^sub>D H\<^sub>D r' f\ \ F.right_map \' = F.right_cmp (r', f) \\<^sub>C (\ \\<^sub>C F.right_map f) \\<^sub>C \" using 1 by auto have "tabulation V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D r' \' f g" proof - have "C.inv (F.right_cmp (r', f)) \\<^sub>C F.right_map \' = (\ \\<^sub>C F.right_map f) \\<^sub>C \" using r' f \' C.comp_assoc C.comp_cod_arr C.invert_side_of_triangle(1) by (metis D.adjoint_pair_antipar(1) D.arrI D.in_hhomE E.G.cmp_components_are_iso E.G.preserves_arr) thus ?thesis using \ \' G.reflects_tabulation by (simp add: D.left_adjoint_is_ide f r') qed thus "\r' \'. tabulation V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D r' \' f g" by auto qed qed qed subsection "Span(C) is a Bicategory of Spans" text \ We first consider an arbitrary 1-cell \r\ in \Span(C)\, and show that it has a tabulation as a span of maps. This is CKS Proposition 3 (stated more strongly to assert that the ``output leg'' can also be taken to be a map, which the proof shows already). \ locale identity_arrow_in_span_bicategory = (* 20 sec *) span_bicategory C prj0 prj1 + r: identity_arrow_of_spans C r for C :: "'a comp" (infixr "\" 55) and prj0 :: "'a \ 'a \ 'a" ("\

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

\<^sub>1[_, _]") and r :: "'a arrow_of_spans_data" begin text \ CKS say: ``Suppose \r = (r\<^sub>0, R, r\<^sub>1): A \ B\ and put \f = (1, R, r\<^sub>0)\, \g = (1, R, r\<^sub>1)\. Let \k\<^sub>0, k\<^sub>1\ form a kernel pair for \r\<^sub>0\ and define \\\ by \k\<^sub>0\ = k\<^sub>1\ = 1\<^sub>R\.'' \ abbreviation ra where "ra \ r.dom.apex" abbreviation r0 where "r0 \ r.dom.leg0" abbreviation r1 where "r1 \ r.dom.leg1" abbreviation f where "f \ mkIde ra r0" abbreviation g where "g \ mkIde ra r1" abbreviation k0 where "k0 \ \

\<^sub>0[r0, r0]" abbreviation k1 where "k1 \ \

\<^sub>1[r0, r0]" text \ Here \ra\ is the apex \R\ of the span \(r\<^sub>0, R, r\<^sub>1)\, and the spans \f\ and \g\ also have that same 0-cell as their apex. The tabulation 2-cell \\\ has to be an arrow of spans from \g = (1, R, r\<^sub>1)\ to \r \ f\, which is the span \(k\<^sub>0, r\<^sub>1 \ k\<^sub>1)\. \ abbreviation \ where "\ \ \Chn = \ra \r0, r0\ ra\, Dom = \Leg0 = ra, Leg1 = r1\, Cod = \Leg0 = k0, Leg1 = r1 \ k1\\" lemma has_tabulation: shows "tabulation vcomp hcomp assoc unit src trg r \ f g" and "is_left_adjoint f" and "is_left_adjoint g" proof - have ide_f: "ide f" using ide_mkIde r.dom.leg_in_hom(1) C.arr_dom C.dom_dom r.dom.apex_def r.dom.leg_simps(1) by presburger interpret f: identity_arrow_of_spans C f using ide_f ide_char' by auto have ide_g: "ide g" using ide_mkIde r.dom.leg_in_hom by (metis C.arr_dom C.dom_dom r.dom.leg_simps(3) r.dom.leg_simps(4)) interpret g: identity_arrow_of_spans C g using ide_g ide_char' by auto show "is_left_adjoint f" using is_left_adjoint_char [of f] ide_f by simp show "is_left_adjoint g" using is_left_adjoint_char [of g] ide_g by simp have ide_r: "ide r" using ide_char' r.identity_arrow_of_spans_axioms by auto have src_r: "src r = mkObj (C.cod r0)" by (simp add: ide_r src_def) have trg_r: "trg r = mkObj (C.cod r1)" by (simp add: ide_r trg_def) have src_f: "src f = mkObj ra" using ide_f src_def by auto have trg_f: "trg f = mkObj (C.cod r0)" using ide_f trg_def by auto have src_g: "src g = mkObj ra" using ide_g src_def by auto have trg_g: "trg g = mkObj (C.cod r1)" using ide_g trg_def by auto have hseq_rf: "hseq r f" using ide_r ide_f src_r trg_f by simp interpret rf: two_composable_arrows_of_spans C prj0 prj1 r f using hseq_rf hseq_char by unfold_locales auto interpret rf: two_composable_identity_arrows_of_spans C prj0 prj1 r f .. interpret rf: identity_arrow_of_spans C \r \ f\ using rf.ide_composite ide_char' by auto let ?rf = "r \ f" (* TODO: Put this expansion into two_composable_identity_arrows_of_spans. *) have rf: "?rf = \Chn = r0 \\ r0, Dom = \Leg0 = k0, Leg1 = r1 \ k1\, Cod = \Leg0 = k0, Leg1 = r1 \ k1\\" unfolding hcomp_def chine_hcomp_def using hseq_rf C.comp_cod_arr by auto interpret Cod_rf: span_in_category C \\Leg0 = k0, Leg1 = r1 \ k1\\ using ide_r ide_f rf C.comp_cod_arr by unfold_locales auto have Dom_g: "Dom g = \Leg0 = ra, Leg1 = r1\" by simp interpret Dom_g: span_in_category C \\Leg0 = ra, Leg1 = r1\\ using Dom_g g.dom.span_in_category_axioms by simp interpret Dom_\: span_in_category C \Dom \\ using Dom_g g.dom.span_in_category_axioms by simp interpret Cod_\: span_in_category C \Cod \\ using rf Cod_rf.span_in_category_axioms by simp interpret \: arrow_of_spans C \ using Dom_\.apex_def Cod_\.apex_def C.comp_assoc C.comp_arr_dom by unfold_locales auto have \: "\\ : g \ r \ f\" using rf ide_g arr_char dom_char cod_char \.arrow_of_spans_axioms ideD(2) Cod_rf.apex_def g.dom.leg_simps(4) by auto show "tabulation vcomp hcomp assoc unit src trg r \ f g" proof - interpret T: tabulation_data vcomp hcomp assoc unit src trg r \ f g using ide_f \ by unfold_locales auto show ?thesis proof show T1: "\u \. \ ide u; \\ : dom \ \ r \ u\ \ \ \w \ \. ide w \ \\ : f \ w \ u\ \ \\ : dom \ \ g \ w\ \ iso \ \ T.composite_cell w \ \ \ = \" proof - fix u \ assume u: "ide u" assume \: "\\ : dom \ \ r \ u\" show "\w \ \. ide w \ \\ : f \ w \ u\ \ \\ : dom \ \ g \ w\ \ iso \ \ T.composite_cell w \ \ \ = \" proof - interpret u: identity_arrow_of_spans C u using u ide_char' by auto have v: "ide (dom \)" using \ by auto interpret v: identity_arrow_of_spans C \dom \\ using v ide_char' by auto interpret \: arrow_of_spans C \ using \ arr_char by auto have hseq_ru: "hseq r u" using u \ ide_cod by fastforce interpret ru: two_composable_arrows_of_spans C prj0 prj1 r u using hseq_ru hseq_char by unfold_locales auto interpret ru: two_composable_identity_arrows_of_spans C prj0 prj1 r u .. text \ CKS say: ``We must show that \(f, \, g)\ is a wide tabulation of \r\. Take \u = (u\<^sub>0, U, u\<^sub>1): X \ A\, \v = (v\<^sub>0, V, v\<^sub>1): X \ B\, \\: v \ ru\ as in \T1\. Let \P\ be the pullback of \u\<^sub>1, r\<^sub>0\. Let \w = (v\<^sub>0, V, p\<^sub>1\): X \ R\, \\ = p\<^sub>0\: fw \ u\, \\ = 1: v \ gw\; so \\ = (r\)(\w)\\ as required.'' \ let ?R = "r.apex" let ?A = "C.cod r0" let ?B = "C.cod r1" let ?U = "u.apex" let ?u0 = "u.leg0" let ?u1 = "u.leg1" let ?X = "C.cod ?u0" let ?V = "v.apex" let ?v0 = "v.leg0" let ?v1 = "v.leg1" let ?\ = "\.chine" let ?P = "r0 \\ ?u1" let ?p0 = "\

\<^sub>0[r0, ?u1]" let ?p1 = "\

\<^sub>1[r0, ?u1]" let ?w1 = "?p1 \ ?\" define w where "w = mkIde ?v0 ?w1" let ?Q = "?R \\ ?w1" let ?q1 = "\

\<^sub>0[?R, ?w1]\. What CKS say is only correct if the projection \\

\<^sub>0[?R, ?w1]\ is an identity, which can't be guaranteed for an arbitrary choice of pullbacks. \ define \ where "\ = \Chn = ?p0 \ ?\ \ \

\<^sub>0[?R, ?w1], Dom = Dom (f \ w), Cod = Cod u\" interpret Dom_\: span_in_category C \Dom \\ using \_def fw.dom.span_in_category_axioms by simp interpret Cod_\: span_in_category C \Cod \\ using \_def u.cod.span_in_category_axioms by simp have Dom_\_leg0_eq: "Dom_\.leg0 = ?v0 \ \

\<^sub>0[?R, ?w1]" using w_def \_def hcomp_def hseq_fw hseq_char by simp have Dom_\_leg1_eq: "Dom_\.leg1 = r0 \ ?q1" using w_def \_def hcomp_def hseq_fw hseq_char by simp have Cod_\_leg0_eq: "Cod_\.leg0 = ?u0" using w_def \_def hcomp_def hseq_fw hseq_char by simp have Cod_\_leg1_eq: "Cod_\.leg1 = ?u1" using w_def \_def hcomp_def hseq_fw hseq_char by simp have Chn_\_eq: "Chn \ = ?p0 \ ?\ \ \

\<^sub>0[?R, ?w1]" using \_def by simp interpret \: arrow_of_spans C \ proof show 1: "\Chn \ : Dom_\.apex \\<^sub>C Cod_\.apex\" using \_def Chn_\ ru.legs_form_cospan fw_apex_eq by (intro C.in_homI, auto) show "Cod_\.leg0 \ Chn \ = Dom_\.leg0" proof - have "Cod_\.leg0 \ Chn \ = ?u0 \ ?p0 \ ?\ \ \

\<^sub>0[?R, ?w1]" using Cod_\_leg0_eq Chn_\_eq by simp also have "... = ?v0 \ \

\<^sub>0[?R, ?w1]" proof - have "?u0 \ ?p0 \ ?\ \ \

\<^sub>0[?R, ?w1] = (?u0 \ ?p0 \ ?\) \ \

\<^sub>0[?R, ?w1]" using C.comp_assoc by simp also have "... = ?v0 \ \

\<^sub>0[?R, ?w1]" proof - have "?u0 \ ?p0 \ ?\ = (?u0 \ ?p0) \ ?\" using C.comp_assoc by simp also have "... = \.cod.leg0 \ ?\" by (metis \ arrow_of_spans_data.select_convs(2) cod_char in_homE ru.leg0_composite) also have "... = \.dom.leg0" using \.leg0_commutes by simp also have "... = ?v0" using \ dom_char by auto finally show ?thesis by simp qed finally show ?thesis by simp qed also have "... = Dom_\.leg0" using Dom_\_leg0_eq by simp finally show ?thesis by blast qed show "Cod_\.leg1 \ Chn \ = Dom_\.leg1" proof - have "Cod_\.leg1 \ Chn \ = ?u1 \ ?p0 \ ?\ \ \

\<^sub>0[?R, ?w1]" using Cod_\_leg1_eq Chn_\_eq by simp also have "... = r0 \ ?q1" proof - have "?u1 \ ?p0 \ ?\ \ \

\<^sub>0[?R, ?w1] = ((?u1 \ ?p0) \ ?\) \ \

\<^sub>0[?R, ?w1]" using C.comp_assoc by fastforce also have "... = ((r0 \ ?p1) \ ?\) \ \

\<^sub>0[?R, ?w1]" using C.pullback_commutes' ru.legs_form_cospan by simp also have "... = r0 \ ?w1 \ \

\<^sub>0[?R, ?w1]" using C.comp_assoc by fastforce also have "... = r0 \ ?R \ ?q1" using \ C.in_homE C.pullback_commutes' w1 by auto also have "... = r0 \ ?q1" using \ w1 C.comp_cod_arr by auto finally show ?thesis by simp qed also have "... = Dom_\.leg1" using Dom_\_leg1_eq by simp finally show ?thesis by blast qed qed text \ Similarly, CKS say to take \\ = 1: v \ gw\, but obviously this can't be interpreted literally, either, because \v.apex\ and \gw.apex\ are not equal. Instead, we have to define \\\ so that \Chn \ = C.inv \

\<^sub>0[?R, ?w1]\, noting that \\

\<^sub>0[?R, ?w1]\ is the pullback of an identity, hence is an isomorphism. Then \?v0 \ \

\<^sub>0[?R, ?w1] \ Chn \ = ?v0\ and \?v1 \ \

\<^sub>1[?R, ?w1] \ Chn \ = ?v1 \ ?w1\, showing that \\\ is an arrow of spans. \ let ?\' = "\

\<^sub>0[?R, ?w1]" define \ where "\ = \Chn = C.inv ?\', Dom = Dom (dom \), Cod = Cod (g \ w)\" have iso_\: "C.inverse_arrows ?\' (Chn \)" using \_def \ w1 C.iso_pullback_ide by (metis C.inv_is_inverse C.seqE arrow_of_spans_data.select_convs(1) r.chine_eq_apex r.chine_simps(1) r.chine_simps(3) r.cod_simps(1) r.dom.apex_def r.dom.ide_apex r.dom.is_span r1w1 v.dom.leg_simps(3)) text \ $$ \xymatrix{ && X \\ && V \ar[u]_{v_0} \ar[d]_{\omega} \ar@/ul50pt/[ddddll]_{v_1} \ar@/l30pt/[dd]_<>(0.7){w_1}\\ & R\downarrow\downarrow w_1 \ar[ur]^{\nu'} \ar[dd]_{q_1} & r_0\downarrow\downarrow u_1 \ar[d]_{p_1} \ar@/dl10pt/[drr]_<>(0.4){p_0} & R\downarrow\downarrow w_1 \ar[ul]_{\nu'} \ar[dd]^<>(0.7){q_1} \ar@ {.>}[dr]_{\theta}\\ && R && U \ar@/ur20pt/[uuull]_{u_0} \ar[dd]^{u_1} \\ & R \ar[dl]_{r_1} \ar@ {<->}[ur]_{R} \ar@ {.>}[dr]^{\rho} \ar@/dl5pt/[ddr]_<>(0.4){R} && R \ar@ {<->}[ul]^{R} \ar[dr]^{r_0} \ar[ur]_{r_1}\\ B && r_0\downarrow\downarrow r_0 \ar[uu]_{k_0} \ar[d]^{k_1} \ar[uu] \ar[ur]_{k_0} && A \\ & & R \ar[ull]^{r_1} \ar[urr]_{r_0} \\ } $$ \ have w1_eq: "?w1 = ?q1 \ C.inv ?\'" proof - have "?R \ ?q1 = ?w1 \ ?\'" using iso_\ \ w1 C.pullback_commutes [of ?R ?w1] by blast hence "?q1 = ?w1 \ ?\'" using \ w1 C.comp_cod_arr by auto thus ?thesis using iso_\ C.invert_side_of_triangle(2) by (metis C.isoI C.prj1_simps(1) arrow_of_spans_data.select_convs(3) fw.legs_form_cospan(2) span_data.simps(1-2) w_def) qed interpret Dom_\: span_in_category C \Dom \\ using \_def v.dom.span_in_category_axioms by simp interpret Cod_\: span_in_category C \Cod \\ using \_def gw.cod.span_in_category_axioms by simp interpret \: arrow_of_spans C \ proof show 1: "\Chn \ : Dom_\.apex \\<^sub>C Cod_\.apex\" proof show "C.arr (Chn \)" using iso_\ by auto show "C.dom (Chn \) = Dom_\.apex" using \_def iso_\ w1 gw_apex_eq by fastforce show "C.cod (Chn \) = Cod_\.apex" using \_def iso_\ gw_apex_eq C.comp_inv_arr C.pbdom_def by (metis C.cod_comp arrow_of_spans_data.select_convs(3) gw.apex_composite gw.chine_composite gw.chine_simps(1,3)) qed show "Cod_\.leg0 \ Chn \ = Dom_\.leg0" using w_def \_def 1 iso_\ hcomp_def hseq_gw C.comp_arr_inv C.comp_assoc v.leg0_commutes by auto show "Cod_\.leg1 \ Chn \ = Dom_\.leg1" using w_def \_def hcomp_def hseq_gw C.comp_assoc w1_eq r1w1 by auto qed text \ Now we can proceed to establishing the conclusions of \T1\. \ have "ide w \ \\ : f \ w \ u\ \ \\ : dom \ \ dom \ \ w\ \ iso \ \ T.composite_cell w \ \ \ = \" proof (intro conjI) show ide_w: "ide w" using w by blast show \: "\\ : f \ w \ u\" using \_def \.arrow_of_spans_axioms arr_char dom_char cod_char hseq_fw hseq_char hcomp_def fw.chine_eq_apex by auto show \: "\\ : dom \ \ dom \ \ w\" proof - have "\\ : dom \ \ g \ w\" using \_def \ \.arrow_of_spans_axioms arr_char dom_char cod_char hseq_gw hseq_char hcomp_def gw.chine_eq_apex by auto thus ?thesis using T.tab_in_hom by simp qed show "iso \" using iso_\ iso_char arr_char \.arrow_of_spans_axioms by auto show "T.composite_cell w \ \ \ = \" proof (intro arr_eqI) have \w: "\\ \ w : g \ w \ (r \ f) \ w\" using w_def \ ide_w hseq_rf hseq_fw hseq_gw by auto have r\: "\r \ \ : r \ f \ w \ r \ u\" using arfw ide_r \ fw.composite_simps(2) rf.composable by auto have 1: "\T.composite_cell w \ \ \ : dom \ \ r \ u\" using \ \w arfw r\ by auto show 3: "par (T.composite_cell w \ \ \) \" using 1 \ by (elim in_homE, auto) show "Chn (T.composite_cell w \ \ \) = ?\" proof - have 2: "Chn (T.composite_cell w \ \ \) = Chn (r \ \) \ Chn \[r, f, w] \ Chn (\ \ w) \ Chn \" using 1 3 Chn_vcomp C.comp_assoc by (metis (full_types) seqE) also have "... = ?\" proof - let ?LHS = "Chn (r \ \) \ Chn \[r, f, w] \ Chn (\ \ w) \ Chn \" have Chn_r\: "Chn (r \ \) = \r.chine \ \

\<^sub>1[r0, r0 \ ?q1] \r0, ?u1\ \.chine \ \

\<^sub>0[r0, r0 \ ?q1]\" using r\ hcomp_def \_def chine_hcomp_def Dom_\_leg1_eq by (metis arrI arrow_of_spans_data.select_convs(1,3) hseq_char r.cod_simps(2) u.cod_simps(3)) have Chn_arfw: "Chn \[r, f, w] = rfw.chine_assoc" using \_ide ide_f rf.composable fw.composable w by auto have Chn_\w: "Chn (\ \ w) = \?\ \ ?q1 \k0, ?w1\ ?\'\" proof - have "Chn (\ \ w) = chine_hcomp \Chn = ?\, Dom = \Leg0 = ?R, Leg1 = r1\, Cod = \Leg0 = k0, Leg1 = r1 \ k1\\ \Chn = v.apex, Dom = \Leg0 = ?v0, Leg1 = ?w1\, Cod = \Leg0 = ?v0, Leg1 = ?w1\\" using \ ide_w hseq_rf hseq_char hcomp_def src_def trg_def by (metis (no_types, lifting) \w arrI arrow_of_spans_data.select_convs(1) v.dom.apex_def w_def) also have "... = \?\ \ ?q1 \k0, ?w1\ ?V \ ?\'\" unfolding chine_hcomp_def by simp also have "... = \?\ \ ?q1 \k0, ?w1\ ?\'\" proof - have "?V \ ?\' = ?\'" using C.comp_ide_arr v.dom.ide_apex \ w1 by auto thus ?thesis by simp qed finally show ?thesis by blast qed have 3: "C.seq ?p1 ?\" using w1 by blast moreover have 4: "C.seq ?p1 ?LHS" proof show "\?LHS : v.apex \\<^sub>C ru.apex\" by (metis (no_types, lifting) 1 2 Chn_in_hom ru.chine_eq_apex v.chine_eq_apex) show "\?p1 : ru.apex \\<^sub>C ?R\" using P C.prj1_in_hom ru.legs_form_cospan by fastforce qed moreover have "?p0 \ ?LHS = ?p0 \ ?\" proof - have "?p0 \ ?LHS = (?p0 \ Chn (r \ \)) \ Chn \[r, f, w] \ Chn (\ \ w) \ Chn \" using C.comp_assoc by simp also have "... = (\.chine \ \

\<^sub>0[r0, r0 \ ?q1]) \ Chn \[r, f, w] \ Chn (\ \ w) \ Chn \" proof - have "?p0 \ Chn (r \ \) = \.chine \ \

\<^sub>0[r0, r0 \ ?q1]" by (metis C.prj_tuple(1) Chn_r\ \_def arrI Dom_\_leg1_eq arrow_of_spans_data.select_convs(3) chine_hcomp_props(2) hseq_char r.cod_simps(2) r\ u.cod_simps(3)) thus ?thesis by argo qed also have "... = ?p0 \ ?\ \ (rfw.Prj\<^sub>0\<^sub>0 \ Chn \[r, f, w]) \ Chn (\ \ w) \ Chn \" using w_def \_def C.comp_assoc by simp also have "... = ?p0 \ ?\ \ (rfw.Prj\<^sub>0 \ Chn (\ \ w)) \ Chn \" using Chn_arfw rfw.prj_chine_assoc C.comp_assoc by simp also have "... = ?p0 \ ?\ \ ?\' \ Chn \" proof - have "rfw.Prj\<^sub>0 \ Chn (\ \ w) = \

\<^sub>0[k0, ?w1] \ \?\ \ ?q1 \k0, ?w1\ ?\'\" using w_def Chn_\w C.comp_cod_arr by simp also have "... = ?\'" by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(1) C.seqE C.tuple_is_extensional Chn_\w 4) finally have "rfw.Prj\<^sub>0 \ Chn (\ \ w) = ?\'" by blast thus ?thesis by simp qed also have "... = ?p0 \ ?\" using iso_\ C.comp_arr_dom by (metis (no_types, lifting) C.comp_arr_inv C.dom_comp \_def \.chine_simps(1) 3 arrow_of_spans_data.simps(1) w1_eq) finally show ?thesis by blast qed moreover have "?p1 \ ?LHS = ?w1" proof - have "?p1 \ ?LHS = (?p1 \ Chn (r \ \)) \ Chn \[r, f, w] \ Chn (\ \ w) \ Chn \" using C.comp_assoc by simp also have "... = (r.chine \ \

\<^sub>1[r0, r0 \ ?q1]) \ Chn \[r, f, w] \ Chn (\ \ w) \ Chn \" by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(2) C.seqE C.tuple_is_extensional Chn_r\ 4) also have "... = r.chine \ (rfw.Prj\<^sub>1 \ Chn \[r, f, w]) \ Chn (\ \ w) \ Chn \" using w_def Dom_\_leg1_eq C.comp_assoc by simp also have "... = r.chine \ (rfw.Prj\<^sub>1\<^sub>1 \ Chn (\ \ w)) \ Chn \" using Chn_arfw rfw.prj_chine_assoc(1) C.comp_assoc by simp also have "... = r.chine \ ?q1 \ Chn \" proof - have "rfw.Prj\<^sub>1\<^sub>1 \ Chn (\ \ w) = (k1 \ \

\<^sub>1[k0, ?w1]) \ \?\ \ ?q1 \k0, ?w1\ ?\'\" using w_def Chn_\w C.comp_cod_arr by simp also have "... = k1 \ \

\<^sub>1[k0, ?w1] \ \?\ \ ?q1 \k0, ?w1\ ?\'\" using C.comp_assoc by simp also have "... = k1 \ ?\ \ ?q1" by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(2) C.seqE C.tuple_is_extensional Chn_\w 4) also have "... = (k1 \ ?\) \ ?q1" using C.comp_assoc by presburger also have "... = ?R \ ?q1" by simp also have "... = ?q1" by (metis Dom_\_leg1_eq C.comp_ide_arr C.prj1_simps(3) C.prj1_simps_arr C.seqE C.seqI Dom_\.leg_simps(3) r.dom.ide_apex) finally have "rfw.Prj\<^sub>1\<^sub>1 \ Chn (\ \ w) = ?q1" by blast thus ?thesis by simp qed also have "... = (r.chine \ ?p1) \ ?\" using \_def w1_eq C.comp_assoc by simp also have "... = ?w1" using C.comp_cod_arr r.chine_eq_apex ru.prj_simps(1) by auto finally show ?thesis by blast qed ultimately show ?thesis using ru.legs_form_cospan C.prj_joint_monic by blast qed finally show ?thesis by argo qed qed qed thus ?thesis using w_def by auto qed qed show T2: "\u w w' \ \' \. \ ide w; ide w'; \\ : f \ w \ u\; \\' : f \ w' \ u\; \\ : g \ w \ g \ w'\; T.composite_cell w \ = T.composite_cell w' \' \ \ \ \ \!\. \\ : w \ w'\ \ \ = g \ \ \ \ = \' \ (f \ \)" proof - fix u w w' \ \' \ assume ide_w: "ide w" assume ide_w': "ide w'" assume \: "\\ : f \ w \ u\" assume \': "\\' : f \ w' \ u\" assume \: "\\ : g \ w \ g \ w'\" assume E: "T.composite_cell w \ = T.composite_cell w' \' \ \" interpret T: uw\w'\'\ vcomp hcomp assoc unit src trg r \ f g u w \ w' \' \ using ide_w ide_w' \ \' \ E comp_assoc by unfold_locales auto show "\!\. \\ : w \ w'\ \ \ = g \ \ \ \ = \' \ (f \ \)" proof interpret u: identity_arrow_of_spans C u using T.uw\.u_simps(1) ide_char' by auto interpret w: identity_arrow_of_spans C w using ide_w ide_char' by auto interpret w': identity_arrow_of_spans C w' using ide_w' ide_char' by auto let ?u0 = u.leg0 let ?u1 = u.leg1 let ?w0 = w.leg0 let ?w1 = w.leg1 let ?wa = "w.apex" let ?w0' = w'.leg0 let ?w1' = w'.leg1 let ?wa' = "w'.apex" let ?R = ra let ?p0 = "\

\<^sub>0[?R, ?w1]" let ?p0' = "\

\<^sub>0[?R, ?w1']" let ?p1 = "\

\<^sub>1[?R, ?w1]" let ?p1' = "\

\<^sub>1[?R, ?w1']" interpret fw: two_composable_identity_arrows_of_spans C prj0 prj1 f w using hseq_char by unfold_locales auto interpret fw': two_composable_identity_arrows_of_spans C prj0 prj1 f w' using hseq_char by unfold_locales auto have hseq_gw: "hseq g w" using T.leg1_in_hom by auto interpret gw: two_composable_identity_arrows_of_spans C prj0 prj1 g w using hseq_gw hseq_char by unfold_locales auto have hseq_gw': "hseq g w'" using T.leg1_in_hom by auto interpret gw': two_composable_identity_arrows_of_spans C prj0 prj1 g w' using hseq_gw' hseq_char by unfold_locales auto interpret rfw: three_composable_identity_arrows_of_spans C prj0 prj1 r f w .. interpret rfw: identity_arrow_of_spans C \r \ f \ w\ using rfw.composites_are_identities ide_char' by auto interpret rfw': three_composable_arrows_of_spans C prj0 prj1 r f w' .. interpret rfw': three_composable_identity_arrows_of_spans C prj0 prj1 r f w' .. interpret rfw': identity_arrow_of_spans C \r \ f \ w'\ using rfw'.composites_are_identities ide_char' by auto have \w: "\\ \ w : g \ w \ (r \ f) \ w\" using \ hseq_gw by blast interpret \w: two_composable_arrows_of_spans C prj0 prj1 \ w using \w by unfold_locales auto have \w': "\\ \ w' : g \ w' \ (r \ f) \ w'\" using \ hseq_gw' by blast interpret \w': two_composable_arrows_of_spans C prj0 prj1 \ w' using \w' by unfold_locales auto have arfw: "\\[r, f, w] : (r \ f) \ w \ r \ f \ w\" using fw.composable ide_f ide_r ide_w rf.composable by auto have arfw': "\\[r, f, w'] : (r \ f) \ w' \ r \ f \ w'\" using fw'.composable ide_f ide_r ide_w' rf.composable by auto have r\: "\r \ \ : r \ f \ w \ r \ u\" by fastforce interpret Dom_\: span_in_category C \Dom \\ using fw.dom.span_in_category_axioms by (metis \ arrow_of_spans_data.select_convs(2) in_homE dom_char) interpret Cod_\: span_in_category C \Cod \\ using \ u.cod.span_in_category_axioms cod_char by auto interpret \: arrow_of_spans C \ using arr_char T.uw\.\_simps(1) by auto interpret r\: two_composable_arrows_of_spans C prj0 prj1 r \ using r\ by unfold_locales auto have r\': "\r \ \' : r \ f \ w' \ r \ u\" by fastforce interpret Dom_\': span_in_category C \Dom \'\ using fw'.dom.span_in_category_axioms by (metis \' arrow_of_spans_data.select_convs(2) in_homE dom_char) interpret Cod_\': span_in_category C \Cod \'\ using \' u.cod.span_in_category_axioms cod_char by auto interpret \': arrow_of_spans C \' using arr_char T.uw'\'.\_simps(1) by auto interpret r\': two_composable_arrows_of_spans C prj0 prj1 r \' using r\' by unfold_locales auto have 7: "\T.composite_cell w' \' \ \ : g \ w \ r \ u\" using \ \w' arfw' r\' by auto have 8: "\T.composite_cell w \ : g \ w \ r \ u\" using \w arfw r\ by auto interpret ru: two_composable_identity_arrows_of_spans C prj0 prj1 r u using hseq_char by unfold_locales auto interpret Dom_\: span_in_category C \Dom \\ using \ fw.dom.span_in_category_axioms arr_char by (metis comp_arr_dom in_homE gw.cod.span_in_category_axioms seq_char) interpret Cod_\: span_in_category C \Cod \\ using \ fw.cod.span_in_category_axioms arr_char by (metis (no_types, lifting) comp_arr_dom ideD(2) in_homI gw'.cod.span_in_category_axioms gw'.chine_is_identity hseq_gw' seqI' seq_char ide_char) interpret \: arrow_of_spans C \ using \ arr_char by auto text \ CKS say: ``Take \u\, \w\, \w'\, \\\, \\'\ as in \T2\ and note that \fw = (w\<^sub>0, W, r\<^sub>0 w\<^sub>1)\, \gw = (w\<^sub>0, W, r\<^sub>1 w\<^sub>1)\, \emph{etc}. So \\: W \ W'\ satisfies \w\<^sub>0 = w\<^sub>0' \\, \r\<^sub>1 w\<^sub>1 = r\<^sub>1 w\<^sub>1' \\. But the equation \(r\)(\w) = (r\')(\w')\\ gives \w\<^sub>1 = w\<^sub>1'\. So \\ = \ : w \ w'\ is unique with \\ = g \, \ = \' (f \).\'' Once again, there is substantial punning in the proof sketch given by CKS. We can express \fw\ and \gw\ almost in the form they indicate, but projections are required. \ have cospan: "C.cospan ?R ?w1" using hseq_char [of \ w] src_def trg_def by auto have cospan': "C.cospan ?R ?w1'" using hseq_char [of \ w'] src_def trg_def by auto have fw: "f \ w = \Chn = ?R \\ ?w1, Dom = \Leg0 = ?w0 \ ?p0, Leg1 = r0 \ ?p1\, Cod = \Leg0 = ?w0 \ ?p0, Leg1 = r0 \ ?p1\\" using ide_f hseq_char hcomp_def chine_hcomp_def fw.dom.apex_def cospan fw.chine_eq_apex by auto have gw: "g \ w = \Chn = ?R \\ ?w1, Dom = \Leg0 = ?w0 \ ?p0, Leg1 = r1 \ ?p1\, Cod = \Leg0 = ?w0 \ ?p0, Leg1 = r1 \ ?p1\\" using hseq_gw hseq_char hcomp_def chine_hcomp_def gw.dom.apex_def cospan gw.chine_eq_apex by auto have fw': "f \ w' = \Chn = ?R \\ ?w1', Dom = \Leg0 = ?w0' \ ?p0', Leg1 = r0 \ ?p1'\, Cod = \Leg0 = ?w0' \ ?p0', Leg1 = r0 \ ?p1'\\" using ide_f hseq_char hcomp_def chine_hcomp_def fw'.dom.apex_def cospan' fw'.chine_eq_apex by auto have gw': "g \ w' = \Chn = ?R \\ ?w1', Dom = \Leg0 = ?w0' \ ?p0', Leg1 = r1 \ ?p1'\, Cod = \Leg0 = ?w0' \ ?p0', Leg1 = r1 \ ?p1'\\" using hseq_gw' hseq_char hcomp_def chine_hcomp_def gw'.dom.apex_def cospan' gw'.chine_eq_apex by auto text \ Note that \?p0\ and \?p0'\ are only isomorphisms, not identities, and we have \?p1\ (which equals \?w1 \ ?p0\) and \?p1'\ (which equals \?w1' \ ?p0'\) in place of \?w1\ and \?w1'\. \ text \ The following diagram summarizes the various given and defined arrows involved in the proof. We have deviated slightly here from the nomenclature used in in CKS. We prefer to use \W\ and \W'\ to denote the apexes of \w\ and \w'\, respectively. We already have the expressions \?R \\ ?w1\ and \?R \\ ?w1'\ for the apexes of \fw\ and \fw'\ (which are the same as the apexes of \gw\ and \gw'\, respectively) and we will not use any abbreviation for them. \ text \ $$ \xymatrix{ &&& X \\ && W \ar[ur]^{w_0} \ar[dr]_{w_1} \ar@ {.>}[rr]^{\gamma} && W' \ar[ul]_{w_0'} \ar[dl]^{w_1'} && U \ar@/r10pt/[dddl]^{u_1} \ar@/u7pt/[ulll]_{u_0}\\ & R\downarrow\downarrow w_1 \ar[ur]_{p_0} \ar[dr]^{p_1} \ar@/d15pt/[rrrr]_{\beta} \ar@/u100pt/[urrrrr]^{\theta} && R && R \downarrow\downarrow w_1' \ar[ul]^{p_0'} \ar[dl]^{p_1'} \ar[ur]_{\theta'} \\ && R \ar@ {.>}[dr]_{\rho} \ar@/dl7pt/[ddr]_{R} \ar[ur]_{R} \ar[dl]_{r_1} \ar@ {<->}[rr]_{R} && R \ar[ul]^{R} \ar[dr]_{r_0} \\ & B && r_0 \downarrow\downarrow r_0 \ar[d]^{k_1} \ar[ur]_{k_0} && A \\ &&& R \ar@/dr10pt/[urr]_{r_0} \ar@/dl5pt/[ull]^{r_1} } $$ \ have Chn_\: "\\.chine: ?R \\ ?w1 \\<^sub>C ?R \\ ?w1'\" using gw gw' Chn_in_hom \ gw'.chine_eq_apex gw.chine_eq_apex by force have \_eq: "\ = \Chn = \.chine, Dom = \Leg0 = ?w0 \ ?p0, Leg1 = r1 \ ?p1\, Cod = \Leg0 = ?w0' \ ?p0', Leg1 = r1 \ ?p1'\\" using \ gw gw' dom_char cod_char by auto have Dom_\_eq: "Dom \ = \Leg0 = ?w0 \ ?p0, Leg1 = r1 \ ?p1\" using \ gw gw' dom_char cod_char by auto have Cod_\_eq: "Cod \ = \Leg0 = ?w0' \ ?p0', Leg1 = r1 \ ?p1'\" using \ gw gw' dom_char cod_char by auto have \0: "?w0 \ ?p0 = ?w0' \ ?p0' \ \.chine" using Dom_\_eq Cod_\_eq \.leg0_commutes C.comp_assoc by simp have \1: "r1 \ ?p1 = r1 \ ?p1' \ \.chine" using Dom_\_eq Cod_\_eq \.leg1_commutes C.comp_assoc by simp have Dom_\_0: "Dom_\.leg0 = ?w0 \ ?p0" using arrI dom_char fw T.uw\.\_simps(4) by auto have Cod_\_0: "Cod_\.leg0 = ?u0" using \ cod_char by auto have Dom_\_1: "Dom_\.leg1 = r0 \ ?p1" using arrI dom_char fw T.uw\.\_simps(4) by auto have Cod_\_1: "Cod_\.leg1 = ?u1" using T.uw\.\_simps(5) cod_char by auto have Dom_\'_0: "Dom_\'.leg0 = ?w0' \ ?p0'" using dom_char fw' T.uw'\'.\_simps(4) by auto have Cod_\'_0: "Cod_\'.leg0 = ?u0" using T.uw'\'.\_simps(5) cod_char by auto have Dom_\'_1: "Dom_\'.leg1 = r0 \ ?p1'" using dom_char fw' T.uw'\'.\_simps(4) by auto have Cod_\'_1: "Cod_\'.leg1 = ?u1" using T.uw'\'.\_simps(5) cod_char by auto have Dom_\_0: "Dom_\.leg0 = ?R" by simp have Dom_\_1: "Dom_\.leg1 = r1" by simp have Cod_\_0: "Cod_\.leg0 = k0" by simp have Cod_\_1: "Cod_\.leg1 = r1 \ k1" by simp have Chn_r\: "\r\.chine : rfw.chine \\<^sub>C ru.chine\" using r\.chine_composite_in_hom ru.chine_composite rfw.chine_composite Cod_\_1 Dom_\_1 fw.leg1_composite by auto have Chn_r\_eq: "r\.chine = \\

\<^sub>1[r0, r0 \ ?p1] \r0, ?u1\ \.chine \ \

\<^sub>0[r0, r0 \ ?p1]\" using r\.chine_composite Cod_\_1 Dom_\_1 fw.leg1_composite C.comp_cod_arr by (metis arrow_of_spans_data.simps(2) fw r.chine_eq_apex r.cod_simps(2) rfw.prj_simps(10) rfw.prj_simps(16) span_data.simps(2)) have r\_cod_apex_eq: "r\.cod.apex = r0 \\ ?u1" using Cod_\_1 r\.chine_composite_in_hom by auto hence r\'_cod_apex_eq: "r\'.cod.apex = r0 \\ ?u1" using Cod_\'_1 r\'.chine_composite_in_hom by auto have Chn_r\': "\r\'.chine : rfw'.chine \\<^sub>C ru.chine\" using r\'.chine_composite_in_hom ru.chine_composite rfw'.chine_composite Cod_\'_1 Dom_\'_1 fw'.leg1_composite by auto have Chn_r\'_eq: "r\'.chine = \\

\<^sub>1[r0, r0 \ ?p1'] \r0, ?u1\ \'.chine \ \

\<^sub>0[r0, r0 \ ?p1']\" using r\'.chine_composite Cod_\'_1 Dom_\'_1 fw'.leg1_composite C.comp_cod_arr by (metis arrow_of_spans_data.simps(2) fw' r.chine_eq_apex r.cod_simps(2) rfw'.prj_simps(10) rfw'.prj_simps(16) span_data.simps(2)) have Chn_\w: "\\w.chine : ?R \\ ?w1 \\<^sub>C k0 \\ ?w1\" using \w.chine_composite_in_hom by simp have Chn_\w_eq: "\w.chine = \\.chine \ ?p1 \k0, ?w1\ ?p0\" using \w.chine_composite C.comp_cod_arr ide_w by (simp add: chine_hcomp_arr_ide hcomp_def) have Chn_\w': "\\w'.chine : ?R \\ ?w1' \\<^sub>C k0 \\ ?w1'\" using \w'.chine_composite_in_hom by simp have Chn_\w'_eq: "\w'.chine = \\.chine \ ?p1' \k0, ?w1'\ ?p0'\" using \w'.chine_composite C.comp_cod_arr ide_w' Dom_\_0 Cod_\_0 by (metis \w'.composite_is_arrow chine_hcomp_arr_ide chine_hcomp_def hseq_char w'.cod_simps(3)) text \ The following are some collected commutativity properties that are used subsequently. \ have "C.commutative_square r0 ?u1 ?p1 \.chine" using ru.legs_form_cospan(1) Dom_\.is_span Dom_\_1 Cod_\_1 \.leg1_commutes apply (intro C.commutative_squareI) by auto have "C.commutative_square r0 ?u1 (?p1' \ \.chine) (\'.chine \ \.chine)" proof have 1: "r0 \ ?p1' = ?u1 \ \'.chine" using \'.leg1_commutes Cod_\'_1 Dom_\'_1 fw'.leg1_composite by simp show "C.cospan r0 ?u1" using ru.legs_form_cospan(1) by blast show "C.span (?p1' \ \.chine) (\'.chine \ \.chine)" using \.chine_in_hom \'.chine_in_hom by (metis "1" C.dom_comp C.in_homE C.prj1_simps(1) C.prj1_simps(2) C.seqI Cod_\'_1 Dom_\'.leg_simps(3) Chn_\ \'.leg1_commutes cospan') show "C.dom r0 = C.cod (?p1' \ \.chine)" using \.chine_in_hom by (metis C.cod_comp C.prj1_simps(3) \C.span (?p1' \ \.chine) (\'.chine \ \.chine)\ cospan' r.dom.apex_def r.chine_eq_apex r.chine_simps(2)) show "r0 \ ?p1' \ \.chine = ?u1 \ \'.chine \ \.chine" using 1 \.chine_in_hom C.comp_assoc by metis qed have "C.commutative_square r0 ?u1 \

\<^sub>1[r0, r0 \ ?p1] (\.chine \ \

\<^sub>0[r0, r0 \ ?p1])" using ru.legs_form_cospan(1) Dom_\.is_span Dom_\_1 C.comp_assoc C.pullback_commutes' r\.legs_form_cospan(1) apply (intro C.commutative_squareI) apply auto by (metis C.comp_assoc Cod_\_1 \.leg1_commutes) hence "C.commutative_square r0 ?u1 \

\<^sub>1[r0, r0 \ ?p1] (\.chine \ \

\<^sub>0[r0, r0 \ ?p1])" using fw.leg1_composite by auto have "C.commutative_square r0 ?u1 \

\<^sub>1[r0, r0 \ ?p1'] (\'.chine \ \

\<^sub>0[r0, r0 \ ?p1'])" using C.tuple_is_extensional Chn_r\'_eq r\'.chine_simps(1) fw' by force have "C.commutative_square ra ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0" using C.pullback_commutes' gw.legs_form_cospan(1) rfw.prj_simps(2) C.comp_assoc C.comp_cod_arr apply (intro C.commutative_squareI) by auto have "C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0" using cospan' apply (intro C.commutative_squareI) apply simp_all by (metis C.comp_assoc C.prj0_simps_arr C.pullback_commutes' arrow_of_spans_data.select_convs(2) rfw'.prj_simps(3) span_data.select_convs(1-2)) have "C.commutative_square r0 (r0 \ ?p1) rfw.Prj\<^sub>1\<^sub>1 \rfw.Prj\<^sub>0\<^sub>1 \ra, ?w1\ rfw.Prj\<^sub>0\" proof - have "C.arr rfw.chine_assoc" by (metis C.seqE rfw.prj_chine_assoc(1) rfw.prj_simps(1)) thus ?thesis using C.tuple_is_extensional rfw.chine_assoc_def by fastforce qed have "C.commutative_square r0 (r0 \ ?p1') rfw'.Prj\<^sub>1\<^sub>1 \rfw'.Prj\<^sub>0\<^sub>1 \ra, ?w1'\ rfw'.Prj\<^sub>0\" by (metis (no_types, lifting) C.not_arr_null C.seqE C.tuple_is_extensional arrow_of_spans_data.select_convs(2) rfw'.chine_assoc_def rfw'.prj_chine_assoc(1) rfw'.prj_simps(1) span_data.select_convs(1-2)) have "C.commutative_square k0 ?w1 (\.chine \ ?p1) ?p0" using C.tuple_is_extensional Chn_\w_eq \w.chine_simps(1) by fastforce have "C.commutative_square k0 ?w1' (\.chine \ ?p1') (w'.chine \ ?p0')" using C.tuple_is_extensional \w'.chine_composite \w'.chine_simps(1) by force have "C.commutative_square k0 ?w1' (\.chine \ ?p1') ?p0'" using C.tuple_is_extensional Chn_\w'_eq \w'.chine_simps(1) by force text \ Now, derive the consequences of the equation: \[ \(r \ \) \ \[r, ?f, w] \ (?\ \ w) = (r \ \') \ \[r, ?f, w'] \ (?\ \ w') \ \\ \] The strategy is to expand and simplify the left and right hand side to tuple form, then compose with projections and equate corresponding components. We first work on the right-hand side. \ have R: "Chn (T.composite_cell w' \' \ \) = \?p1' \ \.chine \r0, ?u1\ \'.chine \ \.chine\" proof - have "Chn (T.composite_cell w' \' \ \) = r\'.chine \ Chn \[r, f, w'] \ \w'.chine \ \.chine" proof - have 1: "\T.composite_cell w' \' \ \ : g \ w \ r \ u\" using \ \w' arfw' r\' by auto have "Chn (T.composite_cell w' \' \ \) = Chn (T.composite_cell w' \') \ \.chine" using 1 Chn_vcomp by blast also have "... = (r\'.chine \ Chn (\[r, f, w'] \ (\ \ w'))) \ \.chine" proof - have "seq (r \ \') (\[r, f, w'] \ (\ \ w'))" using 1 by blast thus ?thesis using 1 Chn_vcomp by presburger qed also have "... = (r\'.chine \ Chn \[r, f, w'] \ \w'.chine) \ \.chine" proof - have "seq \[r, f, w'] (\ \ w')" using 1 by blast thus ?thesis using 1 Chn_vcomp by presburger qed finally show ?thesis using C.comp_assoc by auto qed also have "... = \?p1' \ \.chine \r0, ?u1\ \'.chine \ \.chine\" proof - let ?LHS = "r\'.chine \ Chn \[r, f, w'] \ \w'.chine \ \.chine" let ?RHS = "\?p1' \ \.chine \r0, ?u1\ \'.chine \ \.chine\" have LHS: "\?LHS : ?R \\ ?w1 \\<^sub>C r\'.cod.apex\" proof (intro C.comp_in_homI) show "\\.chine : ?R \\ ?w1 \\<^sub>C ?R \\ ?w1'\" using Chn_\ by simp show "\\w'.chine : ?R \\ ?w1' \\<^sub>C Cod_\.leg0 \\ w'.cod.leg1\" using Chn_\w' by simp show "\Chn \[r, f, w'] : Cod_\.leg0 \\ w'.cod.leg1 \\<^sub>C rfw'.chine\" using arfw' by (metis (no_types, lifting) Chn_in_hom Cod_\_0 arrow_of_spans_data.simps(2) rf rf.leg0_composite rfw'.chine_composite(1) span_data.select_convs(1) w'.cod_simps(3)) show "\r\'.chine : rfw'.chine \\<^sub>C r\'.cod.apex\" using Chn_r\' by auto qed have 2: "C.commutative_square r0 ?u1 (?p1' \ \.chine) (\'.chine \ \.chine)" by fact have RHS: "\?RHS : ?R \\ ?w1 \\<^sub>C r\'.cod.apex\" using 2 Chn_\ r\'_cod_apex_eq C.tuple_in_hom [of r0 ?u1 "?p1' \ \.chine" "\'.chine \ \.chine"] by fastforce show ?thesis proof (intro C.prj_joint_monic [of r0 ?u1 ?LHS ?RHS]) show "C.cospan r0 ?u1" using ru.legs_form_cospan(1) by blast show "C.seq ru.prj\<^sub>1 ?LHS" using LHS r\'_cod_apex_eq by auto show "C.seq ru.prj\<^sub>1 ?RHS" using RHS r\'_cod_apex_eq by auto show "ru.prj\<^sub>0 \ ?LHS = ru.prj\<^sub>0 \ ?RHS" proof - have "ru.prj\<^sub>0 \ ?LHS = (ru.prj\<^sub>0 \ r\'.chine) \ Chn \[r, f, w'] \ \w'.chine \ \.chine" using C.comp_assoc by simp also have "... = ((\'.chine \ \

\<^sub>0[r0, r0 \ ?p1']) \ Chn \[r, f, w']) \ \w'.chine \ \.chine" using Chn_r\'_eq C.comp_assoc fw' \C.commutative_square r0 ?u1 \

\<^sub>1[r0, r0 \ ?p1'] (\'.chine \ \

\<^sub>0[r0, r0 \ ?p1'])\ by simp also have "... = \'.chine \ (\

\<^sub>0[r0, r0 \ ?p1'] \ Chn \[r, f, w']) \ \w'.chine \ \.chine" using C.comp_assoc by simp also have "... = \'.chine \ (\rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine) \ \.chine" using ide_f hseq_rf hseq_char \_ide C.comp_assoc rfw'.chine_assoc_def fw'.leg1_composite C.prj_tuple(1) \C.commutative_square r0 (r0 \ ?p1') rfw'.Prj\<^sub>1\<^sub>1 \rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\\ by simp also have "... = \'.chine \ \.chine" proof - have "\rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine = gw'.apex" proof (intro C.prj_joint_monic [of ?R ?w1' "\rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine" gw'.apex]) show "C.cospan ?R ?w1'" using fw'.legs_form_cospan(1) by simp show "C.seq ?p1' (\rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine)" proof (intro C.seqI' C.comp_in_homI) show "\\w'.chine : Dom_\.leg0 \\ w'.leg1 \\<^sub>C Cod_\.leg0 \\ w'.cod.leg1\" using \w'.chine_composite_in_hom by simp show "\\rfw'.Prj\<^sub>0\<^sub>1 \?R, w'.leg1\ rfw'.Prj\<^sub>0\ : Cod_\.leg0 \\ w'.cod.leg1 \\<^sub>C ?R \\ w'.leg1\" using \C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\ C.tuple_in_hom [of ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0] rf rf.leg0_composite by auto show "\?p1' : ?R \\ w'.leg1 \\<^sub>C f.apex\" using fw'.prj_in_hom(1) by auto qed show "C.seq ?p1' gw'.apex" using gw'.dom.apex_def gw'.leg0_composite fw'.prj_in_hom by auto show "?p0' \ \rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine = ?p0' \ gw'.apex" proof - have "?p0' \ \rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine = (?p0' \ \rfw'.Prj\<^sub>0\<^sub>1 \?R, ?w1'\ rfw'.Prj\<^sub>0\) \ \w'.chine" using C.comp_assoc by simp also have "... = rfw'.Prj\<^sub>0 \ \w'.chine" using \C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\ by auto also have "... = \

\<^sub>0[k0, ?w1'] \ \\.chine \ ?p1' \k0, ?w1'\ w'.chine \ ?p0'\" using \w'.chine_composite Dom_\_0 Cod_\_0 C.comp_cod_arr by simp also have "... = w'.chine \ ?p0'" using \C.commutative_square k0 ?w1' (\.chine \ ?p1') (w'.chine \ ?p0')\ by simp also have "... = ?p0' \ gw'.apex" using cospan C.comp_cod_arr C.comp_arr_dom gw'.chine_is_identity gw'.chine_eq_apex gw'.chine_composite fw'.prj_in_hom by auto finally show ?thesis by simp qed show "?p1' \ \rfw'.Prj\<^sub>0\<^sub>1 \ra, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine = ?p1' \ gw'.apex" proof - have "?p1' \ \rfw'.Prj\<^sub>0\<^sub>1 \ra, ?w1'\ rfw'.Prj\<^sub>0\ \ \w'.chine = (?p1' \ \rfw'.Prj\<^sub>0\<^sub>1 \ra, ?w1'\ rfw'.Prj\<^sub>0\) \ \w'.chine" using C.comp_assoc by simp also have "... = rfw'.Prj\<^sub>0\<^sub>1 \ \w'.chine" using \C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\ by simp also have "... = k0 \ \

\<^sub>1[k0, ?w1'] \ \\.chine \ ?p1' \k0, ?w1'\ w'.chine \ ?p0'\" using \w'.chine_composite Cod_\_0 C.comp_assoc C.comp_cod_arr by simp also have "... = k0 \ \.chine \ ?p1'" using \C.commutative_square k0 ?w1' (\.chine \ ?p1') (w'.chine \ ?p0')\ by simp also have "... = (k0 \ \.chine) \ ?p1'" using C.comp_assoc by metis also have "... = ?p1'" using \.leg0_commutes C.comp_cod_arr cospan' by simp also have "... = ?p1' \ gw'.apex" using C.comp_arr_dom cospan' gw'.chine_eq_apex gw'.chine_composite by simp finally show ?thesis by simp qed qed thus ?thesis using Chn_\ C.comp_cod_arr gw'.apex_composite by auto qed also have "... = \

\<^sub>0[r0, ?u1] \ ?RHS" using RHS 2 C.prj_tuple [of r0 ?u1] by simp finally show ?thesis by simp qed show "ru.prj\<^sub>1 \ ?LHS = ru.prj\<^sub>1 \ ?RHS" proof - have "ru.prj\<^sub>1 \ ?LHS = (ru.prj\<^sub>1 \ r\'.chine) \ Chn \[r, f, w'] \ \w'.chine \ \.chine" using C.comp_assoc by simp also have "... = \

\<^sub>1[r0, fw'.leg1] \ Chn \[r, f, w'] \ \w'.chine \ \.chine" using Chn_r\' Chn_r\'_eq fw' \C.commutative_square r0 ?u1 \

\<^sub>1[r0, r0 \ ?p1'] (\'.chine \ \

\<^sub>0[r0, r0 \ ?p1'])\ by simp also have "... = (rfw'.Prj\<^sub>1 \ rfw'.chine_assoc) \ \w'.chine \ \.chine" using ide_f ide_w' hseq_rf hseq_char \_ide fw'.leg1_composite C.comp_assoc by auto also have "... = (rfw'.Prj\<^sub>1\<^sub>1 \ \w'.chine) \ \.chine" using rfw'.prj_chine_assoc C.comp_assoc by simp also have "... = ((k1 \ \

\<^sub>1[k0, ?w1']) \ \w'.chine) \ \.chine" using C.comp_cod_arr by simp also have "... = (k1 \ \

\<^sub>1[k0, ?w1'] \ \w'.chine) \ \.chine" using C.comp_assoc by simp also have "... = (k1 \ \.chine \ ?p1') \ \.chine" using Chn_\w'_eq Dom_\_0 Cod_\_0 \C.commutative_square k0 ?w1' (\.chine \ ?p1') ?p0'\ by simp also have "... = (k1 \ \.chine) \ ?p1' \ \.chine" using C.comp_assoc by metis also have "... = (?R \ ?p1') \ \.chine" using C.comp_assoc by simp also have "... = ?p1' \ \.chine" using C.comp_cod_arr C.prj1_in_hom [of ?R ?w1'] cospan' by simp also have "... = ru.prj\<^sub>1 \ ?RHS" using RHS 2 by simp finally show ?thesis by simp qed qed qed finally show ?thesis by simp qed text \ Now we work on the left-hand side. \ have L: "Chn (T.composite_cell w \) = \?p1 \r0, ?u1\ \.chine\" proof - have "Chn (T.composite_cell w \) = r\.chine \ Chn \[r, f, w] \ \w.chine" using Chn_vcomp arfw C.comp_assoc by auto moreover have "... = \?p1 \r0, ?u1\ \.chine\" proof - let ?LHS = "r\.chine \ Chn \[r, f, w] \ \w.chine" let ?RHS = "\?p1 \r0, ?u1\ \.chine\" have 2: "C.commutative_square r0 ?u1 ?p1 \.chine" by fact have LHS: "\?LHS : ?R \\ ?w1 \\<^sub>C r0 \\ ?u1\" using Chn_r\ Chn_\w rfw.chine_assoc_in_hom by (metis (no_types, lifting) "8" Chn_in_hom Dom_\_0 arrow_of_spans_data.simps(2) calculation gw.chine_composite r\_cod_apex_eq ru.chine_composite) have RHS: "\?RHS : ?R \\ ?w1 \\<^sub>C r0 \\ ?u1\" using 2 C.tuple_in_hom [of r0 ?u1 "?p1" \.chine] cospan r\_cod_apex_eq by simp show ?thesis proof (intro C.prj_joint_monic [of r0 ?u1 ?LHS ?RHS]) show "C.cospan r0 ?u1" using ru.legs_form_cospan(1) by blast show "C.seq ru.prj\<^sub>1 ?LHS" using LHS r\_cod_apex_eq by auto show "C.seq ru.prj\<^sub>1 ?RHS" using RHS r\_cod_apex_eq by auto show "ru.prj\<^sub>0 \ ?LHS = ru.prj\<^sub>0 \ ?RHS" proof - have "ru.prj\<^sub>0 \ ?LHS = (ru.prj\<^sub>0 \ r\.chine) \ Chn \[r, f, w] \ \w.chine" using C.comp_assoc by simp also have "... = (\.chine \ \

\<^sub>0[r0, f.leg1 \ fw.prj\<^sub>1]) \ Chn \[r, f, w] \ \w.chine" using Chn_r\_eq Dom_\_1 Cod_\_1 fw.leg1_composite \C.commutative_square r0 ?u1 \

\<^sub>1[r0, r0 \ ?p1] (\.chine \ \

\<^sub>0[r0, r0 \ ?p1])\ by simp also have "... = \.chine \ (\

\<^sub>0[r0, r0 \ ?p1] \ Chn \[r, f, w]) \ \w.chine" using C.comp_assoc by simp also have "... = \.chine \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine" proof - have "Chn \[r, f, w] = rfw.chine_assoc" using ide_f ide_w hseq_rf hseq_char \_ide by auto moreover have "\

\<^sub>0[r0, r0 \ ?p1] \ rfw.chine_assoc = \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\" using rfw.chine_assoc_def \C.commutative_square r0 (r0 \ ?p1) rfw.Prj\<^sub>1\<^sub>1 \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\\ by simp ultimately show ?thesis by simp qed also have "... = \.chine \ (?R \\ ?w1)" proof - have "\rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine = ?R \\ ?w1" proof (intro C.prj_joint_monic [of ?R ?w1 "\rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine" "?R \\ ?w1"]) show "C.cospan ?R ?w1" by fact show "C.seq ?p1 (\rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine)" proof - have "C.seq rfw.Prj\<^sub>0\<^sub>1 \w.chine" by (meson C.seqI' Chn_in_hom \w rfw.prj_in_hom(2) \C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\) thus ?thesis using \C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\ by (metis (no_types) C.comp_assoc C.prj_tuple(2)) qed show "C.seq ?p1 (?R \\ ?w1)" using gw.dom.apex_def gw.leg0_composite gw.prj_in_hom by auto show "?p0 \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine = ?p0 \ (?R \\ ?w1)" proof - have "?p0 \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine = (?p0 \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\) \ \w.chine" using C.comp_assoc by simp also have "... = rfw.Prj\<^sub>0 \ \w.chine" using \C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\ by simp also have "... = \

\<^sub>0[k0, ?w1] \ \\.chine \ ?p1 \k0, ?w1\ ?p0\" using Chn_\w_eq C.comp_cod_arr by simp also have "... = ?p0" using \C.commutative_square k0 ?w1 (\.chine \ ?p1) ?p0\ C.prj_tuple(1) by blast also have "... = ?p0 \ (?R \\ ?w1)" using C.comp_arr_dom gw.chine_eq_apex gw.chine_is_identity by (metis C.arr_dom_iff_arr C.pbdom_def Dom_g gw.chine_composite gw.chine_simps(1) span_data.select_convs(1)) finally show ?thesis by simp qed show "?p1 \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine = ?p1 \ (?R \\ ?w1)" proof - have "?p1 \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\ \ \w.chine = (?p1 \ \rfw.Prj\<^sub>0\<^sub>1 \?R, ?w1\ rfw.Prj\<^sub>0\) \ \w.chine" using C.comp_assoc by simp also have "... = rfw.Prj\<^sub>0\<^sub>1 \ \w.chine" using \C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\ by simp also have "... = (k0 \ \

\<^sub>1[k0, ?w1]) \ \\.chine \ ?p1 \k0, ?w1\ ?p0\" using Chn_\w_eq C.comp_cod_arr by simp also have "... = k0 \ \

\<^sub>1[k0, ?w1] \ \\.chine \ ?p1 \k0, ?w1\ ?p0\" using C.comp_assoc by simp also have "... = k0 \ \.chine \ ?p1" using \C.commutative_square k0 ?w1 (\.chine \ ?p1) ?p0\ by simp also have "... = (k0 \ \.chine) \ ?p1" using C.comp_assoc by metis also have "... = ?p1 \ (?R \\ ?w1)" using C.comp_arr_dom C.comp_cod_arr cospan by simp finally show ?thesis by blast qed qed thus ?thesis by simp qed also have "... = \.chine" using C.comp_arr_dom \.chine_in_hom gw.chine_eq_apex gw.chine_is_identity Dom_\_0 Cod_\_0 Dom_\.apex_def Cod_\.apex_def by (metis Dom_g \.chine_simps(1) \.chine_simps(2) gw.chine_composite gw.dom.apex_def gw.leg0_composite span_data.select_convs(1)) also have "... = ru.prj\<^sub>0 \ ?RHS" using 2 by simp finally show ?thesis by blast qed show "ru.prj\<^sub>1 \ ?LHS = ru.prj\<^sub>1 \ ?RHS" proof - have "ru.prj\<^sub>1 \ ?LHS = (ru.prj\<^sub>1 \ r\.chine) \ Chn \[r, f, w] \ \w.chine" using C.comp_assoc by simp also have "... = (r.chine \ \

\<^sub>1[r0, r0 \ ?p1]) \ Chn \[r, f, w] \ \w.chine" proof - have "r\.chine \ C.null \ \

\<^sub>1[r.cod.leg0, Cod_\.leg1] \ r\.chine = r.chine \ \

\<^sub>1[r0, Dom_\.leg1]" by (metis (lifting) C.prj_tuple(2) C.tuple_is_extensional r.cod_simps(2) r\.chine_composite) thus ?thesis using Cod_\_1 Dom_\_1 r\.chine_simps(1) fw by fastforce qed also have "... = r.chine \ (rfw.Prj\<^sub>1 \ Chn \[r, f, w]) \ \w.chine" using C.comp_assoc fw.leg1_composite by simp also have "... = r.chine \ rfw.Prj\<^sub>1\<^sub>1 \ \w.chine" using ide_f ide_w hseq_rf hseq_char \_ide rfw.prj_chine_assoc(1) by auto also have "... = r.chine \ k1 \ \

\<^sub>1[k0, ?w1] \ \w.chine" using C.comp_cod_arr C.comp_assoc by simp also have "... = r.chine \ k1 \ \.chine \ \

\<^sub>1[Dom_\.leg0, ?w1]" using Chn_\w_eq \C.commutative_square k0 ?w1 (\.chine \ \

\<^sub>1[ra, w.leg1]) \

\<^sub>0[ra, w.leg1]\ by auto also have "... = r.chine \ (k1 \ \.chine) \ ?p1" using C.comp_assoc Dom_\_0 by metis also have "... = r.chine \ ra \ ?p1" by simp also have "... = r.chine \ ?p1" using C.comp_cod_arr by (metis C.comp_assoc r.cod_simps(1) r.chine_eq_apex r.chine_simps(1) r.chine_simps(3)) also have "... = ?p1" using C.comp_cod_arr r.chine_eq_apex r.chine_is_identity by (metis 2 C.commutative_squareE r.dom.apex_def) also have "... = ru.prj\<^sub>1 \ ?RHS" using 2 by simp finally show ?thesis by simp qed qed qed ultimately show ?thesis by simp qed text \ This is the main point: the equation E boils down to the following: \[ \?p1' \ \.chine = ?p1 \ \'.chine \ \.chine = \.chine\ \] The first equation gets us close to what we need, but we still need \?p1 \ C.inv ?p0 = ?w1\, which follows from the fact that ?p0 is the pullback of ?R. \ have *: "\?p1' \ \.chine \r0, ?u1\ \'.chine \ \.chine\ = \?p1 \r0, ?u1\ \.chine\" using L R E by simp have **: "?p1' \ \.chine = ?p1" by (metis "*" C.in_homE C.not_arr_null C.prj_tuple(2) C.tuple_in_hom C.tuple_is_extensional \C.commutative_square r0 u.leg1 (\

\<^sub>1[ra, w'.leg1] \ \.chine) (\'.chine \ \.chine)\) have ***: "\'.chine \ \.chine = \.chine" by (metis "*" C.prj_tuple(1) \C.commutative_square r0 ?u1 (?p1' \ \.chine) (\'.chine \ \.chine)\ \C.commutative_square r0 ?u1 ?p1 \.chine\) text \ CKS say to take \\ = \\, but obviously this cannot work as literally described, because \\\ : g \ w \ g \ w'\\, whereas we must have \\\ : w \ w'\\. Instead, we have to define \\\ by transporting \\\ along the projections from \?R \\ ?w1\ to \?W\ and \?R \\ ?w1'\ to \?W'\. These are isomorphisms by virtue of their being pullbacks of identities, but they are not themselves necessarily identities. Specifically, we take \Chn \ = ?p0' \ Chn \ \ C.inv ?p0\. \ let ?\ = "\Chn = ?p0' \ \.chine \ C.inv ?p0, Dom = Dom w, Cod = Cod w'\" interpret Dom_\: span_in_category C \Dom ?\\ using w.dom.span_in_category_axioms by simp interpret Cod_\: span_in_category C \Cod ?\\ using w'.cod.span_in_category_axioms by simp text \ It has to be shown that \\\ is an arrow of spans. \ interpret \: arrow_of_spans C ?\ proof show "\Chn ?\ : Dom_\.apex \\<^sub>C Cod_\.apex\" proof - have "\Chn \: gw.apex \\<^sub>C gw'.apex\" using Chn_in_hom \ gw'.chine_eq_apex gw.chine_eq_apex by force moreover have "\?p0' : gw'.apex \\<^sub>C w'.apex\" using cospan' hseq_gw' hseq_char hcomp_def gw'.dom.apex_def w'.dom.apex_def by auto moreover have "\C.inv ?p0 : w.apex \\<^sub>C gw.apex\" using cospan hseq_gw hseq_char hcomp_def gw.dom.apex_def w.dom.apex_def C.iso_pullback_ide by auto ultimately show ?thesis using Dom_\.apex_def Cod_\.apex_def by auto qed text \ The commutativity property for the ``input leg'' follows directly from that for \\\. \ show "Cod_\.leg0 \ Chn ?\ = Dom_\.leg0" using C.comp_assoc C.comp_arr_dom cospan C.iso_pullback_ide C.comp_arr_inv' by (metis C.invert_side_of_triangle(2) Dom_\.leg_simps(1) Dom_\_eq \0 arrow_of_spans_data.select_convs(1,3) arrow_of_spans_data.simps(2) r.dom.ide_apex span_data.select_convs(1) w'.cod_simps(2)) text \ The commutativity property for the ``output leg'' is a bit more subtle. \ show "Cod_\.leg1 \ Chn ?\ = Dom_\.leg1" proof - have "Cod_\.leg1 \ Chn ?\ = ((?w1' \ ?p0') \ \.chine) \ C.inv ?p0" using C.comp_assoc by simp also have "... = ((?R \ ?p1') \ Chn \) \ C.inv ?p0" using cospan' C.pullback_commutes [of ?R ?w1'] by auto also have "... = (?p1' \ \.chine) \ C.inv ?p0" using cospan' C.comp_cod_arr by simp also have "... = ?p1 \ C.inv ?p0" using ** by simp also have "... = ?w1" text \ Sledgehammer found this at a time when I was still struggling to understand what was going on. \ by (metis C.comp_cod_arr C.invert_side_of_triangle(2) C.iso_pullback_ide C.prj1_simps(1,3) C.pullback_commutes' cospan r.dom.ide_apex r.chine_eq_apex r.chine_simps(2)) also have "... = Dom_\.leg1" by auto finally show ?thesis by simp qed qed text \ What remains to be shown is that \\\ is unique with the properties asserted by \T2\; \emph{i.e.} \\\ : w \ w'\ \ \ = g \ \ \ \ = \' \ (f \ \)\. CKS' assertion that the equation \(r\)(\w) = (r\')(\w')\\ gives \w\<^sub>1 = w\<^sub>1'\ does not really seem to be true. The reason \\\ is unique is because it is obtained by transporting \\\ along isomorphisms. \ have \: "\?\ : w \ w'\" using \.arrow_of_spans_axioms arr_char dom_char cod_char by auto have hseq_f\: "hseq f ?\" using \ src_def trg_def arrI fw.composable rf.are_arrows(2) by auto have hseq_g\: "hseq g ?\" using \ src_def trg_def fw.composable gw.are_arrows(1) src_f by auto interpret f\: two_composable_arrows_of_spans C prj0 prj1 f ?\ using hseq_f\ hseq_char by (unfold_locales, simp) interpret f\: arrow_of_spans C \f \ ?\\ using f\.composite_is_arrow arr_char by simp interpret g\: two_composable_arrows_of_spans C prj0 prj1 g ?\ using hseq_g\ hseq_char by (unfold_locales, simp) interpret g\: arrow_of_spans C \g \ ?\\ using g\.composite_is_arrow arr_char by simp have Chn_g\: "Chn (g \ ?\) = \?p1 \?R, ?w1'\ ?p0' \ \.chine\" proof - have "Chn (g \ ?\) = \?R \ ?p1 \?R, ?w1'\ (?p0' \ \.chine \ C.inv ?p0) \ ?p0\" using g\.chine_composite by simp also have "... = \?p1 \?R, ?w1'\ (?p0' \ \.chine \ C.inv ?p0) \ ?p0\" using C.comp_cod_arr cospan by simp also have "... = \?p1 \?R, ?w1'\ ?p0' \ \.chine\" proof - have "(?p0' \ \.chine \ C.inv ?p0) \ ?p0 = ?p0' \ \.chine" using C.comp_assoc C.iso_pullback_ide [of ?R ?w1] C.comp_inv_arr C.comp_arr_dom Chn_\ by (metis C.comp_inv_arr' C.in_homE C.pbdom_def cospan r.dom.ide_apex) thus ?thesis by simp qed ultimately show ?thesis by simp qed have Chn_\_eq: "\.chine = Chn (g \ ?\)" proof - have "Chn (g \ ?\) = \?p1 \?R, ?w1'\ ?p0' \ Chn \\" using Chn_g\ by simp also have "... = \.chine" text \Here was another score by sledgehammer while I was still trying to understand it.\ using ** C.prj_joint_monic by (metis C.prj1_simps(1) C.tuple_prj cospan cospan') finally show ?thesis by simp qed have \_eq_g\: "\ = g \ ?\" proof (intro arr_eqI) show "par \ (g \ ?\)" proof - have "\g \ ?\ : g \ w \ g \ w'\" using ide_g \ T.leg1_simps(3) by (intro hcomp_in_vhom, auto) thus ?thesis using \ by (elim in_homE, auto) qed show "\.chine = Chn (g \ ?\)" using Chn_\_eq by simp qed moreover have "\ = \' \ (f \ ?\)" proof (intro arr_eqI) have f\: "\f \ ?\ : f \ w \ f \ w'\" using \ ide_f by auto show par: "par \ (\' \ (f \ ?\))" using \ \' f\ by (elim in_homE, auto) show "\.chine = Chn (\' \ (f \ ?\))" using par "***" Chn_vcomp calculation f\.chine_composite g\.chine_composite by auto qed ultimately show 2: "\?\ : w \ w'\ \ \ = g \ ?\ \ \ = \' \ (f \ ?\)" using \ by simp show "\\'. \\' : w \ w'\ \ \ = g \ \' \ \ = \' \ (f \ \') \ \' = ?\" proof - fix \' assume 1: "\\' : w \ w'\ \ \ = g \ \' \ \ = \' \ (f \ \')" interpret \': arrow_of_spans C \' using 1 arr_char by auto have hseq_g\': \hseq g \'\ using 1 \ by auto interpret g\': two_composable_arrows_of_spans C prj0 prj1 g \' using hseq_g\' hseq_char by unfold_locales auto interpret g\': arrow_of_spans C \g \ \'\ using g\'.composite_is_arrow arr_char by simp show "\' = ?\" proof (intro arr_eqI) show par: "par \' ?\" using 1 \ by fastforce show "\'.chine = \.chine" proof - have "C.commutative_square ?R ?w1' (g.chine \ ?p1) (\'.chine \ ?p0)" proof show "C.cospan ?R ?w1'" by fact show 3: "C.span (g.chine \ ?p1) (\'.chine \ ?p0)" proof (intro conjI) show "C.seq g.chine ?p1" using cospan by auto show "C.seq \'.chine ?p0" using cospan 2 par arrow_of_spans_data.simps(1) dom_char in_homE w.chine_eq_apex by auto thus "C.dom (g.chine \ ?p1) = C.dom (\'.chine \ ?p0)" using g.chine_eq_apex cospan by simp qed show "C.dom ra = C.cod (g.chine \ ?p1)" using cospan by auto show "?R \ g.chine \ ?p1 = ?w1' \ \'.chine \ ?p0" proof - have "?w1' \ \'.chine \ ?p0 = (?w1' \ \'.chine) \ ?p0" using C.comp_assoc by simp moreover have "... = ?w1 \ ?p0" using 1 \'.leg1_commutes dom_char cod_char by auto also have "... = ?R \ ?p1" using cospan C.pullback_commutes [of ra ?w1] by auto also have "... = ?R \ g.chine \ ?p1" using 3 C.comp_cod_arr g.chine_is_identity g.chine_eq_apex g.dom.apex_def by auto finally show ?thesis by auto qed qed have "C.commutative_square ?R ?w1' (g.chine \ ?p1) (\.chine \ ?p0)" proof show "C.cospan ?R ?w1'" by fact show 3: "C.span (g.chine \ ?p1) (\.chine \ ?p0)" using cospan \.chine_in_hom by auto show "C.dom ?R = C.cod (g.chine \ ?p1)" using cospan by auto show "?R \ g.chine \ ?p1 = ?w1' \ \.chine \ ?p0" proof - have "?w1' \ \.chine \ ?p0 = (?w1' \ \.chine) \ ?p0" using C.comp_assoc by simp moreover have "... = ?w1 \ ?p0" using 1 \.leg1_commutes dom_char cod_char by auto also have "... = ?R \ ?p1" using cospan C.pullback_commutes [of ra ?w1] by auto also have "... = ?R \ g.chine \ ?p1" using 3 C.comp_cod_arr g.chine_is_identity g.chine_eq_apex g.dom.apex_def by auto finally show ?thesis by auto qed qed have "\'.chine \ ?p0 = \.chine \ ?p0" proof - have "\'.chine \ ?p0 = ?p0' \ g\'.chine" using 1 dom_char cod_char g\'.chine_composite \C.commutative_square ?R ?w1' (g.chine \ ?p1) (\'.chine \ ?p0)\ by auto also have "... = ?p0' \ \.chine" using 1 by simp also have "... = ?p0' \ g\.chine" using Chn_\_eq by simp also have "... = \.chine \ ?p0" using g\.chine_composite \C.commutative_square ?R ?w1' (g.chine \ ?p1) (\.chine \ ?p0)\ by simp finally show ?thesis by simp qed thus ?thesis using C.iso_pullback_ide C.iso_is_retraction C.retraction_is_epi C.epiE [of "?p0" \'.chine \.chine] cospan \.chine_in_hom \'.chine_in_hom by auto qed qed qed qed qed qed qed qed end context span_bicategory begin interpretation chosen_right_adjoints vcomp hcomp assoc unit src trg .. notation some_right_adjoint ("_\<^sup>*" [1000] 1000) (* TODO: Why is this needed? *) notation isomorphic (infix "\" 50) text \ \Span(C)\ is a bicategory of spans. \ lemma is_bicategory_of_spans: shows "bicategory_of_spans vcomp hcomp assoc unit src trg" proof text \ Every 1-cell \r\ is isomorphic to the composition of a map and the right adjoint of a map. The proof is to obtain a tabulation of \r\ as a span of maps \(f, g)\ and then observe that \r\ is isomorphic to \g \ f\<^sup>*\. \ show "\r. ide r \ \f g. is_left_adjoint f \ is_left_adjoint g \ r \ g \ f\<^sup>*" proof - fix r assume r: "ide r" interpret r: identity_arrow_of_spans C r using r ide_char' by auto interpret r: identity_arrow_in_span_bicategory C prj0 prj1 r .. have \: "tabulation (\) (\) assoc unit src trg r r.\ r.f r.g \ is_left_adjoint r.f \ is_left_adjoint r.g" using r r.has_tabulation by blast interpret \: tabulation vcomp hcomp assoc unit src trg r r.\ r.f r.g using \ by fast have 1: "r \ r.g \ r.f\<^sup>*" using \ \.yields_isomorphic_representation' \.T0.is_map left_adjoint_extends_to_adjoint_pair isomorphic_def [of "r.g \ r.f\<^sup>*" r] isomorphic_symmetric by auto thus "\f g. is_left_adjoint f \ is_left_adjoint g \ r \ g \ f\<^sup>*" using \ by blast qed text \ Every span of maps extends to a tabulation. \ show "\f g. \ is_left_adjoint f; is_left_adjoint g; src f = src g \ \ \r \. tabulation (\) (\) assoc unit src trg r \ f g" proof - text \ The proof idea is as follows: Let maps \f = (f\<^sub>1, f\<^sub>0)\ and \g = (g\<^sub>1, g\<^sub>0)\ be given. Let \f' = (f\<^sub>1 \ C.inv f\<^sub>0, C.cod f\<^sub>0)\ and \g' = (g\<^sub>1 \ C.inv g\<^sub>0, C.cod g\<^sub>0)\; then \f'\ and \g'\ are maps isomorphic to \f\ and \g\, respectively. By a previous result, \f'\ and \g'\ extend to a tabulation \(f', \, g')\ of \r = (f\<^sub>1 \ C.inv f\<^sub>0, g\<^sub>1 \ C.inv g\<^sub>0)\. Compose with isomorphisms \\\ : f' \ f\\ and \\\ : g \ g'\\ to obtain \(f, (r \ \) \ \ \ \, g)\ and show it must also be a tabulation. \ fix f g assume f: "is_left_adjoint f" assume g: "is_left_adjoint g" assume fg: "src f = src g" show "\r \. tabulation (\) (\) assoc unit src trg r \ f g" proof - text \We have to unpack the hypotheses to get information about f and g.\ obtain f\<^sub>a \\<^sub>f \\<^sub>f where ff\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \\<^sub>f \\<^sub>f" using f adjoint_pair_def by auto interpret ff\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \\<^sub>f \\<^sub>f using ff\<^sub>a by simp interpret f: arrow_of_spans C f using ide_char [of f] by simp interpret f: identity_arrow_of_spans C f using ide_char [of f] by unfold_locales auto obtain g\<^sub>a \\<^sub>g \\<^sub>g where G: "adjunction_in_bicategory vcomp hcomp assoc unit src trg g g\<^sub>a \\<^sub>g \\<^sub>g" using g adjoint_pair_def by auto interpret gg\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg g g\<^sub>a \\<^sub>g \\<^sub>g using G by simp interpret g: arrow_of_spans C g using ide_char [of g] by simp interpret g: identity_arrow_of_spans C g using ide_char [of g] by unfold_locales auto let ?f' = "mkIde (C.cod f.leg0) (f.dom.leg1 \ C.inv f.leg0)" have f': "ide ?f'" proof - have "C.span (C.cod f.leg0) (f.leg1 \ C.inv f.leg0)" using f is_left_adjoint_char by auto thus ?thesis using ide_mkIde by blast qed interpret f': arrow_of_spans C ?f' using f' ide_char by blast interpret f': identity_arrow_of_spans C ?f' using f' ide_char by unfold_locales auto let ?g' = "mkIde (C.cod g.leg0) (g.dom.leg1 \ C.inv g.leg0)" have g': "ide ?g'" proof - have "C.span (C.cod g.leg0) (g.leg1 \ C.inv g.leg0)" using g is_left_adjoint_char by auto thus ?thesis using ide_mkIde by blast qed interpret g': arrow_of_spans C ?g' using g' ide_char by blast interpret g': identity_arrow_of_spans C ?g' using g' ide_char by unfold_locales auto let ?r = "mkIde (f'.leg1) (g'.leg1)" have r: "ide ?r" proof - have "C.span (f'.leg1) (g'.leg1)" using f g fg src_def is_left_adjoint_char by simp thus ?thesis using ide_mkIde by blast qed interpret r: arrow_of_spans C ?r using r ide_char by blast interpret r: identity_arrow_of_spans C ?r using r ide_char by unfold_locales auto interpret r: identity_arrow_in_span_bicategory C prj0 prj1 ?r .. have "r.f = ?f'" using f r.chine_eq_apex is_left_adjoint_char by auto have "r.g = ?g'" using f r.chine_eq_apex fg src_def is_left_adjoint_char by simp interpret \: tabulation \(\)\ \(\)\ assoc unit src trg ?r r.\ r.f r.g using r.has_tabulation by simp have \_eq: "r.\ = \Chn = \C.cod f.leg0 \f'.leg1, f'.leg1\ C.cod f.leg0\, Dom = \Leg0 = C.cod f.leg0, Leg1 = g'.leg1\, Cod = \Leg0 = \

\<^sub>0[f'.leg1, f'.leg1], Leg1 = g'.leg1 \ \

\<^sub>1[f'.leg1, f'.leg1]\\" using \r.f = ?f'\ by auto text \Obtain the isomorphism from \f'\ to \f\.\ let ?\ = "\Chn = C.inv f.leg0, Dom = Dom ?f', Cod = Dom f\" interpret Dom_\: span_in_category C \Dom \Chn = C.inv f.leg0, Dom = Dom (mkIde f.dsrc (f.leg1 \ C.inv f.leg0)), Cod = Dom f\\ using f'.dom.span_in_category_axioms by simp interpret Cod_\: span_in_category C \Cod \Chn = C.inv f.leg0, Dom = Dom (mkIde f.dsrc (f.leg1 \ C.inv f.leg0)), Cod = Dom f\\ using f.dom.span_in_category_axioms by simp interpret \: arrow_of_spans C ?\ proof show "\Chn \Chn = C.inv f.leg0, Dom = Dom (mkIde f.dsrc (f.leg1 \ C.inv f.leg0)), Cod = Dom f\ : Dom_\.apex \\<^sub>C Cod_\.apex\" using f f.dom.apex_def f'.dom.apex_def is_left_adjoint_char by auto show "Cod_\.leg0 \ Chn \Chn = C.inv f.leg0, Dom = Dom (mkIde f.dsrc (f.leg1 \ C.inv f.leg0)), Cod = Dom f\ = Dom_\.leg0" using f f.dom.apex_def is_left_adjoint_char C.comp_arr_inv C.inv_is_inverse by simp show "Cod_\.leg1 \ Chn \Chn = C.inv f.leg0, Dom = Dom (mkIde f.dsrc (f.leg1 \ C.inv f.leg0)), Cod = Dom f\ = Dom_\.leg1" by simp qed have \: "\?\ : ?f' \ f\ \ iso ?\" using f is_left_adjoint_char iso_char arr_char dom_char cod_char \.arrow_of_spans_axioms f'.dom.apex_def f.dom.apex_def by auto text \ Obtain the isomorphism from \g\ to \g'\. Recall: \g' = mkIde (C.cod g.leg0) (g.dom.leg1 \ C.inv g.leg0)\. The isomorphism is given by \g.leg0\. \ let ?\ = "\Chn = g.leg0, Dom = Dom g, Cod = Dom ?g'\" interpret Dom_\: span_in_category C \Dom \Chn = g.leg0, Dom = Dom g, Cod = Dom (mkIde g.dsrc (g.leg1 \ C.inv g.leg0))\\ using g.dom.span_in_category_axioms by simp interpret Cod_\: span_in_category C \Cod \Chn = g.leg0, Dom = Dom g, Cod = Dom (mkIde g.dsrc (g.leg1 \ C.inv g.leg0))\\ using g'.dom.span_in_category_axioms by simp interpret \: arrow_of_spans C ?\ proof show "\Chn \Chn = g.leg0, Dom = Dom g, Cod = Dom (mkIde g.dsrc (g.leg1 \ C.inv g.leg0))\ : Dom_\.apex \\<^sub>C Cod_\.apex\" using g g.dom.apex_def g'.dom.apex_def is_left_adjoint_char by auto show "Cod_\.leg0 \ Chn \Chn = g.leg0, Dom = Dom g, Cod = Dom (mkIde g.dsrc (g.leg1 \ C.inv g.leg0))\ = Dom_\.leg0" using C.comp_cod_arr by simp show "Cod_\.leg1 \ Chn \Chn = g.leg0, Dom = Dom g, Cod = Dom (mkIde g.dsrc (g.leg1 \ C.inv g.leg0))\ = Dom_\.leg1" using g g.dom.apex_def is_left_adjoint_char C.comp_inv_arr C.inv_is_inverse C.comp_assoc C.comp_arr_dom by simp qed have \: "\?\ : g \ ?g'\ \ iso ?\" using g is_left_adjoint_char iso_char arr_char dom_char cod_char \.arrow_of_spans_axioms g.dom.apex_def g'.dom.apex_def by auto have \\: "tabulation (\) (\) assoc unit src trg ?r (r.\ \ ?\) r.f g" using \ \r.g = ?g'\ r.has_tabulation \.preserved_by_output_iso by simp interpret \\: tabulation vcomp hcomp assoc unit src trg ?r \r.\ \ ?\\ r.f g using \\ by auto have "tabulation (\) (\) assoc unit src trg ?r ((?r \ ?\) \ r.\ \ ?\) f g" using \ \r.f = ?f'\ \\.preserved_by_input_iso [of ?\ f] by argo thus ?thesis by auto qed qed text \The sub-bicategory of maps is locally essentially discrete.\ show "\f f' \ \'. \ is_left_adjoint f; is_left_adjoint f'; \\ : f \ f'\; \\' : f \ f'\ \ \ iso \ \ iso \' \ \ = \'" proof - fix f f' \ \' assume f: "is_left_adjoint f" and f': "is_left_adjoint f'" assume \: "\\ : f \ f'\" and \': "\\' : f \ f'\" obtain f\<^sub>a \ \ where f\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \ \" using f adjoint_pair_def by auto obtain f'\<^sub>a \' \' where f'\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f' f'\<^sub>a \' \'" using f' adjoint_pair_def adjunction_def by auto interpret f\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \ \ using f\<^sub>a by simp interpret f'\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg f' f'\<^sub>a \' \' using f'\<^sub>a by simp interpret f: identity_arrow_of_spans C f using ide_char' [of f] by simp interpret f': identity_arrow_of_spans C f' using ide_char' [of f'] by simp interpret \: arrow_of_spans C \ using \ arr_char by auto interpret \': arrow_of_spans C \' using \' arr_char by auto have 1: "C.iso f.leg0 \ C.iso f'.leg0" using f f' is_left_adjoint_char by simp have 2: "\.chine = C.inv f'.leg0 \ f.leg0" using \ 1 dom_char cod_char \.leg0_commutes C.invert_side_of_triangle by auto moreover have "\'.chine = C.inv f'.leg0 \ f.leg0" using \' 1 dom_char cod_char \'.leg0_commutes C.invert_side_of_triangle by auto ultimately have 3: "\.chine = \'.chine" by simp have "iso \" using 1 2 \ C.isos_compose dom_char cod_char iso_char arr_char by auto hence "iso \'" using 3 iso_char arr_char \'.arrow_of_spans_axioms by simp moreover have "\ = \'" using 3 \ \' dom_char cod_char by fastforce ultimately show "iso \ \ iso \' \ \ = \'" by simp qed qed text \ We can now prove the easier half of the main result (CKS Theorem 4): If \B\ is biequivalent to \Span(C)\, where \C\ is a category with pullbacks, then \B\ is a bicategory of spans. (Well, it is easier given that we have already done the work to show that the notion ``bicategory of spans'' is respected by equivalence of bicategories.) \ theorem equivalent_implies_bicategory_of_spans: assumes "equivalent_bicategories vcomp hcomp assoc unit src trg V\<^sub>1 H\<^sub>1 \\<^sub>1 \\<^sub>1 src\<^sub>1 trg\<^sub>1" shows "bicategory_of_spans V\<^sub>1 H\<^sub>1 \\<^sub>1 \\<^sub>1 src\<^sub>1 trg\<^sub>1" using assms is_bicategory_of_spans bicategory_of_spans_respects_equivalence by blast end subsection "Properties of Bicategories of Spans" text \ We now develop consequences of the axioms for a bicategory of spans, in preparation for proving the other half of the main result. \ context bicategory_of_spans begin notation isomorphic (infix "\" 50) text \ The following is a convenience version of \BS2\ that gives us what we generally want: given specified \f, g\ obtain \\\ that makes \(f, \, g)\ a tabulation of \g \ f\<^sup>*\, not a tabulation of some \r\ isomorphic to \g \ f\<^sup>*\. \ lemma BS2': assumes "is_left_adjoint f" and "is_left_adjoint g" and "src f = src g" and "isomorphic (g \ f\<^sup>*) r" shows "\\. tabulation V H \ \ src trg r \ f g" proof - have 1: "is_left_adjoint f \ is_left_adjoint g \ g \ f\<^sup>* \ r" using assms BS1 by simp obtain \ where \: "\\ : g \ f\<^sup>* \ r\ \ iso \" using 1 isomorphic_def by blast obtain r' \' where \': "tabulation V H \ \ src trg r' \' f g" using assms 1 BS2 by blast interpret \': tabulation V H \ \ src trg r' \' f g using \' by simp let ?\ = "\'.T0.trnr\<^sub>\ r' \'" have \: "\?\ : g \ f\<^sup>* \ r'\ \ iso ?\" using \'.yields_isomorphic_representation by blast have "\\ \ inv ?\ : r' \ r\ \ iso (\ \ inv ?\)" using \ \ isos_compose by blast hence 3: "tabulation V H \ \ src trg r ((\ \ inv ?\ \ f) \ \') f g" using \'.is_preserved_by_base_iso by blast hence "\\. tabulation V H \ \ src trg r \ f g" by blast thus ?thesis using someI_ex [of "\\. tabulation V H \ \ src trg r \ f g"] by simp qed text \ The following observation is made by CKS near the beginning of the proof of Theorem 4: If \w\ is an arbitrary 1-cell, and \g\ and \g \ w\ are maps, then \w\ is in fact a map. It is applied frequently. \ lemma BS4: assumes "is_left_adjoint g" and "ide w" and "is_left_adjoint (g \ w)" shows "is_left_adjoint w" proof - text \ CKS say: ``by (i) there are maps \m, n\ with \w \ nm\<^sup>*\, so, by (ii), we have two tabulations \(1, \, gw)\, \(m, \, gn)\ of \gw\; since tabulations are unique up to equivalence, \m\ is invertible and \w \ nm\<^sup>*\ is a map.'' \ have ex_\: "\\. tabulation V H \ \ src trg (g \ w) \ (src w) (g \ w)" proof - have "(g \ w) \ src w \ g \ w" by (metis assms(3) iso_runit ideD(1) isomorphic_def left_adjoint_is_ide runit_in_hom(2) src_hcomp) moreover have "isomorphic ((g \ w) \ (src w)\<^sup>*) (g \ w)" proof - have "(g \ w) \ src (g \ w) \ g \ w" using calculation isomorphic_implies_ide(2) by auto moreover have "isomorphic (src (g \ w)) (src w)\<^sup>*" proof - interpret src_w: map_in_bicategory V H \ \ src trg \src w\ using assms obj_is_self_adjoint by unfold_locales auto interpret src_w: adjunction_in_bicategory V H \ \ src trg \src w\ \(src w)\<^sup>*\ src_w.\ src_w.\ using src_w.is_map left_adjoint_extends_to_adjunction by simp have "adjoint_pair (src w) (src w)" using assms obj_is_self_adjoint by simp moreover have "adjoint_pair (src w) (src w)\<^sup>*" using adjoint_pair_def src_w.adjunction_in_bicategory_axioms by auto ultimately have "src w \ (src w)\<^sup>*" using left_adjoint_determines_right_up_to_iso by simp moreover have "src w = src (g \ w)" using assms isomorphic_def hcomp_simps(1) left_adjoint_is_ide by simp ultimately show ?thesis by simp qed moreover have "src (g \ w) = trg (src (g \ w))" using assms left_adjoint_is_ide by simp ultimately show ?thesis using assms left_adjoint_is_ide isomorphic_transitive isomorphic_symmetric hcomp_ide_isomorphic by blast qed ultimately show ?thesis using assms obj_is_self_adjoint left_adjoint_is_ide BS2' [of "src w" "g \ w" "g \ w"] by auto qed obtain \ where \: "tabulation V H \ \ src trg (g \ w) \ (src w) (g \ w)" using ex_\ by auto obtain m n where mn: "is_left_adjoint m \ is_left_adjoint n \ isomorphic w (n \ m\<^sup>*)" using assms BS1 [of w] by auto have m\<^sub>a: "adjoint_pair m m\<^sup>* \ isomorphic w (n \ m\<^sup>*)" using mn adjoint_pair_def left_adjoint_extends_to_adjoint_pair by blast have ex_\: "\\. tabulation V H \ \ src trg (g \ w) \ m (g \ n)" proof - have "hseq n m\<^sup>*" using mn isomorphic_implies_ide by auto have "trg (n \ m\<^sup>*) = trg w" using mn m\<^sub>a isomorphic_def by (metis (no_types, lifting) dom_inv in_homE trg_dom trg_inv) hence "trg n = trg w" using mn by (metis assms(2) ideD(1) trg.preserves_reflects_arr trg_hcomp) hence "hseq g n" using assms mn left_adjoint_is_ide ideD(1) by (metis hseq_char) have "hseq g w" using assms left_adjoint_is_ide by simp have "src m = src n" using mn m\<^sub>a \hseq n m\<^sup>*\ adjoint_pair_antipar [of m "m\<^sup>*"] by fastforce have "is_left_adjoint (g \ n)" using assms mn left_adjoints_compose \hseq g n\ by blast moreover have "src m = src (g \ n)" using assms mn \hseq g n\ \src m = src n\ by simp moreover have "(g \ n) \ m\<^sup>* \ g \ w" proof - have 1: "src g = trg (n \ m\<^sup>*)" using assms \trg (n \ m\<^sup>*) = trg w\ \hseq g w\ by fastforce hence "(g \ n) \ m\<^sup>* \ g \ n \ m\<^sup>*" using assms mn m\<^sub>a assoc_in_hom iso_assoc \hseq g n\ \hseq n m\<^sup>*\ isomorphic_def left_adjoint_is_ide right_adjoint_is_ide by (metis hseqE ideD(2) ideD(3)) also have "... \ g \ w" using assms 1 mn m\<^sub>a isomorphic_symmetric hcomp_ide_isomorphic left_adjoint_is_ide by simp finally show ?thesis using isomorphic_transitive by blast qed ultimately show ?thesis using assms mn m\<^sub>a BS2' by blast qed obtain \ where \: "tabulation V H \ \ src trg (g \ w) \ m (g \ n)" using ex_\ by auto interpret \: tabulation V H \ \ src trg \g \ w\ \ \src w\ \g \ w\ using \ by auto interpret \: tabulation V H \ \ src trg \g \ w\ \ m \g \ n\ using \ by auto text \ As usual, the sketch given by CKS seems more suggestive than it is a precise recipe. We can obtain an equivalence map \\e : src w \ src m\\ and \\\ such that \\\ : m \ e \ src w\\. We can also obtain an equivalence map \\e' : src m \ src w\\ and \\'\ such that \\\' : src w \ e' \ m\\. If \\'\ can be taken to be an isomorphism; then we have \e' \ src w \ e' \ m\. Since \e'\ is an equivalence, this shows \m\ is an equivalence, hence its right adjoint \m\<^sup>*\ is also an equivalence and therefore a map. But \w = n \ m\<^sub>a\, so this shows that \w\ is a map. Now, we may assume without loss of generality that \e\ and \e'\ are part of an adjoint equivalence. We have \\\ : m \ e \ src w\\ and \\\' : src w \ e' \ m\\. We may take the transpose of \\\ to obtain \\\ : m \ src w \ e'\\; then \\\' \ \ : m \ m\\ and \\\ \ \' : src w \ e' \ src w \ e'\\. Since \m\ and \src w \ e'\ are maps, by \BS3\ it must be that \\\ and \\'\ are inverses. \ text \ {\bf Note:} CKS don't cite \BS3\ here. I am not sure whether this result can be proved without \BS3\. For example, I am interested in knowing whether it can still be proved under the the assumption that 2-cells between maps are unique, but not necessarily invertible, or maybe even in a more general situation. It looks like the invertibility part of \BS3\ is not used in the proof below. \ have 2: "\e e' \ \ \ \ \' \'. equivalence_in_bicategory (\) (\) \ \ src trg e e' \ \ \ \\' : src w \ e' \ m\ \ \\ : g \ n \ (g \ w) \ e'\ \ iso \ \ \ = \.composite_cell e' \' \ \ \ \\ : m \ e \ src w\ \ \\' : g \ w \ (g \ n) \ e\ \ iso \' \ \ = ((g \ w) \ \) \ \[g \ w, m, e] \ (\ \ e) \ \'" using \ \.apex_unique_up_to_equivalence [of \ "src w" "g \ w"] comp_assoc by metis obtain e e' \ \ \ \ \' \' where *: "equivalence_in_bicategory (\) (\) \ \ src trg e e' \ \ \ \\' : src w \ e' \ m\ \ \\ : g \ n \ (g \ w) \ e'\ \ iso \ \ \ = \.composite_cell e' \' \ \ \ \\ : m \ e \ src w\ \ \\' : g \ w \ (g \ n) \ e\ \ iso \' \ \ = \.composite_cell e \ \ \'" using 2 comp_assoc by auto interpret ee': equivalence_in_bicategory \(\)\ \(\)\ \ \ src trg e e' \ \ using * by simp have equiv_e: "equivalence_map e" using ee'.equivalence_in_bicategory_axioms equivalence_map_def by auto obtain \' where \': "adjoint_equivalence_in_bicategory (\) (\) \ \ src trg e e' \ \'" using equivalence_refines_to_adjoint_equivalence [of e e' \] ee'.unit_in_hom(2) ee'.unit_is_iso ee'.antipar equiv_e by auto interpret ee': adjoint_equivalence_in_bicategory \(\)\ \(\)\ \ \ src trg e e' \ \' using \' by simp interpret e'e: adjoint_equivalence_in_bicategory \(\)\ \(\)\ \ \ src trg e' e \inv \'\ \inv \\ using * ee'.dual_adjoint_equivalence by simp have equiv_e': "equivalence_map e'" using e'e.equivalence_in_bicategory_axioms equivalence_map_def by auto have "hseq m e" using * ide_dom [of \] apply (elim conjE in_homE) by simp have "hseq (src w) e'" using * ide_dom [of \'] apply (elim conjE in_homE) by simp have "e'e.trnr\<^sub>\ m \ \ hom m (src w \ e')" proof - have "src m = trg e" using \hseq m e\ by auto moreover have "src (src w) = trg e'" using \hseq (src w) e'\ by auto moreover have "ide m" using mn left_adjoint_is_ide by simp moreover have "ide (src w)" using assms by simp ultimately show ?thesis using * e'e.adjoint_transpose_right(1) by blast qed hence 3: "\e'e.trnr\<^sub>\ m \ : m \ src w \ e'\" by simp hence "\\' \ e'e.trnr\<^sub>\ m \ : m \ m\ \ \e'e.trnr\<^sub>\ m \ \ \' : src w \ e' \ src w \ e'\" using * by auto moreover have "\m : m \ m\ \ \src w \ e' : src w \ e' \ src w \ e'\" using mn 3 ide_cod [of "e'e.trnr\<^sub>\ m \"] left_adjoint_is_ide by fastforce moreover have 4: "is_left_adjoint (src w \ e')" proof - have "is_left_adjoint (src w)" using assms obj_is_self_adjoint by simp moreover have "is_left_adjoint e'" using e'e.adjunction_in_bicategory_axioms adjoint_pair_def by auto ultimately show ?thesis using left_adjoints_compose \hseq (src w) e'\ by auto qed ultimately have "\' \ e'e.trnr\<^sub>\ m \ = m \ e'e.trnr\<^sub>\ m \ \ \' = src w \ e'" using mn BS3 [of m m "\' \ e'e.trnr\<^sub>\ m \" m] BS3 [of "src w \ e'" "src w \ e'" "e'e.trnr\<^sub>\ m \ \ \'" "src w \ e'"] by auto hence "inverse_arrows \' (e'e.trnr\<^sub>\ m \)" using mn 4 left_adjoint_is_ide inverse_arrows_def by simp hence 5: "iso \'" by auto have "equivalence_map (src w \ e')" using assms obj_is_equivalence_map equiv_e' \hseq (src w) e'\ equivalence_maps_compose by auto hence "equivalence_map m" using * 5 equivalence_map_preserved_by_iso isomorphic_def by auto hence "equivalence_map m\<^sup>*" using mn m\<^sub>a right_adjoint_to_equivalence_is_equivalence by simp hence "is_left_adjoint m\<^sup>*" using equivalence_is_left_adjoint by simp moreover have "hseq n m\<^sup>*" using mn isomorphic_implies_ide by auto ultimately have "is_left_adjoint (n \ m\<^sup>*)" using mn left_adjoints_compose by blast thus ?thesis using mn left_adjoint_preserved_by_iso isomorphic_def isomorphic_symmetric by metis qed end subsection "Choosing Tabulations" context bicategory_of_spans begin notation isomorphic (infix "\" 50) notation iso_class ("\_\") text \ We will ultimately need to have chosen a specific tabulation for each 1-cell. This has to be done carefully, to avoid unnecessary choices. We start out by using \BS1\ to choose a specific factorization of the form \r \ tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*\ for each 1-cell \r\. This has to be done in such a way that all elements of an isomorphism class are assigned the same factorization. \ abbreviation isomorphic_rep where "isomorphic_rep r f g \ is_left_adjoint f \ is_left_adjoint g \ g \ f\<^sup>* \ r" definition tab\<^sub>0 where "tab\<^sub>0 r \ SOME f. \g. isomorphic_rep (iso_class_rep \r\) f g" definition tab\<^sub>1 where "tab\<^sub>1 r \ SOME g. isomorphic_rep (iso_class_rep \r\) (tab\<^sub>0 r) g" definition rep where "rep r \ SOME \. \\ : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\ \ iso \" lemma rep_props: assumes "ide r" shows "\rep r : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\" and "iso (rep r)" and "r \ iso_class_rep \r\" and "isomorphic_rep r (tab\<^sub>0 r) (tab\<^sub>1 r)" and "tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r" proof - have 1: "isomorphic_rep r (tab\<^sub>0 r) (tab\<^sub>1 r)" proof - have "\f g. isomorphic_rep (iso_class_rep \r\) f g" using assms BS1 isomorphic_symmetric rep_iso_class isomorphic_transitive by blast hence "isomorphic_rep (iso_class_rep \r\) (tab\<^sub>0 r) (tab\<^sub>1 r)" using assms tab\<^sub>0_def tab\<^sub>1_def someI_ex [of "\f. \g. isomorphic_rep (iso_class_rep \r\) f g"] someI_ex [of "\g. isomorphic_rep (iso_class_rep \r\) (tab\<^sub>0 r) g"] by simp thus ?thesis using assms isomorphic_symmetric isomorphic_transitive rep_iso_class by blast qed hence "\\. \\ : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\ \ iso \" using isomorphic_def by blast hence 2: "\rep r : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\ \ iso (rep r)" using someI_ex [of "\\. \\ : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\ \ iso \"] rep_def by auto show "\rep r : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\" using 2 by simp show "iso (rep r)" using 2 by simp show "r \ iso_class_rep \r\" using assms rep_iso_class isomorphic_symmetric by simp thus "isomorphic_rep r (tab\<^sub>0 r) (tab\<^sub>1 r)" using 1 isomorphic_transitive by blast thus "tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r" by simp qed lemma tab\<^sub>0_in_hom [intro]: assumes "ide r" shows "\tab\<^sub>0 r : src (tab\<^sub>0 r) \ src r\" and "\tab\<^sub>0 r : tab\<^sub>0 r \ tab\<^sub>0 r\" proof - show "\tab\<^sub>0 r : tab\<^sub>0 r \ tab\<^sub>0 r\" using assms rep_props left_adjoint_is_ide by auto have "trg (tab\<^sub>0 r) = src r" using assms rep_props by (metis ideD(1) isomorphic_implies_hpar(1) isomorphic_implies_hpar(3) right_adjoint_simps(2) src_hcomp) thus "\tab\<^sub>0 r : src (tab\<^sub>0 r) \ src r\" using assms rep_props left_adjoint_is_ide by (intro in_hhomI, auto) qed lemma tab\<^sub>0_simps [simp]: assumes "ide r" shows "ide (tab\<^sub>0 r)" and "is_left_adjoint (tab\<^sub>0 r)" and "trg (tab\<^sub>0 r) = src r" and "dom (tab\<^sub>0 r) = tab\<^sub>0 r" and "cod (tab\<^sub>0 r) = tab\<^sub>0 r" using assms tab\<^sub>0_in_hom rep_props ide_dom left_adjoint_is_ide by auto lemma tab\<^sub>1_in_hom [intro]: assumes "ide r" shows "\tab\<^sub>1 r : src (tab\<^sub>0 r) \ trg r\" and "\tab\<^sub>1 r : tab\<^sub>1 r \ tab\<^sub>1 r\" proof - show "\tab\<^sub>1 r : tab\<^sub>1 r \ tab\<^sub>1 r\" using assms rep_props left_adjoint_is_ide by auto have "trg (tab\<^sub>1 r) = trg r" using assms rep_props by (metis ideD(1) isomorphic_implies_hpar(1) isomorphic_implies_hpar(4) trg_hcomp) moreover have "src (tab\<^sub>0 r) = src (tab\<^sub>1 r)" using assms rep_props by fastforce ultimately show "\tab\<^sub>1 r : src (tab\<^sub>0 r) \ trg r\" using assms rep_props left_adjoint_is_ide by (intro in_hhomI, auto) qed lemma tab\<^sub>1_simps [simp]: assumes "ide r" shows "ide (tab\<^sub>1 r)" and "is_left_adjoint (tab\<^sub>1 r)" and "src (tab\<^sub>1 r) = src (tab\<^sub>0 r)" and "trg (tab\<^sub>1 r) = trg r" and "dom (tab\<^sub>1 r) = tab\<^sub>1 r" and "cod (tab\<^sub>1 r) = tab\<^sub>1 r" using assms tab\<^sub>1_in_hom rep_props ide_dom left_adjoint_is_ide by auto lemma rep_in_hom [intro]: assumes "ide r" shows "\rep r : src r \ trg r\" and "\rep r : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\" proof - show "\rep r : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\" using assms rep_props by auto thus "\rep r : src r \ trg r\" using arrI vconn_implies_hpar(1-4) by force qed lemma rep_simps [simp]: assumes "ide r" shows "arr (rep r)" and "src (rep r) = src r" and "trg (rep r) = trg r" and "dom (rep r) = tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*" and "cod (rep r) = r" using assms rep_in_hom by auto lemma iso_rep: assumes "ide r" shows "iso (rep r)" using assms rep_props by simp end text \ Next, we assign a specific tabulation to each 1-cell r. We can't just do this any old way if we ultimately expect to obtain a mapping that is functorial with respect to vertical composition. What we have to do is to assign the representative \tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*\ its canonical tabulation, obtained as the adjoint transpose of the identity, and then translate this to a tabulation of \r\ via the chosen isomorphism \\rep r : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ r\\. \ locale identity_in_bicategory_of_spans = bicategory_of_spans + fixes r :: 'a assumes is_ide: "ide r" begin interpretation tab\<^sub>0: map_in_bicategory V H \ \ src trg \tab\<^sub>0 r\ using is_ide rep_props by unfold_locales auto interpretation tab\<^sub>1: map_in_bicategory V H \ \ src trg \tab\<^sub>1 r\ using is_ide rep_props by unfold_locales auto text \ A tabulation \(tab\<^sub>0 r, tab, tab\<^sub>1 r)\ of \r\ can be obtained as the adjoint transpose of the isomorphism \\rep r : (tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>* \ r\\. It is essential to define it this way if we expect the mapping from 2-cells of the underlying bicategory to arrows of spans to preserve vertical composition. \ definition tab where "tab \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) (rep r)" text \ In view of \BS2'\, the 1-cell \(tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*\ has the canonical tabulation obtained via adjoint transpose of an identity. In fact, this tabulation generates the chosen tabulation of \r\ in the same isomorphism class by translation along the isomorphism \\rep r : (tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>* \ r\\. This fact is used to show that the mapping from 2-cells to arrows of spans preserves identities. \ lemma canonical_tabulation: shows "tabulation V H \ \ src trg ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*) (tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)) (tab\<^sub>0 r) (tab\<^sub>1 r)" proof - have "\\. tabulation V H \ \ src trg ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*) \ (tab\<^sub>0 r) (tab\<^sub>1 r)" by (simp add: bicategory_of_spans.BS2' bicategory_of_spans_axioms is_ide isomorphic_reflexive) thus ?thesis using is_ide tab\<^sub>0.canonical_tabulation by simp qed lemma tab_def_alt: shows "tab = (rep r \ tab\<^sub>0 r) \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" and "(inv (rep r) \ tab\<^sub>0 r) \ tab = tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" proof - have "tab = tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) (rep r \ ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*))" using tab_def is_ide rep_in_hom [of r] comp_arr_dom by auto also have "... = (rep r \ tab\<^sub>0 r) \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" using is_ide tab\<^sub>0.trnr\<^sub>\_comp by auto finally show 1: "tab = (rep r \ tab\<^sub>0 r) \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" by simp have "(inv (rep r) \ tab\<^sub>0 r) \ tab = ((inv (rep r) \ tab\<^sub>0 r) \ (rep r \ tab\<^sub>0 r)) \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" unfolding 1 using comp_assoc by presburger also have "... = tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" proof - have 1: "(inv (rep r) \ tab\<^sub>0 r) \ (rep r \ tab\<^sub>0 r) = ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*) \ tab\<^sub>0 r" using whisker_right [of "tab\<^sub>0 r" "inv (rep r)" "rep r"] iso_rep rep_in_hom inv_is_inverse comp_inv_arr by (simp add: comp_inv_arr' is_ide) show ?thesis proof - have "\tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*) : tab\<^sub>1 r \ (tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*) \ tab\<^sub>0 r\" by (meson canonical_tabulation tabulation_data.tab_in_hom(2) tabulation_def) hence "((tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*) \ tab\<^sub>0 r) \ tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*) = tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" using 1 comp_cod_arr by blast thus ?thesis using 1 by simp qed qed finally show "(inv (rep r) \ tab\<^sub>0 r) \ tab = tab\<^sub>0.trnr\<^sub>\ (tab\<^sub>1 r) ((tab\<^sub>1 r) \ (tab\<^sub>0 r)\<^sup>*)" by blast qed lemma tab_is_tabulation: shows "tabulation V H \ \ src trg r tab (tab\<^sub>0 r) (tab\<^sub>1 r)" by (metis bicategory_of_spans.iso_rep bicategory_of_spans.rep_in_hom(2) bicategory_of_spans_axioms is_ide canonical_tabulation tab_def_alt(1) tabulation.is_preserved_by_base_iso) (* * TODO: If I pull the interpretation "tab" out of the following, Isabelle warns that * the lemma is a redundant introduction rule and is being "ignored" for that purpose. * However, the redundancy is only in the present context: if the enclosing locale is * interpreted elsewhere, then the rule is not redundant. In order to make sure that * the rule is not "ignored", I have put the interpretation "tab" into the proof to * avoid the warning. *) lemma tab_in_hom [intro]: shows "\tab : src (tab\<^sub>0 r) \ trg r\" and "\tab : tab\<^sub>1 r \ r \ tab\<^sub>0 r\" proof - interpret tab: tabulation V H \ \ src trg r tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ using tab_is_tabulation by simp show "\tab : src (tab\<^sub>0 r) \ trg r\" using tab.tab_in_hom by auto show "\tab : tab\<^sub>1 r \ r \ tab\<^sub>0 r\" using tab.tab_in_hom by auto qed lemma tab_simps [simp]: shows "arr tab" and "src tab = src (tab\<^sub>0 r)" and "trg tab = trg r" and "dom tab = tab\<^sub>1 r" and "cod tab = r \ tab\<^sub>0 r" using tab_in_hom by auto end text \ The following makes the chosen tabulation conveniently available whenever we are considering a particular 1-cell. \ sublocale identity_in_bicategory_of_spans \ tabulation V H \ \ src trg r tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ using is_ide tab_is_tabulation by simp context identity_in_bicategory_of_spans begin interpretation tab\<^sub>0: map_in_bicategory V H \ \ src trg \tab\<^sub>0 r\ using is_ide rep_props by unfold_locales auto interpretation tab\<^sub>1: map_in_bicategory V H \ \ src trg \tab\<^sub>1 r\ using is_ide rep_props by unfold_locales auto text \ The fact that adjoint transpose is a bijection allows us to invert the definition of \tab\ in terms of \rep\ to express rep in terms of tab. \ lemma rep_in_terms_of_tab: shows "rep r = T0.trnr\<^sub>\ r tab" using is_ide T0.adjoint_transpose_right(3) [of r "tab\<^sub>1 r" "rep r"] tab_def by fastforce lemma isomorphic_implies_same_tab: assumes "isomorphic r r'" shows "tab\<^sub>0 r = tab\<^sub>0 r'" and "tab\<^sub>1 r = tab\<^sub>1 r'" using assms tab\<^sub>0_def tab\<^sub>1_def iso_class_eqI by auto text \ ``Every 1-cell has a tabulation as a span of maps.'' Has a nice simple ring to it, but maybe not so useful for us, since we generally really need to know that the tabulation has a specific form. \ lemma has_tabulation: shows "\\ f g. is_left_adjoint f \ is_left_adjoint g \ tabulation V H \ \ src trg r \ f g" using is_ide tab_is_tabulation rep_props by blast end subsection "Tabulations in a Bicategory of Spans" context bicategory_of_spans begin abbreviation tab_of_ide where "tab_of_ide r \ identity_in_bicategory_of_spans.tab V H \ \ src trg r" abbreviation prj\<^sub>0 where "prj\<^sub>0 h k \ tab\<^sub>0 (k\<^sup>* \ h)" abbreviation prj\<^sub>1 where "prj\<^sub>1 h k \ tab\<^sub>1 (k\<^sup>* \ h)" lemma prj_in_hom [intro]: assumes "ide h" and "is_left_adjoint k" and "trg h = trg k" shows "\prj\<^sub>0 h k : src (prj\<^sub>0 h k) \ src h\" and "\prj\<^sub>1 h k : src (prj\<^sub>0 h k) \ src k\" and "\prj\<^sub>0 h k : prj\<^sub>0 h k \ prj\<^sub>0 h k\" and "\prj\<^sub>1 h k : prj\<^sub>1 h k \ prj\<^sub>1 h k\" by (intro in_hhomI, auto simp add: assms(1-3)) lemma prj_simps [simp]: assumes "ide h" and "is_left_adjoint k" and "trg h = trg k" shows "trg (prj\<^sub>0 h k) = src h" and "src (prj\<^sub>1 h k) = src (prj\<^sub>0 h k)" and "trg (prj\<^sub>1 h k) = src k" and "dom (prj\<^sub>0 h k) = prj\<^sub>0 h k" and "cod (prj\<^sub>0 h k) = prj\<^sub>0 h k" and "dom (prj\<^sub>1 h k) = prj\<^sub>1 h k" and "cod (prj\<^sub>1 h k) = prj\<^sub>1 h k" and "is_left_adjoint (prj\<^sub>0 h k)" and "is_left_adjoint (prj\<^sub>1 h k)" using assms prj_in_hom by auto end text \ Many of the commutativity conditions that we would otherwise have to worry about when working with tabulations in a bicategory of spans reduce to trivialities. The following locales try to exploit this to make our life more manageable. \ locale span_of_maps = bicategory_of_spans + fixes leg\<^sub>0 :: 'a and leg\<^sub>1 :: 'a assumes leg0_is_map: "is_left_adjoint leg\<^sub>0" and leg1_is_map : "is_left_adjoint leg\<^sub>1" text \ The purpose of the somewhat strange-looking assumptions in this locale is to cater to the form of data that we obtain from \T1\. Under the assumption that we are in a bicategory of spans and that the legs of \r\ and \s\ are maps, the hypothesized 2-cells will be uniquely determined isomorphisms, and an arrow of spans \w\ from \r\ to \s\ will be a map. We want to prove this once and for all under the weakest assumptions we can manage. \ locale arrow_of_spans_of_maps = bicategory_of_spans V H \ \ src trg + r: span_of_maps V H \ \ src trg r\<^sub>0 r\<^sub>1 + s: span_of_maps V H \ \ src trg s\<^sub>0 s\<^sub>1 for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and r\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a and w :: 'a + assumes is_ide: "ide w" and leg0_lax: "\\. \\ : s\<^sub>0 \ w \ r\<^sub>0\" and leg1_iso: "\\. \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ iso \" begin notation isomorphic (infix "\" 50) lemma composite_leg1_is_map: shows "is_left_adjoint (s\<^sub>1 \ w)" using r.leg1_is_map leg1_iso left_adjoint_preserved_by_iso' isomorphic_def isomorphic_symmetric by auto lemma is_map: shows "is_left_adjoint w" using is_ide composite_leg1_is_map s.leg1_is_map BS4 [of s\<^sub>1 w] by auto lemma hseq_leg\<^sub>0: shows "hseq s\<^sub>0 w" by (metis ideD(1) ide_dom in_homE leg0_lax) lemma composite_with_leg0_is_map: shows "is_left_adjoint (s\<^sub>0 \ w)" using left_adjoints_compose is_map s.leg0_is_map hseq_leg\<^sub>0 by blast lemma leg0_uniquely_isomorphic: shows "s\<^sub>0 \ w \ r\<^sub>0" and "\!\. \\ : s\<^sub>0 \ w \ r\<^sub>0\" proof - show 1: "s\<^sub>0 \ w \ r\<^sub>0" using leg0_lax composite_with_leg0_is_map r.leg0_is_map BS3 [of "s\<^sub>0 \ w" r\<^sub>0] isomorphic_def by auto have "\\. \\ : s\<^sub>0 \ w \ r\<^sub>0\ \ iso \" using 1 isomorphic_def by simp moreover have "\\ \'. \\ : s\<^sub>0 \ w \ r\<^sub>0\ \ \\' : s\<^sub>0 \ w \ r\<^sub>0\ \ \ = \'" using BS3 r.leg0_is_map composite_with_leg0_is_map by blast ultimately show "\!\. \\ : s\<^sub>0 \ w \ r\<^sub>0\" by blast qed lemma leg1_uniquely_isomorphic: shows "r\<^sub>1 \ s\<^sub>1 \ w" and "\!\. \\ : r\<^sub>1 \ s\<^sub>1 \ w\" proof - show 1: "r\<^sub>1 \ s\<^sub>1 \ w" using leg1_iso isomorphic_def by auto have "\\. \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ iso \" using leg1_iso isomorphic_def isomorphic_symmetric by simp moreover have "\\ \'. \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ \\' : r\<^sub>1 \ s\<^sub>1 \ w\ \ \ = \'" using BS3 r.leg1_is_map composite_leg1_is_map by blast ultimately show "\!\. \\ : r\<^sub>1 \ s\<^sub>1 \ w\" by blast qed definition the_\ where "the_\ \ THE \. \\ : s\<^sub>0 \ w \ r\<^sub>0\" definition the_\ where "the_\ \ THE \. \\ : r\<^sub>1 \ s\<^sub>1 \ w\" lemma the_\_props: shows "\the_\ : s\<^sub>0 \ w \ r\<^sub>0\" and "iso the_\" proof - show "\the_\ : s\<^sub>0 \ w \ r\<^sub>0\" unfolding the_\_def using the1I2 [of "\\. \\ : s\<^sub>0 \ w \ r\<^sub>0\" "\\. \\ : s\<^sub>0 \ w \ r\<^sub>0\"] leg0_uniquely_isomorphic by simp thus "iso the_\" using BS3 r.leg0_is_map composite_with_leg0_is_map by simp qed lemma the_\_in_hom [intro]: shows "\the_\ : src r\<^sub>0 \ trg r\<^sub>0\" and "\the_\ : s\<^sub>0 \ w \ r\<^sub>0\" using the_\_props apply auto by (metis cod_trg in_hhom_def in_homE isomorphic_implies_hpar(3) leg0_uniquely_isomorphic(1) src_dom trg.preserves_cod) lemma the_\_simps [simp]: shows "arr the_\" and "src the_\ = src r\<^sub>0" and "trg the_\ = trg r\<^sub>0" and "dom the_\ = s\<^sub>0 \ w" and "cod the_\ = r\<^sub>0" using the_\_in_hom by auto lemma the_\_props: shows "\the_\ : r\<^sub>1 \ s\<^sub>1 \ w\" and "iso the_\" proof - show "\the_\ : r\<^sub>1 \ s\<^sub>1 \ w\" unfolding the_\_def using the1I2 [of "\\. \\ : r\<^sub>1 \ s\<^sub>1 \ w\" "\\. \\ : r\<^sub>1 \ s\<^sub>1 \ w\"] leg1_uniquely_isomorphic by simp thus "iso the_\" using BS3 r.leg1_is_map composite_leg1_is_map by simp qed lemma the_\_in_hom [intro]: shows "\the_\ : src r\<^sub>1 \ trg r\<^sub>1\" and "\the_\ : r\<^sub>1 \ s\<^sub>1 \ w\" using the_\_props apply auto by (metis in_hhom_def in_homE isomorphic_implies_hpar(3) leg1_uniquely_isomorphic(1) src_cod trg_dom) lemma the_\_simps [simp]: shows "arr the_\" and "src the_\ = src r\<^sub>1" and "trg the_\ = trg r\<^sub>1" and "dom the_\ = r\<^sub>1" and "cod the_\ = s\<^sub>1 \ w" using the_\_in_hom by auto end (* * TODO: I could probably avoid repeating the declarations of the locale parameters * if I was willing to accept them being given in their order of appearance. *) locale arrow_of_spans_of_maps_to_tabulation_data = bicategory_of_spans V H \ \ src trg + arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 w + \: tabulation_data V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and r\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s :: 'a and \ :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a and w :: 'a text \ The following declaration allows us to inherit the rules and other facts defined in locale @{locale uw\}. It is tedious to prove very much without these in place. \ sublocale arrow_of_spans_of_maps_to_tabulation_data \ uw\ V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 r\<^sub>0 w the_\ using \.tab_in_hom is_ide the_\_props by unfold_locales auto locale arrow_of_spans_of_maps_to_tabulation = arrow_of_spans_of_maps_to_tabulation_data + tabulation V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 locale tabulation_in_maps = span_of_maps V H \ \ src trg s\<^sub>0 s\<^sub>1 + tabulation V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and s :: 'a and \ :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a sublocale tabulation_in_maps \ tabulation V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 .. sublocale identity_in_bicategory_of_spans \ tabulation_in_maps V H \ \ src trg r tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ using is_ide rep_props by unfold_locales auto locale cospan_of_maps_in_bicategory_of_spans = bicategory_of_spans + fixes h :: 'a and k :: 'a assumes h_is_map: "is_left_adjoint h" and k_is_map: "is_left_adjoint k" and cospan: "trg h = trg k" begin text \ The following sublocale declaration is perhaps pushing the limits of sensibility, but the purpose is, given a cospan of maps \(h, k)\, to obtain ready access to the composite \k\<^sup>* \ h\ and its chosen tabulation. \ sublocale identity_in_bicategory_of_spans V H \ \ src trg \k\<^sup>* \ h\ using h_is_map k_is_map cospan left_adjoint_is_ide by unfold_locales auto notation isomorphic (infix "\" 50) interpretation E: self_evaluation_map V H \ \ src trg .. notation E.eval ("\_\") interpretation h: map_in_bicategory V H \ \ src trg h using h_is_map by unfold_locales auto interpretation k: map_in_bicategory V H \ \ src trg k using k_is_map by unfold_locales auto text \ Our goal here is to reformulate the biuniversal properties of the chosen tabulation of \k\<^sup>* \ h\ in terms of its transpose, which yields a 2-cell from \k \ tab\<^sub>1 (k\<^sup>* \ h)\ to \h \ tab\<^sub>0 (k\<^sup>* \ h)\. These results do not depend on \BS3\. \ abbreviation p\<^sub>0 where "p\<^sub>0 \ prj\<^sub>0 h k" abbreviation p\<^sub>1 where "p\<^sub>1 \ prj\<^sub>1 h k" lemma p\<^sub>0_in_hom [intro]: shows "\p\<^sub>0 : src p\<^sub>0 \ src h\" by auto lemma p\<^sub>1_in_hom [intro]: shows "\p\<^sub>1 : src p\<^sub>0 \ src k\" using prj_in_hom cospan h.ide_left k_is_map by blast lemma p\<^sub>0_simps [simp]: shows "trg p\<^sub>0 = src h" by simp lemma p\<^sub>1_simps [simp]: shows "trg p\<^sub>1 = src k" using k.antipar(1) by auto definition \ where "\ \ k.trnl\<^sub>\ (h \ p\<^sub>0) (\[k\<^sup>*, h, p\<^sub>0] \ tab)" lemma \_in_hom [intro]: shows "\\ : src p\<^sub>0 \ trg h\" and "\\ : k \ p\<^sub>1 \ h \ p\<^sub>0\" proof - show 1: "\\ : k \ p\<^sub>1 \ h \ p\<^sub>0\" unfolding \_def using k.antipar cospan k.adjoint_transpose_left(2) [of "h \ p\<^sub>0" "p\<^sub>1"] by fastforce show "\\ : src p\<^sub>0 \ trg h\" using 1 k.antipar arrI cospan vconn_implies_hpar(1-2) by force qed lemma \_simps [simp]: shows "arr \" and "src \ = src p\<^sub>0" and "trg \ = trg h" and "dom \ = k \ p\<^sub>1" and "cod \ = h \ p\<^sub>0" using \_in_hom by auto lemma transpose_\: shows "tab = \\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ k.trnl\<^sub>\ p\<^sub>1 \" proof - have "\\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ k.trnl\<^sub>\ p\<^sub>1 \ = \\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ \[k\<^sup>*, h, p\<^sub>0] \ tab" unfolding \_def using k.antipar cospan k.adjoint_transpose_left(4) [of "h \ p\<^sub>0" "p\<^sub>1" "\[k\<^sup>*, h, p\<^sub>0] \ tab"] by fastforce also have "... = (\\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ \[k\<^sup>*, h, p\<^sub>0]) \ tab" using comp_assoc by presburger also have "... = tab" using k.antipar cospan comp_cod_arr comp_assoc_assoc' by simp finally show ?thesis by simp qed lemma transpose_triangle: assumes "ide w" and "\\ : p\<^sub>0 \ w \ u\" and "\\ : v \ p\<^sub>1 \ w\" shows "k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) \ \) = (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \)" proof - have u: "ide u" using assms(2) by auto have v: "ide v" using assms(3) by auto have 0: "src p\<^sub>0 = trg w" by (metis assms(2) hseqE ideD(1) src.preserves_reflects_arr u vconn_implies_hpar(3)) have 1: "src h = trg u" using assms(1-2) 0 trg_dom trg_cod vconn_implies_hpar(4) by auto have 2: "src k = trg v" using assms(1,3) 0 trg_dom trg_cod hseqI' by (metis ideD(1) leg1_simps(2) leg1_simps(3) p\<^sub>1_simps trg_hcomp vconn_implies_hpar(4)) have 3: "src u = src v \ src u = src w" using assms 0 k.antipar src_dom src_cod hseqI' by (metis ideD(1) leg0_simps(2) leg1_simps(2) leg1_simps(3) src_hcomp vconn_implies_hpar(3)) have 4: "src h = trg \" using assms 1 k.antipar by auto define \ where "\ = \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w)" have \: "\\ : p\<^sub>1 \ w \ k\<^sup>* \ h \ p\<^sub>0 \ w\" unfolding \_def using assms 0 k.antipar cospan by (intro comp_in_homI, auto) have "k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) \ \) = k.trnl\<^sub>\ (h \ u) ((k\<^sup>* \ h \ \) \ \ \ \)" unfolding \_def using assms 1 k.antipar cospan assoc_naturality [of "k\<^sup>*" h \] comp_assoc by (metis "4" h.ide_left ide_char in_homE k.ide_right) also have "... = k.trnl\<^sub>\ (h \ u) (k\<^sup>* \ h \ \) \ (k \ \ \ \)" proof - have "ide (h \ u)" using "1" u assms h.ide_left by blast moreover have "seq (k\<^sup>* \ h \ \) (\ \ \)" using assms 1 k.antipar cospan \ seqI' apply (intro seqI) apply auto apply blast proof - have "dom (k\<^sup>* \ h \ \) = k\<^sup>* \ h \ p\<^sub>0 \ w" using assms by (metis "4" cospan hcomp_simps(2-3) h.ide_left hseqI' ide_char in_homE k.antipar(2) k.ide_right) also have "... = cod \" using \ by auto finally show "dom (k\<^sup>* \ h \ \) = cod \" by simp qed moreover have "src k = trg (k\<^sup>* \ h \ \)" using assms k.antipar cospan calculation(2) by auto ultimately show ?thesis using k.trnl\<^sub>\_comp by simp qed also have "... = k.trnl\<^sub>\ (h \ u) (k\<^sup>* \ h \ \) \ (k \ \) \ (k \ \)" using assms u \ whisker_left by (metis k.ide_left seqI') also have "... = (\[h \ u] \ (k.\ \ h \ u) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ u] \ (k \ k\<^sup>* \ h \ \)) \ (k \ \) \ (k \ \)" unfolding k.trnl\<^sub>\_def by simp also have "... = (h \ \) \ (\[h \ p\<^sub>0 \ w] \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \)) \ (k \ \)" proof - have "\[h \ u] \ (k.\ \ h \ u) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ u] \ (k \ k\<^sup>* \ h \ \) = \[h \ u] \ (k.\ \ h \ u) \ ((k \ k\<^sup>*) \ h \ \) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w]" using assms 4 k.antipar cospan assoc'_naturality [of k "k\<^sup>*" "h \ \"] by fastforce also have "... = \[h \ u] \ ((k.\ \ h \ u) \ ((k \ k\<^sup>*) \ h \ \)) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w]" using comp_assoc by presburger also have "... = (\[h \ u] \ (trg k \ h \ \)) \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w]" proof - have "(k.\ \ h \ u) \ ((k \ k\<^sup>*) \ h \ \) = k.\ \ (k \ k\<^sup>*) \ (h \ u) \ (h \ \)" using assms 1 k.antipar cospan interchange comp_arr_dom comp_cod_arr by fastforce also have "... = k.\ \ h \ \" using assms k.antipar cospan comp_arr_dom comp_cod_arr k.counit_in_hom whisker_left by (metis h.ide_left in_homE) also have "... = (trg k \ h \ \) \ (k.\ \ h \ p\<^sub>0 \ w)" using assms 4 k.antipar cospan whisker_left comp_arr_dom comp_cod_arr interchange [of "trg k" k.\ "h \ \" "h \ p\<^sub>0 \ w"] by auto finally have "(k.\ \ h \ u) \ ((k \ k\<^sup>*) \ h \ \) = (trg k \ h \ \) \ (k.\ \ h \ p\<^sub>0 \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (h \ \) \ \[h \ p\<^sub>0 \ w] \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w]" proof - have "\[h \ u] \ (trg k \ h \ \) = (h \ \) \ \[h \ p\<^sub>0 \ w]" using assms 1 4 k.antipar cospan lunit_naturality [of "h \ \"] by (metis hcomp_simps(3-4) h.ide_left hseqI' ide_char in_homE trg_hcomp) thus ?thesis using comp_assoc by presburger qed finally have "\[h \ u] \ (k.\ \ h \ u) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ u] \ (k \ k\<^sup>* \ h \ \) = (h \ \) \ \[h \ p\<^sub>0 \ w] \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w]" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \)" proof - have "\[h \ p\<^sub>0 \ w] \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w)) = \[h, p\<^sub>0, w] \ \[(h \ p\<^sub>0) \ w] \ (trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w))" proof - have "\[h \ p\<^sub>0 \ w] = \[h, p\<^sub>0, w] \ \[(h \ p\<^sub>0) \ w] \ (trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w])" proof - have "\[h, p\<^sub>0, w] \ \[(h \ p\<^sub>0) \ w] \ (trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) = \[h, p\<^sub>0, w] \ \ ((h \ p\<^sub>0) \ w) \ (trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w])" using assms 0 k.antipar cospan comp_cod_arr \_ide_simp by simp also have "... = \[h, p\<^sub>0, w] \ \ (\\<^sup>-\<^sup>1[h, p\<^sub>0, w])" using assms 0 k.antipar cospan \.is_natural_2 [of "\\<^sup>-\<^sup>1[h, p\<^sub>0, w]"] by simp also have "... = \[h, p\<^sub>0, w] \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w] \ \ (h \ p\<^sub>0 \ w)" using assms 0 k.antipar cospan \.is_natural_1 [of "\\<^sup>-\<^sup>1[h, p\<^sub>0, w]"] comp_assoc by simp also have "... = (\[h, p\<^sub>0, w] \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \ (h \ p\<^sub>0 \ w)" using comp_assoc by presburger also have "... = \ (h \ p\<^sub>0 \ w)" using assms 0 k.antipar cospan comp_cod_arr comp_assoc_assoc' by simp also have "... = \[h \ p\<^sub>0 \ w]" using assms 0 k.antipar cospan \_ide_simp by simp finally show ?thesis by simp qed thus ?thesis using comp_assoc by presburger qed also have "... = \[h, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ \\<^sup>-\<^sup>1[trg h, h \ p\<^sub>0, w] \ ((trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ (k.\ \ h \ p\<^sub>0 \ w)) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w))" using assms 0 k.antipar cospan lunit_hcomp comp_assoc by simp also have "... = \[h, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ (\\<^sup>-\<^sup>1[trg h, h \ p\<^sub>0, w] \ (k.\ \ (h \ p\<^sub>0) \ w)) \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w))" proof - have "(trg h \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ (k.\ \ h \ p\<^sub>0 \ w) = (k.\ \ (h \ p\<^sub>0) \ w) \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w])" using assms 0 k.antipar cospan comp_arr_dom comp_cod_arr interchange [of "trg h" k.\ "\\<^sup>-\<^sup>1[h, p\<^sub>0, w]" "h \ p\<^sub>0 \ w"] interchange [of k.\ "k \ k\<^sup>*" "(h \ p\<^sub>0) \ w" "\\<^sup>-\<^sup>1[h, p\<^sub>0, w]"] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[h, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ ((k.\ \ (h \ p\<^sub>0)) \ w) \ \\<^sup>-\<^sup>1[k \ k\<^sup>*, h \ p\<^sub>0, w] \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w))" using assms 0 k.antipar cospan assoc'_naturality [of k.\ "h \ p\<^sub>0" w] comp_assoc by simp also have "... = \[h, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ ((k.\ \ (h \ p\<^sub>0)) \ w) \ \\<^sup>-\<^sup>1[k \ k\<^sup>*, h \ p\<^sub>0, w] \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w]) \ (k \ \[k\<^sup>* \ h, p\<^sub>0, w]) \ (k \ tab \ w)" proof - have "k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) = (k \ \[k\<^sup>*, h, p\<^sub>0 \ w]) \ (k \ \[k\<^sup>* \ h, p\<^sub>0, w]) \ (k \ tab \ w)" proof - have "seq \[k\<^sup>*, h, p\<^sub>0 \ w] (\[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w))" using \_def assms 0 k.antipar cospan \ by blast thus ?thesis using assms 0 k.antipar cospan whisker_left by auto qed thus ?thesis using comp_assoc by presburger qed also have "... = \[h, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ ((k.\ \ (h \ p\<^sub>0)) \ w) \ (\\<^sup>-\<^sup>1[k \ k\<^sup>*, h \ p\<^sub>0, w] \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w]) \ (k \ \[k\<^sup>* \ h, p\<^sub>0, w]) \ \[k, (k\<^sup>* \ h) \ p\<^sub>0, w]) \ ((k \ tab) \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" proof - have "k \ tab \ w = \[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ \\<^sup>-\<^sup>1[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ (k \ tab \ w)" proof - have "\[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ \\<^sup>-\<^sup>1[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ (k \ tab \ w) = (\[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ \\<^sup>-\<^sup>1[k, (k\<^sup>* \ h) \ p\<^sub>0, w]) \ (k \ tab \ w)" using comp_assoc by presburger also have "... = (k \ ((k\<^sup>* \ h) \ p\<^sub>0) \ w) \ (k \ tab \ w)" using assms k.antipar 0 comp_assoc_assoc' by simp also have "... = k \ tab \ w" using assms k.antipar 0 comp_cod_arr by simp finally show ?thesis by simp qed also have "... = \[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ ((k \ tab) \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" using assms 0 k.antipar cospan assoc'_naturality [of k tab w] by simp finally have "k \ tab \ w = \[k, (k\<^sup>* \ h) \ p\<^sub>0, w] \ ((k \ tab) \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[h, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ w) \ ((k.\ \ h \ p\<^sub>0) \ w) \ (\\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0] \ (k \ \[k\<^sup>*, h, p\<^sub>0]) \ w) \ ((k \ tab) \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" proof - have "\\<^sup>-\<^sup>1[k \ k\<^sup>*, h \ p\<^sub>0, w] \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w]) \ (k \ \[k\<^sup>* \ h, p\<^sub>0, w]) \ \[k, (k\<^sup>* \ h) \ p\<^sub>0, w] = \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0] \ (k \ \[k\<^sup>*, h, p\<^sub>0]) \ w" proof - have "\\<^sup>-\<^sup>1[k \ k\<^sup>*, h \ p\<^sub>0, w] \ ((k \ k\<^sup>*) \ \\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w]) \ (k \ \[k\<^sup>* \ h, p\<^sub>0, w]) \ \[k, (k\<^sup>* \ h) \ p\<^sub>0, w] = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\k\<^bold>\ \<^bold>\ \<^bold>\k\<^sup>*\<^bold>\, \<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\k\<^bold>\ \<^bold>\ \<^bold>\k\<^sup>*\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\k\<^bold>\, \<^bold>\k\<^sup>*\<^bold>\, \<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\k\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\k\<^sup>*\<^bold>\, \<^bold>\h\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\k\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\k\<^sup>*\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\k\<^bold>\, (\<^bold>\k\<^sup>*\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]\" using assms 0 k.antipar cospan \_def \'_def by simp also have "... = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\k\<^bold>\, \<^bold>\k\<^sup>*\<^bold>\, \<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\k\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\k\<^sup>*\<^bold>\, \<^bold>\h\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\w\<^bold>\\" using assms 0 k.antipar cospan by (intro E.eval_eqI, simp_all) also have "... = \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0] \ (k \ \[k\<^sup>*, h, p\<^sub>0]) \ w" using assms 0 k.antipar cospan \_def \'_def by simp finally show ?thesis by simp qed thus ?thesis using comp_assoc by presburger qed also have "... = \[h, p\<^sub>0, w] \ (\[h \ p\<^sub>0] \ (k.\ \ h \ p\<^sub>0) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0] \ (k \ \[k\<^sup>*, h, p\<^sub>0]) \ (k \ tab) \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" using assms 0 k.antipar cospan comp_assoc whisker_right by auto also have "... = \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" unfolding \_def k.trnl\<^sub>\_def using assms 0 k.antipar cospan comp_assoc whisker_left by simp finally have "\[h \ p\<^sub>0 \ w] \ (k.\ \ h \ p\<^sub>0 \ w) \ \\<^sup>-\<^sup>1[k, k\<^sup>*, h \ p\<^sub>0 \ w] \ (k \ \[k\<^sup>*, h, p\<^sub>0 \ w] \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w)) = \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" by blast thus ?thesis using \_def comp_assoc by simp qed finally show ?thesis by simp qed text \ \BS3\ implies that \\\ is the unique 2-cell from \k \ p\<^sub>1\ to \h \ p\<^sub>0\ and is an isomorphism. \ lemma \_uniqueness: shows "\\. \\ : k \ p\<^sub>1 \ h \ p\<^sub>0\ \ \ = \" and "iso \" proof - have 2: "is_left_adjoint (k \ p\<^sub>1)" using k.antipar cospan left_adjoints_compose by (simp add: k_is_map) have 3: "is_left_adjoint (h \ p\<^sub>0)" using k.antipar cospan left_adjoints_compose by (simp add: h_is_map) show "\\. \\ : k \ p\<^sub>1 \ h \ p\<^sub>0\ \ \ = \" using \_in_hom 2 3 BS3 by simp show "iso \" using \_in_hom 2 3 BS3 by simp qed text \ As a consequence, the chosen tabulation of \k\<^sup>* \ h\ is the unique 2-cell from \p\<^sub>1\ to \(k\<^sup>* \ h) \ p\<^sub>0\, and therefore if we are given any such 2-cell we may conclude it yields a tabulation of \k\<^sup>* \ h\. \ lemma tab_uniqueness: assumes "\\ : p\<^sub>1 \ (k\<^sup>* \ h) \ p\<^sub>0\" shows "\ = tab" proof - have "\k.trnl\<^sub>\ (h \ p\<^sub>0) (\[k\<^sup>*, h, p\<^sub>0] \ \) : k \ p\<^sub>1 \ h \ p\<^sub>0\" using assms k.antipar cospan k.adjoint_transpose_left(2) [of "h \ p\<^sub>0" "p\<^sub>1"] assoc_in_hom by force hence "tab = \\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ k.trnl\<^sub>\ p\<^sub>1 (k.trnl\<^sub>\ (h \ p\<^sub>0) (\[k\<^sup>*, h, p\<^sub>0] \ \))" using transpose_\ \_uniqueness(1) by auto also have "... = \\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ \[k\<^sup>*, h, p\<^sub>0] \ \" using assms k.antipar cospan k.adjoint_transpose_left(4) assoc_in_hom by auto also have "... = (\\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \ \[k\<^sup>*, h, p\<^sub>0]) \ \" using comp_assoc by presburger also have "... = \" using assms k.antipar cospan comp_cod_arr comp_assoc_assoc' by auto finally show ?thesis by simp qed text \ The following lemma reformulates the biuniversal property of the canonical tabulation of \k\<^sup>* \ h\ as a biuniversal property of \\\, regarded as a square that commutes up to isomorphism. \ lemma \_biuniversal_prop: assumes "ide u" and "ide v" shows "\\. \\ : k \ v \ h \ u\ \ \w \ \. ide w \ \\ : p\<^sub>0 \ w \ u\ \ \\ : v \ p\<^sub>1 \ w\ \ iso \ \ (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \) = \" and "\w w' \ \' \. \ ide w; ide w'; \\ : p\<^sub>0 \ w \ u\; \\' : p\<^sub>0 \ w' \ u\; \\ : p\<^sub>1 \ w \ p\<^sub>1 \ w'\; (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] = (h \ \') \ \[h, p\<^sub>0, w'] \ (\ \ w') \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \ (k \ \) \ \ \!\. \\ : w \ w'\ \ \ = \' \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = \" proof - fix \ assume \: "\\ : k \ v \ h \ u\" have 1: "src h = trg u" using assms \ ide_cod by (metis ide_def in_homE seq_if_composable) have 2: "src k = trg v" using assms \ ide_dom by (metis ideD(1) in_homE not_arr_null seq_if_composable) let ?\ = "\\<^sup>-\<^sup>1[k\<^sup>*, h, u] \ k.trnl\<^sub>\ v \" have \: "\?\ : v \ (k\<^sup>* \ h) \ u\" using assms \ 1 2 k.antipar cospan k.adjoint_transpose_left(1) [of "h \ u" v] assoc_in_hom by auto obtain w \ \ where w\\: "ide w \ \\ : p\<^sub>0 \ w \ u\ \ \\ : v \ p\<^sub>1 \ w\ \ iso \ \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) \ \ = ?\" using assms \ T1 [of u ?\] comp_assoc by (metis in_homE) have 0: "src p\<^sub>0 = trg w" using w\\ ide_dom by (metis hseqE ideD(1) in_homE) have "\ = k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) \ \)" proof - have "\ = k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ?\)" proof - have "k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ?\) = k.trnl\<^sub>\ (h \ u) ((\[k\<^sup>*, h, u] \ \\<^sup>-\<^sup>1[k\<^sup>*, h, u]) \ k.trnl\<^sub>\ v \)" using comp_assoc by presburger also have "... = k.trnl\<^sub>\ (h \ u) (k.trnl\<^sub>\ v \)" proof - have "(\[k\<^sup>*, h, u] \ \\<^sup>-\<^sup>1[k\<^sup>*, h, u]) \ k.trnl\<^sub>\ v \ = (k\<^sup>* \ h \ u) \ k.trnl\<^sub>\ v \" using comp_assoc_assoc' by (simp add: "1" assms(1) cospan k.antipar(2)) also have "... = k.trnl\<^sub>\ v \" using "1" \ assms(1) comp_ide_arr cospan k.antipar(2) by fastforce finally show ?thesis by simp qed also have "... = \" using assms \ k.antipar cospan 1 2 k.adjoint_transpose_left(3) by simp finally show ?thesis by simp qed thus ?thesis using w\\ by simp qed also have "... = (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \)" using assms k.antipar cospan w\\ transpose_triangle [of w \ u \] by auto finally have "(h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \) = \" by simp thus "\w \ \. ide w \ \\ : p\<^sub>0 \ w \ u\ \ \\ : v \ p\<^sub>1 \ w\ \ iso \ \ (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \) = \" using w\\ by blast next fix w w' \ \' \ assume w: "ide w" assume w': "ide w'" assume \: "\\ : p\<^sub>0 \ w \ u\" assume \': "\\' : p\<^sub>0 \ w' \ u\" assume \: "\\ : p\<^sub>1 \ w \ p\<^sub>1 \ w'\" assume eq: "(h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] = (h \ \') \ \[h, p\<^sub>0, w'] \ (\ \ w') \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \ (k \ \)" have 0: "src p\<^sub>0 = trg w" using \ ide_dom by (metis ideD(1) in_homE not_arr_null seq_if_composable) interpret uw\w'\': uw\w'\' V H \ \ src trg \k\<^sup>* \ h\ tab \p\<^sub>0\ \p\<^sub>1\ u w \ w' \' using w \ w' \' apply (unfold_locales) by auto show "\!\. \\ : w \ w'\ \ \ = \' \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = \" proof - let ?LHS = "\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w)" let ?RHS = "\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \') \ \[k\<^sup>* \ h, p\<^sub>0, w'] \ (tab \ w') \ \" have eq': "?LHS = ?RHS" proof - have "k.trnl\<^sub>\ (h \ u) ?LHS = k.trnl\<^sub>\ (h \ u) (\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) \ (p\<^sub>1 \ w))" using assms 0 w \ \ k.antipar cospan comp_arr_dom by (metis tab_simps(1) tab_simps(4) whisker_right) also have "... = (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ p\<^sub>1 \ w)" using assms w \ \ transpose_triangle by (metis arr_dom ide_hcomp ide_in_hom(2) in_homE ide_leg1 not_arr_null seq_if_composable) also have "... = (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" using assms 0 w k.antipar cospan comp_arr_dom by simp also have "... = (h \ \') \ \[h, p\<^sub>0, w'] \ (\ \ w') \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \ (k \ \)" using eq by blast also have "... = k.trnl\<^sub>\ (h \ u) ?RHS" using assms w' \' \ transpose_triangle by simp finally have 4: "k.trnl\<^sub>\ (h \ u) ?LHS = k.trnl\<^sub>\ (h \ u) ?RHS" by simp have "src k = trg (p\<^sub>1 \ w)" using assms 0 w k.antipar cospan by simp moreover have "src k\<^sup>* = trg (h \ u)" using assms 0 w k.antipar cospan by simp moreover have "ide (h \ u)" using assms 0 w k.antipar cospan by simp moreover have "ide (p\<^sub>1 \ w)" using assms 0 w k.antipar cospan by simp ultimately have "inj_on (k.trnl\<^sub>\ (h \ u)) (hom (p\<^sub>1 \ w) (k\<^sup>* \ h \ u))" using assms 0 w w' k.antipar cospan k.adjoint_transpose_left(6) bij_betw_imp_inj_on by blast moreover have "?LHS \ hom (p\<^sub>1 \ w) (k\<^sup>* \ h \ u)" proof - have "\\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \) \ \[k\<^sup>* \ h, p\<^sub>0, w] \ (tab \ w) : p\<^sub>1 \ w \ k\<^sup>* \ h \ u\" using k.antipar cospan apply (intro comp_in_homI) by auto thus ?thesis by simp qed moreover have "?RHS \ hom (p\<^sub>1 \ w) (k\<^sup>* \ h \ u)" proof - have "\\[k\<^sup>*, h, u] \ ((k\<^sup>* \ h) \ \') \ \[k\<^sup>* \ h, p\<^sub>0, w'] \ (tab \ w') \ \ : p\<^sub>1 \ w \ k\<^sup>* \ h \ u\" using \ k.antipar cospan apply (intro comp_in_homI) by auto thus ?thesis by blast qed ultimately show "?LHS = ?RHS" using assms 4 k.antipar cospan bij_betw_imp_inj_on inj_on_def [of "k.trnl\<^sub>\ (h \ u)" "hom (p\<^sub>1 \ w) (k\<^sup>* \ h \ u)"] by simp qed moreover have "seq \[k\<^sup>*, h, u] (composite_cell w \)" using assms k.antipar cospan tab_in_hom by fastforce moreover have "seq \[k\<^sup>*, h, u] (composite_cell w' \' \ \)" using assms \ k.antipar cospan tab_in_hom by fastforce ultimately have "composite_cell w \ = composite_cell w' \' \ \" using assms 0 w w' \ k.antipar cospan iso_assoc iso_is_section section_is_mono monoE [of "\[k\<^sup>*, h, u]" "composite_cell w \" "composite_cell w' \' \ \"] comp_assoc by simp thus ?thesis using w w' \ \' \ eq' T2 [of w w' \ u \' \] by metis qed qed text \ Using the uniqueness properties established for \\\, we obtain yet another reformulation of the biuniversal property associated with the chosen tabulation of \k\<^sup>* \ h\, this time as a kind of pseudo-pullback. We will use this to show that the category of isomorphism classes of maps has pullbacks. \ lemma has_pseudo_pullback: assumes "is_left_adjoint u" and "is_left_adjoint v" and "isomorphic (k \ v) (h \ u)" shows "\w. is_left_adjoint w \ isomorphic (p\<^sub>0 \ w) u \ isomorphic v (p\<^sub>1 \ w)" and "\w w'. \ is_left_adjoint w; is_left_adjoint w'; p\<^sub>0 \ w \ u; v \ p\<^sub>1 \ w; p\<^sub>0 \ w' \ u; v \ p\<^sub>1 \ w' \ \ w \ w'" proof - interpret u: map_in_bicategory V H \ \ src trg u using assms(1) by unfold_locales auto interpret v: map_in_bicategory V H \ \ src trg v using assms(2) by unfold_locales auto obtain \ where \: "\\ : k \ v \ h \ u\ \ iso \" using assms(3) by auto obtain w \ \ where w\\: "ide w \ \\ : p\<^sub>0 \ w \ u\ \ \\ : v \ p\<^sub>1 \ w\ \ iso \ \ (h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] \ (k \ \) = \" using assms \ \_biuniversal_prop(1) [of u v \] by auto have "is_left_adjoint w \ isomorphic (p\<^sub>0 \ w) u \ isomorphic v (p\<^sub>1 \ w)" proof (intro conjI) show 1: "is_left_adjoint w" using assms(2) w\\ left_adjoint_preserved_by_iso' isomorphic_def BS4 leg1_is_map by blast show "v \ p\<^sub>1 \ w" using w\\ isomorphic_def by blast show "p\<^sub>0 \ w \ u" proof - have "src p\<^sub>0 = trg w" using w\\ ide_dom by (metis ideD(1) in_homE not_arr_null seq_if_composable) hence "is_left_adjoint (p\<^sub>0 \ w)" using 1 left_adjoints_compose by simp thus ?thesis using assms w\\ 1 BS3 isomorphic_def by metis qed qed thus "\w. is_left_adjoint w \ p\<^sub>0 \ w \ u \ v \ p\<^sub>1 \ w" by blast show "\w w'. \ is_left_adjoint w; is_left_adjoint w'; p\<^sub>0 \ w \ u; v \ p\<^sub>1 \ w; p\<^sub>0 \ w' \ u; v \ p\<^sub>1 \ w' \ \ w \ w'" proof - fix w w' assume w: "is_left_adjoint w" and w': "is_left_adjoint w'" assume 1: "p\<^sub>0 \ w \ u" assume 2: "v \ p\<^sub>1 \ w" assume 3: "p\<^sub>0 \ w' \ u" assume 4: "v \ p\<^sub>1 \ w'" obtain \ where \: "\\ : p\<^sub>0 \ w \ u\" using 1 by auto obtain \' where \': "\\' : p\<^sub>0 \ w' \ u\" using 3 by auto obtain \ where \: "\\: v \ p\<^sub>1 \ w\ \ iso \" using 2 by blast obtain \' where \': "\\': v \ p\<^sub>1 \ w'\ \ iso \'" using 4 by blast let ?\ = "\' \ inv \" have \: "\?\ : p\<^sub>1 \ w \ p\<^sub>1 \ w'\" using \ \' by (elim conjE in_homE, auto) interpret uw\: uw\ V H \ \ src trg \k\<^sup>* \ h\ tab \p\<^sub>0\ \p\<^sub>1\ u w \ using w \ left_adjoint_is_ide by unfold_locales auto interpret uw'\': uw\ V H \ \ src trg \k\<^sup>* \ h\ tab \p\<^sub>0\ \p\<^sub>1\ u w' \' using w' \' left_adjoint_is_ide by unfold_locales auto interpret uw\w'\': uw\w'\' V H \ \ src trg \k\<^sup>* \ h\ tab \p\<^sub>0\ \p\<^sub>1\ u w \ w' \' using w w' \ \' left_adjoint_is_ide by unfold_locales have "(h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w] = (h \ \') \ \[h, p\<^sub>0, w'] \ (\ \ w') \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \ (k \ ?\)" proof - let ?LHS = "(h \ \) \ \[h, p\<^sub>0, w] \ (\ \ w) \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w]" let ?RHS = "(h \ \') \ \[h, p\<^sub>0, w'] \ (\ \ w') \ \\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \ (k \ ?\)" have "\?LHS : k \ p\<^sub>1 \ w \ h \ u\" using w k.antipar by fastforce moreover have "\?RHS : k \ p\<^sub>1 \ w \ h \ u\" using w k.antipar \ by fastforce moreover have "is_left_adjoint (k \ p\<^sub>1 \ w)" using w k.is_map left_adjoints_compose by simp moreover have "is_left_adjoint (h \ u)" using assms h.is_map left_adjoints_compose by auto ultimately show "?LHS = ?RHS" using BS3 by blast qed hence "\!\. \\ : w \ w'\ \ \ = \' \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = ?\" using assms left_adjoint_is_ide w w' \ \' \ \' \ \_biuniversal_prop(2) [of u v w w' \ \' ?\] by presburger thus "w \ w'" using w w' BS3 isomorphic_def by metis qed qed end subsubsection "Tabulations in Maps" text \ Here we focus our attention on the properties of tabulations in a bicategory of spans, in the special case in which both legs are maps. \ context tabulation_in_maps begin text \ The following are the conditions under which \w\ is a 1-cell induced via \T1\ by a 2-cell \\\ : dom \ \ s \ r\<^sub>0\\: \w\ is an arrow of spans and \\\ is obtained by composing the tabulation \\\ with \w\ and the isomorphisms that witness \w\ being an arrow of spans. \ abbreviation is_induced_by_cell where "is_induced_by_cell w r\<^sub>0 \ \ arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 (dom \) s\<^sub>0 s\<^sub>1 w \ composite_cell w (arrow_of_spans_of_maps.the_\ V H r\<^sub>0 s\<^sub>0 w) \ (arrow_of_spans_of_maps.the_\ V H (dom \) s\<^sub>1 w) = \" lemma induced_map_unique: assumes "is_induced_by_cell w r\<^sub>0 \" and "is_induced_by_cell w' r\<^sub>0 \" shows "isomorphic w w'" proof - interpret w: arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 \dom \\ s\<^sub>0 s\<^sub>1 w using assms(1) by auto interpret w: arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 \dom \\ s \ s\<^sub>0 s\<^sub>1 w .. interpret w': arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 \dom \\ s\<^sub>0 s\<^sub>1 w' using assms(2) by auto interpret w': arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 \dom \\ s \ s\<^sub>0 s\<^sub>1 w' .. let ?\ = "w'.the_\ \ inv w.the_\" have \: "\?\ : s\<^sub>1 \ w \ s\<^sub>1 \ w'\" using w.the_\_props w'.the_\_props arr_iff_in_hom by auto have 1: "composite_cell w w.the_\ = composite_cell w' w'.the_\ \ (w'.the_\ \ inv w.the_\)" proof - have "composite_cell w' w'.the_\ \ (w'.the_\ \ inv w.the_\) = ((composite_cell w' w'.the_\) \ w'.the_\) \ inv w.the_\" using comp_assoc by presburger also have "... = \ \ inv w.the_\" using assms(2) comp_assoc by simp also have "... = (composite_cell w w.the_\ \ w.the_\) \ inv w.the_\" using assms(1) comp_assoc by simp also have "... = composite_cell w w.the_\ \ w.the_\ \ inv w.the_\" using comp_assoc by presburger also have "... = composite_cell w w.the_\" proof - have "w.the_\ \ inv w.the_\ = s\<^sub>1 \ w" using w.the_\_props comp_arr_inv inv_is_inverse by auto thus ?thesis using composite_cell_in_hom w.ide_w w.the_\_props comp_arr_dom by (metis composite_cell_in_hom in_homE w.w_in_hom(1)) qed finally show ?thesis by auto qed have "\\. \\ : w \ w'\" using 1 \ w.is_ide w'.is_ide w.the_\_props w'.the_\_props T2 [of w w' w.the_\ r\<^sub>0 w'.the_\ ?\] by blast thus ?thesis using BS3 w.is_map w'.is_map by blast qed text \ The object src \s\<^sub>0\ forming the apex of the tabulation satisfies the conditions for being a map induced via \T1\ by the 2-cell \\\ itself. This is ultimately required for the map from 2-cells to arrows of spans to be functorial with respect to vertical composition. \ lemma apex_is_induced_by_cell: shows "is_induced_by_cell (src s\<^sub>0) s\<^sub>0 \" proof - have 1: "arrow_of_spans_of_maps V H \ \ src trg s\<^sub>0 s\<^sub>1 s\<^sub>0 s\<^sub>1 (src s\<^sub>0)" using iso_runit [of s\<^sub>0] iso_runit [of s\<^sub>1] tab_in_hom apply unfold_locales apply simp using ide_leg0 isomorphic_def apply blast using ide_leg1 isomorphic_def leg1_simps(3) runit'_in_vhom [of s\<^sub>1 "src s\<^sub>0"] iso_runit' by blast interpret w: arrow_of_spans_of_maps V H \ \ src trg s\<^sub>0 \dom \\ s\<^sub>0 s\<^sub>1 \src s\<^sub>0\ using 1 tab_in_hom by simp interpret w: arrow_of_spans_of_maps_to_tabulation V H \ \ src trg s\<^sub>0 \dom \\ s \ s\<^sub>0 s\<^sub>1 \src s\<^sub>0\ .. show "is_induced_by_cell (src s\<^sub>0) s\<^sub>0 \" proof (intro conjI) show "arrow_of_spans_of_maps V H \ \ src trg s\<^sub>0 (dom \) s\<^sub>0 s\<^sub>1 (src s\<^sub>0)" using w.arrow_of_spans_of_maps_axioms by simp show "composite_cell (src s\<^sub>0) w.the_\ \ w.the_\ = \" proof - have \: "w.the_\ = \[s\<^sub>0]" using iso_runit [of s\<^sub>0] w.leg0_uniquely_isomorphic w.the_\_props the1_equality [of "\\. \\ : s\<^sub>0 \ src s\<^sub>0 \ s\<^sub>0\ \ iso \"] by auto have \: "w.the_\ = \\<^sup>-\<^sup>1[s\<^sub>1]" using iso_runit' w.leg1_uniquely_isomorphic w.the_\_props leg1_simps(3) the1_equality [of "\\. \\ : s\<^sub>1 \ s\<^sub>1 \ src s\<^sub>0\ \ iso \"] tab_in_vhom' by auto have "composite_cell (src s\<^sub>0) \[s\<^sub>0] \ \\<^sup>-\<^sup>1[s\<^sub>1] = \" proof - have "composite_cell (src s\<^sub>0) \[s\<^sub>0] \ \\<^sup>-\<^sup>1[s\<^sub>1] = ((s \ \[s\<^sub>0]) \ \[s, s\<^sub>0, src s\<^sub>0]) \ (\ \ src s\<^sub>0) \ \\<^sup>-\<^sup>1[s\<^sub>1]" using comp_assoc by presburger also have "... = (\[s \ s\<^sub>0] \ (\ \ src s\<^sub>0)) \ \\<^sup>-\<^sup>1[s\<^sub>1]" using runit_hcomp comp_assoc by simp also have "... = \ \ \[s\<^sub>1] \ \\<^sup>-\<^sup>1[s\<^sub>1]" using runit_naturality tab_in_hom by (metis tab_simps(1) tab_simps(2) tab_simps(4) tab_simps(5) comp_assoc) also have "... = \" using iso_runit tab_in_hom comp_arr_dom comp_arr_inv inv_is_inverse by simp finally show ?thesis by simp qed thus ?thesis using \ \ comp_assoc by simp qed qed qed end subsubsection "Composing Tabulations" text \ Given tabulations \(r\<^sub>0, \, r\<^sub>1)\ of \r\ and \(s\<^sub>0, \, s\<^sub>1)\ of \s\ in a bicategory of spans, where \(r\<^sub>0, r\<^sub>1)\ and \(s\<^sub>0, s\<^sub>1)\ are spans of maps and 1-cells \r\ and \s\ are composable, we can construct a 2-cell that yields a tabulation of \r \ s\. The proof uses the fact that the 2-cell \\\ associated with the cospan \(r\<^sub>0, s\<^sub>1)\ is an isomorphism, which we have proved above (\cospan_of_maps_in_bicategory_of_spans.\_uniqueness\) using \BS3\. However, this is the only use of \BS3\ in the proof, and it seems plausible that it would be possible to establish that \\\ is an isomorphism in more general situations in which the subbicategory of maps is not locally essentially discrete. Alternatively, more general situations could be treated by adding the assertion that \\\ is an isomorphism as part of a weakening of \BS3\. \ locale composite_tabulation_in_maps = bicategory_of_spans V H \ \ src trg + \: tabulation_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 + \: tabulation_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and r :: 'a and \ :: 'a and r\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s :: 'a and \ :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a + assumes composable: "src r = trg s" begin text \ Interpret \(r\<^sub>0, s\<^sub>1)\ as a @{locale cospan_of_maps_in_bicategory_of_spans}, to obtain the isomorphism \\\ in the central diamond, along with the assertion that it is unique. \ interpretation r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg s\<^sub>1 r\<^sub>0 using \.leg0_is_map \.leg1_is_map composable by unfold_locales auto text \ We need access to simps, etc. in the preceding interpretation, yet trying to declare it as a sublocale introduces too many conflicts at the moment. As it confusing elsewhere to figure out exactly how, in other contexts, to express the particular interpretation that is needed, to make things easier we include the following lemma. Then we can just recall the lemma to find out how to express the interpretation required in a given context. \ lemma r\<^sub>0s\<^sub>1_is_cospan: shows "cospan_of_maps_in_bicategory_of_spans V H \ \ src trg s\<^sub>1 r\<^sub>0" .. text \ The following define the projections associated with the natural tabulation of \r\<^sub>0\<^sup>* \ s\<^sub>1\. \ abbreviation p\<^sub>0 where "p\<^sub>0 \ r\<^sub>0s\<^sub>1.p\<^sub>0" abbreviation p\<^sub>1 where "p\<^sub>1 \ r\<^sub>0s\<^sub>1.p\<^sub>1" text \ $$ \xymatrix{ && {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho} && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\ {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s} } $$ \ text \ Next, we define the 2-cell that is the composite of the tabulation \\\, the tabulation \\\, and the central diamond that commutes up to unique isomorphism \\\. \ definition tab where "tab \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ (r \ \ \ p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, p\<^sub>1] \ (\ \ p\<^sub>1)" lemma tab_in_hom [intro]: shows "\tab : r\<^sub>1 \ p\<^sub>1 \ (r \ s) \ s\<^sub>0 \ p\<^sub>0\" using \.T0.antipar(1) r\<^sub>0s\<^sub>1.\_in_hom composable \.leg0_in_hom(1) \.leg1_in_hom(1) composable tab_def by auto interpretation tabulation_data V H \ \ src trg \r \ s\ tab \s\<^sub>0 \ p\<^sub>0\ \r\<^sub>1 \ p\<^sub>1\ using composable tab_in_hom by unfold_locales auto text \ In the subsequent proof we will use coherence to shortcut a few of the calculations. \ interpretation E: self_evaluation_map V H \ \ src trg .. notation E.eval ("\_\") text \ The following is applied twice in the proof of property \T2\ for the composite tabulation. It's too long to repeat. \ lemma technical: assumes "ide w" and "\\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\" and "w\<^sub>r = p\<^sub>1 \ w" and "\\<^sub>r = (s \ \) \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" shows "\.composite_cell w\<^sub>r \\<^sub>r = \[r, s, u] \ composite_cell w \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" text \ $$ \xymatrix{ && X \ar[d]^{w} \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta}} \\ && {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho} && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\ {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s} } $$ \ proof - interpret uw\: uw\ V H \ \ src trg \r \ s\ tab \s\<^sub>0 \ p\<^sub>0\ \r\<^sub>1 \ p\<^sub>1\ u w \ using assms(1-2) composable by unfold_locales auto show ?thesis proof - have "\[r, s, u] \ composite_cell w \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] = (\[r, s, u] \ ((r \ s) \ \)) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using comp_assoc by presburger also have "... = (r \ s \ \) \ \[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using assoc_naturality [of r s \] composable comp_assoc by simp also have "... = (r \ s \ \) \ \[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ ((\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0])) \ (r \ \ \ p\<^sub>0) \ \.composite_cell p\<^sub>1 r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" unfolding tab_def using comp_assoc by presburger also have "... = (r \ s \ \) \ ((\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w))) \ ((r \ \ \ p\<^sub>0) \ \.composite_cell p\<^sub>1 r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using composable \.T0.antipar(1) comp_assoc whisker_right by auto also have "... = (r \ s \ \) \ ((\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w))) \ ((r \ \ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using composable \.T0.antipar(1) whisker_right tab_def tab_in_hom(2) composable comp_assoc by force also have "... = (r \ s \ \) \ ((\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w))) \ ((r \ \ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ ((\[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]) \ (\ \ p\<^sub>1 \ w)" using assoc'_naturality [of \ p\<^sub>1 w] \.T0.antipar(1) r\<^sub>0s\<^sub>1.base_simps(2) comp_assoc by auto also have "... = (r \ s \ \) \ ((\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w))) \ ((r \ \ \ p\<^sub>0) \ w) \ (((r \ r\<^sub>0s\<^sub>1.\) \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w]) \ \.composite_cell (p\<^sub>1 \ w) \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "(\[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w] = \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]" proof - have "(\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) = \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w] \ \\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1 \ w]" using pentagon' \.T0.antipar(1) comp_assoc by simp moreover have 1: "seq (\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w)(\\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]))" using \.T0.antipar(1) by (intro seqI hseqI, auto) ultimately have "\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w] = ((\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \ \[r, r\<^sub>0, p\<^sub>1 \ w]" using \.T0.antipar(1) iso_assoc invert_side_of_triangle(2) [of "(\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])" "\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]" "\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1 \ w]"] by fastforce hence "\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w] = (\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]" using comp_assoc by presburger moreover have "seq (inv (\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w)) \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]" using \.T0.antipar(1) 1 by fastforce ultimately show ?thesis using \.T0.antipar(1) iso_assoc invert_side_of_triangle(1) [of "\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]" "\\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \ w" "\\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]"] by fastforce qed thus ?thesis using comp_assoc by presburger qed also have "... = (r \ s \ \) \ ((\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w))) \ (((r \ \ \ p\<^sub>0) \ w) \ \\<^sup>-\<^sup>1[r, s\<^sub>1 \ p\<^sub>0, w]) \ (r \ r\<^sub>0s\<^sub>1.\ \ w) \ \.composite_cell (p\<^sub>1 \ w) \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "((r \ r\<^sub>0s\<^sub>1.\) \ w) \ \\<^sup>-\<^sup>1[r, r\<^sub>0 \ p\<^sub>1, w] = \\<^sup>-\<^sup>1[r, s\<^sub>1 \ p\<^sub>0, w] \ (r \ r\<^sub>0s\<^sub>1.\ \ w)" using assoc'_naturality [of r r\<^sub>0s\<^sub>1.\ w] r\<^sub>0s\<^sub>1.cospan by auto thus ?thesis using comp_assoc by presburger qed also have "... = (r \ s \ \) \ (\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w) \ \\<^sup>-\<^sup>1[r, (s \ s\<^sub>0) \ p\<^sub>0, w]) \ (r \ (\ \ p\<^sub>0) \ w) \ (r \ r\<^sub>0s\<^sub>1.\ \ w) \ \.composite_cell (p\<^sub>1 \ w) \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "((r \ \ \ p\<^sub>0) \ w) \ \\<^sup>-\<^sup>1[r, s\<^sub>1 \ p\<^sub>0, w] = \\<^sup>-\<^sup>1[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ (r \ (\ \ p\<^sub>0) \ w)" using assoc'_naturality [of r "\ \ p\<^sub>0" w] by (simp add: composable) thus ?thesis using comp_assoc by presburger qed also have "... = (r \ s \ \) \ (r \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ ((r \ (\ \ p\<^sub>0) \ w) \ (r \ r\<^sub>0s\<^sub>1.\ \ w) \ (r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w)" proof - have "\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w) \ \\<^sup>-\<^sup>1[r, (s \ s\<^sub>0) \ p\<^sub>0, w] = r \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ \[s \ s\<^sub>0, p\<^sub>0, w]" proof - have "\[r, s, (s\<^sub>0 \ p\<^sub>0) \ w] \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0] \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ w) \ \\<^sup>-\<^sup>1[r, (s \ s\<^sub>0) \ p\<^sub>0, w] = \\<^bold>\\<^bold>[\<^bold>\r\<^bold>\, \<^bold>\s\<^bold>\, (\<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\r\<^bold>\, \<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\w\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\r\<^bold>\, (\<^bold>\s\<^bold>\ \<^bold>\ \<^bold>\s\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]\" using \_def \'_def composable by simp also have "... = \\<^bold>\r\<^bold>\ \<^bold>\ (\<^bold>\s\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^bold>\ \<^bold>\ \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]\" using composable by (intro E.eval_eqI, simp_all) also have "... = r \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ \[s \ s\<^sub>0, p\<^sub>0, w]" using \_def \'_def composable by simp finally show ?thesis by simp qed thus ?thesis using comp_assoc by presburger qed also have "... = (r \ s \ \) \ (r \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ \.composite_cell (p\<^sub>1 \ w) (((\ \ p\<^sub>0) \ w) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])" using assms(3) arrI \.T0.antipar(1) whisker_left by auto also have "... = (r \ (s \ \) \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\[s \ s\<^sub>0, p\<^sub>0, w] \ ((\ \ p\<^sub>0) \ w)) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w)" using \.T0.antipar(1) comp_assoc whisker_left by auto also have "... = (r \ (s \ \) \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w)" using assoc_naturality [of \ p\<^sub>0 w] comp_assoc by simp finally show ?thesis using assms(3-4) by simp qed qed lemma composite_is_tabulation: shows "tabulation V H \ \ src trg (r \ s) tab (s\<^sub>0 \ p\<^sub>0) (r\<^sub>1 \ p\<^sub>1)" proof show "\u \. \ ide u; \\ : dom \ \ (r \ s) \ u\ \ \ \w \ \. ide w \ \\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\ \ \\ : dom \ \ (r\<^sub>1 \ p\<^sub>1) \ w\ \ iso \ \ composite_cell w \ \ \ = \" proof - fix u \ assume u: "ide u" assume \: "\\ : dom \ \ (r \ s) \ u\" let ?v = "dom \" have 1: "\\[r, s, u] \ \ : ?v \ r \ s \ u\" proof - (* * TODO: I think this highlights the current issue with assoc_in_hom: * it can't be applied automatically here because there isn't any way to * obtain the equations src r = trg s and src s = trg u from assumption \. * Maybe this can be improved with a little bit of thought, but not while * a lot of other stuff is being changed, too. *) have "src r = trg s \ src s = trg u" by (metis \ arr_cod hseq_char in_homE hcomp_simps(1)) thus ?thesis using u \ by fastforce qed obtain w\<^sub>r \\<^sub>r \\<^sub>r where w\<^sub>r\\<^sub>r\\<^sub>r: "ide w\<^sub>r \ \\\<^sub>r : r\<^sub>0 \ w\<^sub>r \ s \ u\ \ \\\<^sub>r : ?v \ r\<^sub>1 \ w\<^sub>r\ \ iso \\<^sub>r \ \.composite_cell w\<^sub>r \\<^sub>r \ \\<^sub>r = \[r, s, u] \ \" using u \ \.T1 [of "s \ u" "\[r, s, u] \ \"] by (metis 1 \.ide_base \.ide_base arr_cod composable ide_hcomp in_homE match_1 not_arr_null seq_if_composable) text \ $$ \xymatrix{ && X \ar@ {.>}[ddl]^{w_r} \ar@/ul20pt/[dddll]_{v} \xtwocell[dddll]{}\omit{^{<1.5>\nu_r}} \ar@/ur20pt/[dddrr]^{u} \xtwocell[dddr]{}\omit{^{\theta_r}} \\ && \\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho} && \\ {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s} } $$ \ text \We need some simps, etc., otherwise the subsequent diagram chase is very painful.\ have w\<^sub>r: "ide w\<^sub>r" using w\<^sub>r\\<^sub>r\\<^sub>r by simp have [simp]: "src w\<^sub>r = src u" using w\<^sub>r\\<^sub>r\\<^sub>r \ 1 comp_arr_dom in_homE seqE hcomp_simps(1) vseq_implies_hpar(1) by (metis src_hcomp) have [simp]: "trg w\<^sub>r = src \" using w\<^sub>r\\<^sub>r\\<^sub>r by (metis 1 arrI not_arr_null seqE seq_if_composable) have \\<^sub>r_in_hom [intro]: "\\\<^sub>r : r\<^sub>0 \ w\<^sub>r \ s \ u\" using w\<^sub>r\\<^sub>r\\<^sub>r by simp have \\<^sub>r_in_hhom [intro]: "\\\<^sub>r : src u \ trg s\" using \\<^sub>r_in_hom src_cod [of \\<^sub>r] trg_cod [of \\<^sub>r] by (metis \src w\<^sub>r = src u\ \.leg1_simps(4) arr_dom in_hhomI in_homE r\<^sub>0s\<^sub>1.cospan src_hcomp trg_hcomp vconn_implies_hpar(1) vconn_implies_hpar(2)) have [simp]: "src \\<^sub>r = src u" using \\<^sub>r_in_hhom by auto have [simp]: "trg \\<^sub>r = trg s" using \\<^sub>r_in_hhom by auto have [simp]: "dom \\<^sub>r = r\<^sub>0 \ w\<^sub>r" using \\<^sub>r_in_hom by blast have [simp]: "cod \\<^sub>r = s \ u" using \\<^sub>r_in_hom by blast have \\<^sub>r_in_hom [intro]: "\\\<^sub>r : ?v \ r\<^sub>1 \ w\<^sub>r\" using w\<^sub>r\\<^sub>r\\<^sub>r by simp have \\<^sub>r_in_hhom [intro]: "\\\<^sub>r : src u \ trg r\" using \\<^sub>r_in_hom src_dom [of \\<^sub>r] trg_dom [of \\<^sub>r] by (metis \src w\<^sub>r = src u\ \.leg1_simps(4) arr_cod in_hhomI in_homE src_hcomp trg_hcomp vconn_implies_hpar(3) vconn_implies_hpar(4)) have [simp]: "src \\<^sub>r = src u" using \\<^sub>r_in_hhom by auto have [simp]: "trg \\<^sub>r = trg r" using \\<^sub>r_in_hhom by auto have [simp]: "dom \\<^sub>r = ?v" using \\<^sub>r_in_hom by auto have [simp]: "cod \\<^sub>r = r\<^sub>1 \ w\<^sub>r" using \\<^sub>r_in_hom by auto have iso_\\<^sub>r: "iso \\<^sub>r" using w\<^sub>r\\<^sub>r\\<^sub>r by simp obtain w\<^sub>s \\<^sub>s \\<^sub>s where w\<^sub>s\\<^sub>s\\<^sub>s: "ide w\<^sub>s \ \\\<^sub>s : s\<^sub>0 \ w\<^sub>s \ u\ \ \\\<^sub>s : r\<^sub>0 \ w\<^sub>r \ s\<^sub>1 \ w\<^sub>s\ \ iso \\<^sub>s \ \.composite_cell w\<^sub>s \\<^sub>s \ \\<^sub>s = \\<^sub>r" using u w\<^sub>r\\<^sub>r\\<^sub>r \.T1 [of u \\<^sub>r] by auto text \ $$ \xymatrix{ && X \ar[ddl]^{w_r} \ar@/ul20pt/[dddll]_{v} \xtwocell[dddll]{}\omit{^{<1.5>\nu_r}} \ar@/ur20pt/[dddrr]^{u} \ar@ {.>}[ddr]_{w_s} \xtwocell[dddrr]{}\omit{^{<-1.5>\theta_s}} \xtwocell[ddd]{}\omit{^{<1>\nu_s}} \\ && \\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho} && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\ {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s} } $$ \ have w\<^sub>s: "ide w\<^sub>s" using w\<^sub>s\\<^sub>s\\<^sub>s by simp have [simp]: "src w\<^sub>s = src u" using w\<^sub>s\\<^sub>s\\<^sub>s src_cod by (metis \.leg0_simps(2) \.tab_simps(2) \\<^sub>r_in_hom arrI hseqI' ideD(1) seqE seq_if_composable src_hcomp vconn_implies_hpar(3)) have [simp]: "trg w\<^sub>s = src \" using w\<^sub>s\\<^sub>s\\<^sub>s by (metis \.tab_simps(2) arr_dom in_homE not_arr_null seq_if_composable) have \\<^sub>s_in_hom [intro]: "\\\<^sub>s : s\<^sub>0 \ w\<^sub>s \ u\" using w\<^sub>s\\<^sub>s\\<^sub>s by simp have \\<^sub>s_in_hhom [intro]: "\\\<^sub>s : src u \ src s\" using \\<^sub>s_in_hom src_cod trg_cod by (metis \\<^sub>r_in_hom arrI hseqE in_hhom_def seqE vconn_implies_hpar(1) vconn_implies_hpar(3) w\<^sub>s\\<^sub>s\\<^sub>s) have [simp]: "src \\<^sub>s = src u" using \\<^sub>s_in_hhom by auto have [simp]: "trg \\<^sub>s = src s" using \\<^sub>s_in_hhom by auto have [simp]: "dom \\<^sub>s = s\<^sub>0 \ w\<^sub>s" using \\<^sub>s_in_hom by blast have [simp]: "cod \\<^sub>s = u" using \\<^sub>s_in_hom by blast have \\<^sub>s_in_hom [intro]: "\\\<^sub>s : r\<^sub>0 \ w\<^sub>r \ s\<^sub>1 \ w\<^sub>s\" using w\<^sub>s\\<^sub>s\\<^sub>s by simp have \\<^sub>s_in_hhom [intro]: "\\\<^sub>s : src u \ trg s\" using \\<^sub>s_in_hom src_dom trg_cod by (metis \src \\<^sub>r = src u\ \trg \\<^sub>r = trg s\ \\<^sub>r_in_hom in_hhomI in_homE src_dom trg_dom) have [simp]: "src \\<^sub>s = src u" using \\<^sub>s_in_hhom by auto have [simp]: "trg \\<^sub>s = trg s" using \\<^sub>s_in_hhom by auto have [simp]: "dom \\<^sub>s = r\<^sub>0 \ w\<^sub>r" using \\<^sub>s_in_hom by auto have [simp]: "cod \\<^sub>s = s\<^sub>1 \ w\<^sub>s" using \\<^sub>s_in_hom by auto have iso_\\<^sub>s: "iso \\<^sub>s" using w\<^sub>s\\<^sub>s\\<^sub>s by simp obtain w \\<^sub>t \\<^sub>t where w\\<^sub>t\\<^sub>t: "ide w \ \\\<^sub>t : p\<^sub>0 \ w \ w\<^sub>s\ \ \\\<^sub>t : w\<^sub>r \ p\<^sub>1 \ w\ \ iso \\<^sub>t \ (s\<^sub>1 \ \\<^sub>t) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t) = \\<^sub>s" using w\<^sub>r\\<^sub>r\\<^sub>r w\<^sub>s\\<^sub>s\\<^sub>s iso_\\<^sub>s r\<^sub>0s\<^sub>1.\_biuniversal_prop(1) [of w\<^sub>s w\<^sub>r \\<^sub>s] by blast text \ $$ \xymatrix{ && X \ar[ddl]_{w_r} \ar@/ul20pt/[dddll]_{v} \xtwocell[dddll]{}\omit{^{<1.5>\nu_r}} \ar@/ur20pt/[dddrr]^{u} \ar[ddr]^{w_s} \xtwocell[dddrr]{}\omit{^{<-1.5>\theta_s}} \ar@ {.>}[d]^{w} \xtwocell[ddl]{}\omit{^<-2>{\nu_t}} \xtwocell[ddr]{}\omit{^<2>{\theta_t}} \\ && {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho} && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\ {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s} } $$ \ text \{\bf Note:} \w\ is not necessarily a map.\ have w: "ide w" using w\\<^sub>t\\<^sub>t by simp have [simp]: "src w = src u" using w\\<^sub>t\\<^sub>t src_cod by (metis \\<^sub>s_in_hom \src \\<^sub>s = src u\ arrI seqE src_hcomp src_vcomp vseq_implies_hpar(1)) have [simp]: "trg w = src p\<^sub>0" using w\\<^sub>t\\<^sub>t by (metis \\<^sub>s_in_hom arrI not_arr_null r\<^sub>0s\<^sub>1.\_simps(2) seqE seq_if_composable) have \\<^sub>t_in_hom [intro]: "\\\<^sub>t : p\<^sub>0 \ w \ w\<^sub>s\" using w\\<^sub>t\\<^sub>t by simp have \\<^sub>t_in_hhom [intro]: "\\\<^sub>t : src u \ src \\" using \\<^sub>t_in_hom src_cod trg_cod \src w\<^sub>s = src u\ \trg w\<^sub>s = src \\ by fastforce have [simp]: "src \\<^sub>t = src u" using \\<^sub>t_in_hhom by auto have [simp]: "trg \\<^sub>t = src \" using \\<^sub>t_in_hhom by auto have [simp]: "dom \\<^sub>t = p\<^sub>0 \ w" using \\<^sub>t_in_hom by blast have (* [simp]: *) "cod \\<^sub>t = w\<^sub>s" using \\<^sub>t_in_hom by blast have \\<^sub>t_in_hom [intro]: "\\\<^sub>t : w\<^sub>r \ p\<^sub>1 \ w\" using w\\<^sub>t\\<^sub>t by simp have \\<^sub>t_in_hhom [intro]: "\\\<^sub>t : src u \ src \\" using \\<^sub>t_in_hom src_dom trg_dom \src w\<^sub>r = src u\ \trg w\<^sub>r = src \\ by fastforce have [simp]: "src \\<^sub>t = src u" using \\<^sub>t_in_hhom by auto have [simp]: "trg \\<^sub>t = src \" using \\<^sub>t_in_hhom by auto have (* [simp]: *) "dom \\<^sub>t = w\<^sub>r" using \\<^sub>t_in_hom by auto have [simp]: "cod \\<^sub>t = p\<^sub>1 \ w" using \\<^sub>t_in_hom by auto have iso_\\<^sub>t: "iso \\<^sub>t" using w\\<^sub>t\\<^sub>t by simp define \ where "\ = \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t) \ \[s\<^sub>0, p\<^sub>0, w]" have \: "\\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\" proof (unfold \_def, intro comp_in_homI) show "\\[s\<^sub>0, p\<^sub>0, w] : (s\<^sub>0 \ p\<^sub>0) \ w \ s\<^sub>0 \ p\<^sub>0 \ w\" using w\\<^sub>t\\<^sub>t by auto show "\s\<^sub>0 \ \\<^sub>t : s\<^sub>0 \ p\<^sub>0 \ w \ s\<^sub>0 \ w\<^sub>s\" using w\\<^sub>t\\<^sub>t by auto show "\\\<^sub>s : s\<^sub>0 \ w\<^sub>s \ u\" using w\<^sub>s\\<^sub>s\\<^sub>s by simp qed define \ where "\ = \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" have \: "\\ : ?v \ (r\<^sub>1 \ p\<^sub>1) \ w\" proof (unfold \_def, intro comp_in_homI) show "\\\<^sub>r : ?v \ r\<^sub>1 \ w\<^sub>r\" using w\<^sub>r\\<^sub>r\\<^sub>r by blast show "\r\<^sub>1 \ \\<^sub>t : r\<^sub>1 \ w\<^sub>r \ r\<^sub>1 \ p\<^sub>1 \ w\" using w\\<^sub>t\\<^sub>t by (intro hcomp_in_vhom, auto) show "\\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] : r\<^sub>1 \ p\<^sub>1 \ w \ (r\<^sub>1 \ p\<^sub>1) \ w\" using w\\<^sub>t\\<^sub>t assoc_in_hom by (simp add: \.T0.antipar(1)) qed have iso_\: "iso \" using \ w\\<^sub>t\\<^sub>t w\<^sub>r\\<^sub>r\\<^sub>r \.T0.antipar(1) by (unfold \_def, intro isos_compose) auto have *: "arr ((s \ \\<^sub>s) \ \[s, s\<^sub>0, w\<^sub>s] \ (\ \ w\<^sub>s) \ (s\<^sub>1 \ \\<^sub>t) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t))" using w\<^sub>s\\<^sub>s\\<^sub>s w\\<^sub>t\\<^sub>t \\<^sub>r_in_hom comp_assoc by auto have "((r \ s) \ \) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \ = \" proof - have "seq (r \ \\<^sub>r) (\[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r)" using w\<^sub>r\\<^sub>r\\<^sub>r \.base_simps(2) composable by fastforce hence "\ = \\<^sup>-\<^sup>1[r, s, u] \ \.composite_cell w\<^sub>r \\<^sub>r \ \\<^sub>r" using w\<^sub>r\\<^sub>r\\<^sub>r invert_side_of_triangle(1) iso_assoc by (metis 1 \.ide_base \.ide_base arrI assoc'_eq_inv_assoc composable hseq_char' seqE seq_if_composable u vconn_implies_hpar(2) vconn_implies_hpar(4) w\<^sub>s\\<^sub>s\\<^sub>s) also have "... = \\<^sup>-\<^sup>1[r, s, u] \ \.composite_cell w\<^sub>r (\.composite_cell w\<^sub>s \\<^sub>s \ \\<^sub>s) \ \\<^sub>r" using w\<^sub>s\\<^sub>s\\<^sub>s by simp also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ \[s, s\<^sub>0, w\<^sub>s] \ (\ \ w\<^sub>s) \ (s\<^sub>1 \ \\<^sub>t) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)) \ \[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r" using w\\<^sub>t\\<^sub>t comp_assoc by simp text \Rearrange to create \\\ and \\\, leaving \tab\ in the middle.\ also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ \[s, s\<^sub>0, w\<^sub>s] \ ((\ \ w\<^sub>s) \ (s\<^sub>1 \ \\<^sub>t)) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)) \ \[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r" using comp_assoc by presburger also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (\[s, s\<^sub>0, w\<^sub>s] \ ((s \ s\<^sub>0) \ \\<^sub>t)) \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)) \ \[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r" proof - have "(\ \ w\<^sub>s) \ (s\<^sub>1 \ \\<^sub>t) = \ \ \\<^sub>t" using comp_arr_dom comp_cod_arr interchange by (metis \cod \\<^sub>t = w\<^sub>s\ \.tab_simps(1) \.tab_simps(4) arrI w\\<^sub>t\\<^sub>t) also have "... = ((s \ s\<^sub>0) \ \\<^sub>t) \ (\ \ p\<^sub>0 \ w)" using comp_arr_dom comp_cod_arr interchange w\<^sub>s\\<^sub>s\\<^sub>s w\\<^sub>t\\<^sub>t \.tab_in_hom by (metis \dom \\<^sub>t = p\<^sub>0 \ w\ \.tab_simps(5) arrI) finally have "(\ \ w\<^sub>s) \ (s\<^sub>1 \ \\<^sub>t) = ((s \ s\<^sub>0) \ \\<^sub>t) \ (\ \ p\<^sub>0 \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)) \ \[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r" using assoc_naturality [of s s\<^sub>0 \\<^sub>t] w\\<^sub>t\\<^sub>t comp_assoc \cod \\<^sub>t = w\<^sub>s\ arrI by force also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)) \ \[r, r\<^sub>0, w\<^sub>r] \ (\ \ w\<^sub>r) \ \\<^sub>r" proof - have "r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t) = (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t))" proof - have "seq ((s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) (\[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t))" proof - have "seq (s \ \\<^sub>s) ((s \ s\<^sub>0 \ \\<^sub>t) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t))" using \\\[r, s, u] \ \ : dom \ \ r \ s \ u\\ calculation by blast thus ?thesis using comp_assoc by presburger qed thus ?thesis using whisker_left [of r "(s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)" "\[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ (r\<^sub>0 \ \\<^sub>t)"] w\<^sub>s\\<^sub>s\\<^sub>s w\\<^sub>t\\<^sub>t comp_assoc by simp qed thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ ((r \ r\<^sub>0 \ \\<^sub>t) \ \[r, r\<^sub>0, w\<^sub>r]) \ (\ \ w\<^sub>r) \ \\<^sub>r" proof - have "seq (\[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) (r\<^sub>0 \ \\<^sub>t)" using 1 r\<^sub>0s\<^sub>1.p\<^sub>1_simps w\\<^sub>t\\<^sub>t apply (intro seqI' comp_in_homI) by auto hence "r \ (\[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ (r\<^sub>0 \ \\<^sub>t) = (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ (r \ r\<^sub>0 \ \\<^sub>t)" using whisker_left by simp thus ?thesis using comp_assoc by simp qed also have "... = \\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (((r \ r\<^sub>0) \ \\<^sub>t) \ (\ \ w\<^sub>r)) \ \\<^sub>r" proof - have "(r \ r\<^sub>0 \ \\<^sub>t) \ \[r, r\<^sub>0, w\<^sub>r] = \[r, r\<^sub>0, p\<^sub>1 \ w] \ ((r \ r\<^sub>0) \ \\<^sub>t)" using assoc_naturality [of r r\<^sub>0 \\<^sub>t] \\<^sub>t_in_hom by auto thus ?thesis using comp_assoc by presburger qed also have "... = (\\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t))) \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "((r \ r\<^sub>0) \ \\<^sub>t) \ (\ \ w\<^sub>r) = \ \ \\<^sub>t" using comp_arr_dom comp_cod_arr interchange by (metis \dom \\<^sub>t = w\<^sub>r\ \.tab_simps(1) \.tab_simps(5) arrI w\\<^sub>t\\<^sub>t) also have "... = (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t)" using comp_arr_dom comp_cod_arr interchange by (metis \cod \\<^sub>t = p\<^sub>1 \ w\ \trg \\<^sub>t = src \\ \.T0.antipar(1) \.tab_simps(1) \.tab_simps(2) \.tab_simps(4) r\<^sub>0s\<^sub>1.base_simps(2) trg.preserves_reflects_arr trg_hcomp) finally have "((r \ r\<^sub>0) \ \\<^sub>t) \ (\ \ w\<^sub>r) = (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ ((\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w]) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "\\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w]" proof - have "seq (s \ \\<^sub>s) (s \ s\<^sub>0 \ \\<^sub>t)" using \\<^sub>s_in_hom \\<^sub>s_in_hhom \\<^sub>t_in_hom \\<^sub>t_in_hhom 1 calculation by blast moreover have "src r = trg (s \ \\<^sub>s)" using composable hseqI by force ultimately have "\\<^sup>-\<^sup>1[r, s, u] \ (r \ (s \ \\<^sub>s) \ (s \ s\<^sub>0 \ \\<^sub>t)) = (\\<^sup>-\<^sup>1[r, s, u] \ (r \ s \ \\<^sub>s)) \ (r \ s \ s\<^sub>0 \ \\<^sub>t)" using whisker_left comp_assoc by simp also have "... = ((r \ s) \ \\<^sub>s) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ w\<^sub>s] \ (r \ s \ s\<^sub>0 \ \\<^sub>t)" using assoc_naturality comp_assoc by (metis \cod \\<^sub>s = u\ \dom \\<^sub>s = s\<^sub>0 \ w\<^sub>s\ \trg \\<^sub>s = src s\ \.base_simps(2-4) \.base_simps(2-4) arrI assoc'_naturality composable w\<^sub>s\\<^sub>s\\<^sub>s) also have "... = (((r \ s) \ \\<^sub>s) \ ((r \ s) \ s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w]" proof - have "\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ w\<^sub>s] \ (r \ s \ s\<^sub>0 \ \\<^sub>t) = ((r \ s) \ s\<^sub>0 \ \\<^sub>t) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w]" using arrI hseq_char assoc'_naturality [of r s "s\<^sub>0 \ \\<^sub>t"] \cod \\<^sub>t = w\<^sub>s\ composable by auto thus ?thesis using comp_assoc by auto qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w]" using \_def \ whisker_left by (metis (full_types) arrI cod_comp ide_base seqE seqI) finally show ?thesis by simp qed thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ ((r \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ \[s \ s\<^sub>0, p\<^sub>0, w] \ ((\ \ p\<^sub>0) \ w) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \ \[r, r\<^sub>0, p\<^sub>1 \ w] \ (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "(\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] = \[s \ s\<^sub>0, p\<^sub>0, w] \ ((\ \ p\<^sub>0) \ w)" using assoc_naturality [of \ p\<^sub>0 w] by (simp add: w\\<^sub>t\\<^sub>t) thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ (r \ (\ \ p\<^sub>0) \ w) \ (r \ r\<^sub>0s\<^sub>1.\ \ w) \ ((r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]) \ (\ \ p\<^sub>1 \ w) \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" using r\<^sub>0s\<^sub>1.p\<^sub>1_simps w\\<^sub>t\\<^sub>t whisker_left comp_assoc by force also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ (r \ (\ \ p\<^sub>0) \ w) \ (r \ r\<^sub>0s\<^sub>1.\ \ w) \ (\[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ (\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]) \ (\ \ p\<^sub>1 \ w)) \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "(r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] = \[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]" proof - have 1: "(r \ \[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w) = \[r, r\<^sub>0, p\<^sub>1 \ w] \ \[r \ r\<^sub>0, p\<^sub>1, w]" using pentagon by (simp add: \.T0.antipar(1) w) moreover have 2: "seq \[r, r\<^sub>0, p\<^sub>1 \ w] \[r \ r\<^sub>0, p\<^sub>1, w]" using \.T0.antipar(1) w by simp moreover have "inv (r \ \[r\<^sub>0, p\<^sub>1, w]) = r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \.T0.antipar(1) w by simp ultimately have "\[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w) = ((r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]) \ \[r \ r\<^sub>0, p\<^sub>1, w]" using \.T0.antipar(1) w comp_assoc invert_side_of_triangle(1) [of "\[r, r\<^sub>0, p\<^sub>1 \ w] \ \[r \ r\<^sub>0, p\<^sub>1, w]" "r \ \[r\<^sub>0, p\<^sub>1, w]" "\[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w)"] by simp hence "(r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w] = (\[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w)) \ \\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w]" using \.T0.antipar(1) w invert_side_of_triangle(2) [of "\[r, r\<^sub>0 \ p\<^sub>1, w] \ (\[r, r\<^sub>0, p\<^sub>1] \ w)" "(r \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \ \[r, r\<^sub>0, p\<^sub>1 \ w]" "\[r \ r\<^sub>0, p\<^sub>1, w]"] using \trg w = src p\<^sub>0\ by simp thus ?thesis using comp_assoc by presburger qed thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ (r \ (\ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\ \ w) \ \[r, r\<^sub>0 \ p\<^sub>1, w]) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "\\<^sup>-\<^sup>1[r \ r\<^sub>0, p\<^sub>1, w] \ (\ \ p\<^sub>1 \ w) = ((\ \ p\<^sub>1) \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using assoc'_naturality [of \ p\<^sub>1 w] by (simp add: \.T0.antipar(1) w\\<^sub>t\\<^sub>t) thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ ((r \ (\ \ p\<^sub>0) \ w) \ \[r, s\<^sub>1 \ p\<^sub>0, w]) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "(r \ r\<^sub>0s\<^sub>1.\ \ w) \ \[r, r\<^sub>0 \ p\<^sub>1, w] = \[r, s\<^sub>1 \ p\<^sub>0, w] \ ((r \ r\<^sub>0s\<^sub>1.\) \ w)" using assoc_naturality [of r r\<^sub>0s\<^sub>1.\ w] r\<^sub>0s\<^sub>1.cospan w\\<^sub>t\\<^sub>t by auto thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ \[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ (((r \ \ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w)) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "(r \ (\ \ p\<^sub>0) \ w) \ \[r, s\<^sub>1 \ p\<^sub>0, w] = \[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ ((r \ \ \ p\<^sub>0) \ w)" proof - have "arr w \ dom w = w \ cod w = w" using ide_char w by blast then show ?thesis using assoc_naturality [of r "\ \ p\<^sub>0" w] composable by auto qed thus ?thesis using comp_assoc by presburger qed also have "... = ((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ \[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ w) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ w)) \ (tab \ w) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ (r\<^sub>1 \ \\<^sub>t) \ \\<^sub>r" proof - have "((r \ \ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w) = (r \ \ \ p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, p\<^sub>1] \ (\ \ p\<^sub>1) \ w" using w \.T0.antipar(1) composable whisker_right by auto also have "... = (((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0]) \ (r \ \[s, s\<^sub>0, p\<^sub>0])) \ (r \ \ \ p\<^sub>0)) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, p\<^sub>1] \ (\ \ p\<^sub>1) \ w" proof - have "((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0]) \ (r \ \[s, s\<^sub>0, p\<^sub>0])) \ (r \ \ \ p\<^sub>0) = r \ \ \ p\<^sub>0" proof - have "((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0]) \ (r \ \[s, s\<^sub>0, p\<^sub>0])) \ (r \ \ \ p\<^sub>0) = ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ ((r \ s \ s\<^sub>0 \ p\<^sub>0) \ (r \ \[s, s\<^sub>0, p\<^sub>0]))) \ (r \ \ \ p\<^sub>0)" using comp_assoc_assoc' by (simp add: composable) also have "... = ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ (r \ \[s, s\<^sub>0, p\<^sub>0])) \ (r \ \ \ p\<^sub>0)" using comp_cod_arr by (simp add: composable) also have "... = ((r \ (s \ s\<^sub>0) \ p\<^sub>0)) \ (r \ \ \ p\<^sub>0)" using whisker_left comp_assoc_assoc' assoc_in_hom hseqI' by (metis \.ide_base \.base_simps(2) \.ide_base \.ide_leg0 \.leg0_simps(2-3) \.leg1_simps(3) r\<^sub>0s\<^sub>1.ide_leg0 r\<^sub>0s\<^sub>1.leg0_simps(2) r\<^sub>0s\<^sub>1.p\<^sub>0_simps hcomp_simps(1)) also have "... = r \ \ \ p\<^sub>0" using comp_cod_arr by (simp add: composable) finally show ?thesis by blast qed thus ?thesis by simp qed also have "... = (r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ \[r, s, s\<^sub>0 \ p\<^sub>0] \ (\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0]) \ (r \ \[s, s\<^sub>0, p\<^sub>0]) \ (r \ \ \ p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, p\<^sub>1] \ (\ \ p\<^sub>1) \ w" using comp_assoc by presburger also have "... = (r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ \[r, s, s\<^sub>0 \ p\<^sub>0] \ tab \ w" using tab_def by simp also have "... = ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ w) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ w) \ (tab \ w)" using w \.T0.antipar(1) composable comp_assoc whisker_right by auto finally have "((r \ \ \ p\<^sub>0) \ w) \ ((r \ r\<^sub>0s\<^sub>1.\) \ w) \ (\[r, r\<^sub>0, p\<^sub>1] \ w) \ ((\ \ p\<^sub>1) \ w) = ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ w) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ w) \ (tab \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ ((r \ s) \ \[s\<^sub>0, p\<^sub>0, w])) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \" proof - have "\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ \[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ w) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ w) = ((r \ s) \ \[s\<^sub>0, p\<^sub>0, w]) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w]" proof - have "\\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ p\<^sub>0 \ w] \ (r \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (r \ \[s \ s\<^sub>0, p\<^sub>0, w]) \ \[r, (s \ s\<^sub>0) \ p\<^sub>0, w] \ ((r \ \\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \ w) \ (\[r, s, s\<^sub>0 \ p\<^sub>0] \ w) = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\r\<^bold>\, \<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^bold>\ \<^bold>\ \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\r\<^bold>\, (\<^bold>\s\<^bold>\ \<^bold>\ \<^bold>\s\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\w\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\r\<^bold>\, \<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\w\<^bold>\)\" using w comp_assoc \'_def \_def composable by simp also have "... = \((\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\s\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\s\<^sub>0\<^bold>\, \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\r\<^bold>\ \<^bold>\ \<^bold>\s\<^bold>\, \<^bold>\s\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\p\<^sub>0\<^bold>\, \<^bold>\w\<^bold>\\<^bold>]\" using w composable apply (intro E.eval_eqI) by simp_all also have "... = ((r \ s) \ \[s\<^sub>0, p\<^sub>0, w]) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w]" using w comp_assoc \'_def \_def composable by simp finally show ?thesis by simp qed thus ?thesis using \_def comp_assoc by simp qed also have "... = ((r \ s) \ \) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \" proof - have "((r \ s) \ \\<^sub>s \ (s\<^sub>0 \ \\<^sub>t)) \ ((r \ s) \ \[s\<^sub>0, p\<^sub>0, w]) = (r \ s) \ \" using \_def w whisker_left composable by (metis \ arrI ide_base comp_assoc) thus ?thesis using comp_assoc by presburger qed finally show "((r \ s) \ \) \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w] \ (tab \ w) \ \ = \" by simp qed thus "\w \ \. ide w \ \\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\ \ \\ : dom \ \ (r\<^sub>1 \ p\<^sub>1) \ w\ \ iso \ \ composite_cell w \ \ \ = \" using w\\<^sub>t\\<^sub>t \ \ iso_\ comp_assoc by metis qed show "\u w w' \ \' \. \ ide w; ide w'; \\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\; \\' : (s\<^sub>0 \ p\<^sub>0) \ w' \ u\; \\ : (r\<^sub>1 \ p\<^sub>1) \ w \ (r\<^sub>1 \ p\<^sub>1) \ w'\; composite_cell w \ = composite_cell w' \' \ \ \ \ \!\. \\ : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \ \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \)" proof - fix u w w' \ \' \ assume w: "ide w" assume w': "ide w'" assume \: "\\ : (s\<^sub>0 \ p\<^sub>0) \ w \ u\" assume \': "\\' : (s\<^sub>0 \ p\<^sub>0) \ w' \ u\" assume \: "\\ : (r\<^sub>1 \ p\<^sub>1) \ w \ (r\<^sub>1 \ p\<^sub>1) \ w'\" assume eq: "composite_cell w \ = composite_cell w' \' \ \" interpret uw\w'\'\: uw\w'\'\ V H \ \ src trg \r \ s\ tab \s\<^sub>0 \ p\<^sub>0\ \r\<^sub>1 \ p\<^sub>1\ u w \ w' \' \ using w w' \ \' \ eq composable tab_in_hom comp_assoc by unfold_locales auto text \ $$ \begin{array}{ll} \xymatrix{ && X \ar[d]_{w'} \xtwocell[ddl]{}\omit{^{\beta}} \ar@/ul20pt/[dddll]_<>(0.25){w}|<>(0.33)@ {>}_<>(0.5){p_1}|<>(0.67)@ {>}_<>(0.75){r_1} \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta'}}\\ && {\rm src}~\phi \ar[dl]^{p_1} \ar[dr]_{p_0} \ddtwocell\omit{^\phi} \\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho} && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\ {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s} } \\ \hspace{5cm} = \qquad \xy/50pt/ \xymatrix{ && X \ar[d]_{w} \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta}}\\ && {\rm src}~\phi \ar[dl]^{p_1} \ar[dr]_{p_0} \ddtwocell\omit{^\phi} \\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho} && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\ {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s} } \endxy \end{array} $$ \ text \ First apply property \\.T2\ using \\\\<^sub>r : r\<^sub>1 \ p\<^sub>1 \ w \ r\<^sub>1 \ p\<^sub>1 \ w'\\ (obtained by composing \\\ with associativities) and ``everything to the right'' as \\\<^sub>r\ and \\\<^sub>r'\. This yields \\\\<^sub>r : p\<^sub>1 \ w \ p\<^sub>1 \ w'\\. Next, apply property \\.T2\ to obtain \\\\<^sub>s : p\<^sub>0 \ w \ p\<^sub>0 \ w'\\. For this use \\\\<^sub>s : s\<^sub>0 \ p\<^sub>0 \ w \ u\\ and \\\\<^sub>s' : s\<^sub>0 \ p\<^sub>0 \ w'\\ obtained by composing \\\ and \\'\ with associativities. We also need \\\\<^sub>s : s\<^sub>1 \ p\<^sub>0 \ w \ s\<^sub>1 \ p\<^sub>0 \ w'\\. To get this, transport \r\<^sub>0 \ \\<^sub>r\ across \\\; we need \\\ to be an isomorphism in order to do this. Finally, apply the biuniversal property of \\\ to get \\\ : w \ w'\\ and verify the required equation. \ let ?w\<^sub>r = "p\<^sub>1 \ w" have w\<^sub>r: "ide ?w\<^sub>r" by simp let ?w\<^sub>r' = "p\<^sub>1 \ w'" have w\<^sub>r': "ide ?w\<^sub>r'" by simp define \\<^sub>r where "\\<^sub>r = (s \ \) \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" have \\<^sub>r: "\\\<^sub>r : r\<^sub>0 \ ?w\<^sub>r \ s \ u\" unfolding \\<^sub>r_def using \.T0.antipar(1) by fastforce define \\<^sub>r' where "\\<^sub>r' = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']" have \\<^sub>r': "\\\<^sub>r' : r\<^sub>0 \ ?w\<^sub>r' \ s \ u\" unfolding \\<^sub>r'_def using \.T0.antipar(1) by fastforce define \\<^sub>r where "\\<^sub>r = \[r\<^sub>1, p\<^sub>1, w'] \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" have \\<^sub>r: "\\\<^sub>r : r\<^sub>1 \ ?w\<^sub>r \ r\<^sub>1 \ ?w\<^sub>r'\" unfolding \\<^sub>r_def using \.T0.antipar(1) by force have eq\<^sub>r: "\.composite_cell ?w\<^sub>r \\<^sub>r = \.composite_cell ?w\<^sub>r' \\<^sub>r' \ \\<^sub>r" text \ $$ \begin{array}{ll} \xymatrix{ && X \ar[ddl]^{w_r'} \xtwocell[dddll]{}\omit{^<2>{\beta_r}} \ar@/ul20pt/[dddll]_<>(0.33){w_r}|<>(0.67)@ {>}_<>(0.75){r_1} \ar@/ur20pt/[dddrr]^{u} \xtwocell[dddr]{}\omit{^{\theta_r'}}\\ && \\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho} && \\ {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s} } \\ \hspace{5cm} =\qquad \xy/50pt/ \xymatrix{ && X \ar[ddl]^{w_r} \ar@/ur20pt/[dddrr]^{u} \xtwocell[dddr]{}\omit{^{\theta_r}}\\ && \\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho} && \\ {\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s} } \endxy \end{array} $$ \ proof - have "\.composite_cell ?w\<^sub>r \\<^sub>r = \[r, s, u] \ composite_cell w \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using \\<^sub>r_def technical uw\w'\'\.uw\.uw\ by blast also have "... = \[r, s, u] \ (((r \ s) \ \') \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w'] \ (tab \ w') \ \) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using eq comp_assoc by simp also have "... = (r \ \\<^sub>r') \ \[r, r\<^sub>0, ?w\<^sub>r'] \ (\ \ ?w\<^sub>r') \ \\<^sub>r" proof - have "\[r, s, u] \ (composite_cell w' \' \ \) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] = \[r, s, u] \ composite_cell w' \' \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using comp_assoc by presburger also have "... = (r \ (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \ \[r, r\<^sub>0, p\<^sub>1 \ w'] \ (\ \ p\<^sub>1 \ w') \ \[r\<^sub>1, p\<^sub>1, w'] \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" proof - have "\[r, s, u] \ composite_cell w' \' \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] = \[r, s, u] \ composite_cell w' \' \ ((\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \[r\<^sub>1, p\<^sub>1, w']) \ \) \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" proof - have "(\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \[r\<^sub>1, p\<^sub>1, w']) \ \ = \" using comp_cod_arr \.T0.antipar(1) \ comp_assoc_assoc' by simp thus ?thesis by argo qed also have "... = (\[r, s, u] \ ((r \ s) \ \') \ \[r \ s, s\<^sub>0 \ p\<^sub>0, w'] \ (tab \ w') \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w']) \ \[r\<^sub>1, p\<^sub>1, w'] \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using comp_assoc by presburger also have "... = ((r \ (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \ \[r, r\<^sub>0, p\<^sub>1 \ w'] \ (\ \ p\<^sub>1 \ w')) \ \[r\<^sub>1, p\<^sub>1, w'] \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using \\<^sub>r'_def technical [of w' \' u ?w\<^sub>r' \\<^sub>r'] comp_assoc by fastforce finally show ?thesis using comp_assoc by simp qed finally show ?thesis using \\<^sub>r'_def \\<^sub>r_def comp_assoc by auto qed finally show ?thesis using comp_assoc by presburger qed have 1: "\!\. \\ : ?w\<^sub>r \ ?w\<^sub>r'\ \ \\<^sub>r = \\<^sub>r' \ (r\<^sub>0 \ \) \ \\<^sub>r = r\<^sub>1 \ \" using eq\<^sub>r \.T2 [of ?w\<^sub>r ?w\<^sub>r' \\<^sub>r "s \ u" \\<^sub>r' \\<^sub>r] w\<^sub>r w\<^sub>r' \\<^sub>r \\<^sub>r' \\<^sub>r by blast obtain \\<^sub>r where \\<^sub>r: "\\\<^sub>r : ?w\<^sub>r \ ?w\<^sub>r'\ \ \\<^sub>r = \\<^sub>r' \ (r\<^sub>0 \ \\<^sub>r) \ \\<^sub>r = r\<^sub>1 \ \\<^sub>r" using 1 by blast let ?w\<^sub>s = "p\<^sub>0 \ w" have w\<^sub>s: "ide ?w\<^sub>s" by simp let ?w\<^sub>s' = "p\<^sub>0 \ w'" have w\<^sub>s': "ide ?w\<^sub>s'" by simp define \\<^sub>s where "\\<^sub>s = \ \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]" have \\<^sub>s: "\\\<^sub>s : s\<^sub>0 \ p\<^sub>0 \ w \ u\" using \\<^sub>s_def by auto define \\<^sub>s' where "\\<^sub>s' = \' \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']" have \\<^sub>s': "\\\<^sub>s' : s\<^sub>0 \ p\<^sub>0 \ w' \ u\" using \\<^sub>s'_def by auto define \\<^sub>s where "\\<^sub>s = \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" have \\<^sub>s: "\\\<^sub>s : s\<^sub>1 \ ?w\<^sub>s \ s\<^sub>1 \ ?w\<^sub>s'\" unfolding \\<^sub>s_def using \\<^sub>r r\<^sub>0s\<^sub>1.\_in_hom(2) r\<^sub>0s\<^sub>1.\_uniqueness(2) \.T0.antipar(1) apply (intro comp_in_homI) apply auto by auto have eq\<^sub>s: "\.composite_cell (p\<^sub>0 \ w) (\ \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) = \.composite_cell (p\<^sub>0 \ w') (\' \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \\<^sub>s" text \ $$ \begin{array}{ll} \xy/67pt/ \xymatrix{ && X \ar[d]^{w'} \ar@/l10pt/[dl]_{w} \ddltwocell\omit{^{\gamma_r}} \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta_s'}}\\ & {\rm src}~\phi \ar[dr]_{p_1} \ar[d]_{p_0} & {\rm src}~\phi \ar[d]^{p_1} \ar[dr]_{p_0} \ddrtwocell\omit{^\phi} \xtwocell[ddl]{}\omit{^\;\;\;\;\phi^{-1}} \\ & {\rm src}~\sigma \ar[dr]_{s_1} & {\rm src}~\rho \ar[d]^{r_0} & {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\ && {\rm src}~r = {\rm trg}~s && {\rm src}~s \ar[ll]^{s} } \endxy \\ \hspace{5cm}= \xy/50pt/ \xymatrix{ & X \ar@/dl15pt/[ddr]_<>(0.5){w_s} \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta_s}}\\ & \\ && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\ & {\rm src}~r = {\rm trg}~s && {\rm src}~s \ar[ll]^{s} } \endxy \end{array} $$ \ proof - have "\.composite_cell (p\<^sub>0 \ w') (\' \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \\<^sub>s = (\\<^sub>r' \ (r\<^sub>0 \ \\<^sub>r)) \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using \\<^sub>s_def \\<^sub>r'_def whisker_left comp_assoc by simp also have "... = \\<^sub>r \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using \\<^sub>r by simp also have "... = ((s \ \) \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])) \ \[s, s\<^sub>0, ?w\<^sub>s] \ (\ \ ?w\<^sub>s) \ \[s\<^sub>1, p\<^sub>0, w] \ ((r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w]) \ (inv r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using \\<^sub>r_def comp_assoc by simp also have "... = (s \ \) \ \.composite_cell (p\<^sub>0 \ w) \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]" proof - have "(\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ ((r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w]) \ (inv r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] = \ \ p\<^sub>0 \ w" proof - have "\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w] = cod (inv r\<^sub>0s\<^sub>1.\ \ w)" using r\<^sub>0s\<^sub>1.\_uniqueness(2) \.T0.antipar(1) comp_assoc_assoc' by simp text \Here the fact that \\\ is a retraction is used.\ moreover have "(r\<^sub>0s\<^sub>1.\ \ w) \ (inv r\<^sub>0s\<^sub>1.\ \ w) = cod \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using r\<^sub>0s\<^sub>1.\_uniqueness(2) comp_arr_inv' whisker_right [of w r\<^sub>0s\<^sub>1.\ "inv r\<^sub>0s\<^sub>1.\"] by simp moreover have "\[s\<^sub>1, p\<^sub>0, w] \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] = dom (\ \ p\<^sub>0 \ w)" using r\<^sub>0s\<^sub>1.base_simps(2) hseq_char comp_assoc_assoc' by auto moreover have "hseq (inv r\<^sub>0s\<^sub>1.\) w" using r\<^sub>0s\<^sub>1.\_uniqueness(2) by (intro hseqI, auto) moreover have "hseq \ (p\<^sub>0 \ w)" by (intro hseqI, auto) ultimately show ?thesis using comp_arr_dom comp_cod_arr by simp qed thus ?thesis using comp_assoc by simp qed also have "... = \.composite_cell (p\<^sub>0 \ w) (\ \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])" using \\<^sub>s_def whisker_left by (metis \.ide_base \\<^sub>s arrI comp_assoc) finally show ?thesis by simp qed hence 2: "\!\. \\ : ?w\<^sub>s \ ?w\<^sub>s'\ \ \\<^sub>s = \\<^sub>s' \ (s\<^sub>0 \ \) \ \\<^sub>s = s\<^sub>1 \ \" using \.T2 [of ?w\<^sub>s ?w\<^sub>s' \\<^sub>s u \\<^sub>s' \\<^sub>s] w\<^sub>s w\<^sub>s' \\<^sub>s \\<^sub>s' \\<^sub>s by (metis \\<^sub>s'_def \\<^sub>s_def) obtain \\<^sub>s where \\<^sub>s: "\\\<^sub>s : ?w\<^sub>s \ ?w\<^sub>s'\ \ \\<^sub>s = \\<^sub>s' \ (s\<^sub>0 \ \\<^sub>s) \ \\<^sub>s = s\<^sub>1 \ \\<^sub>s" using 2 by blast have eq\<^sub>t: "(s\<^sub>1 \ \\<^sub>s) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] = (s\<^sub>1 \ ?w\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>r)" text \ $$ \xy/78pt/ \xymatrix{ & X \ar[d]^{w'} \ar@/ul15pt/[ddl]_{w_r} \xtwocell[ddl]{}\omit{^{\gamma_r}} \\ & {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\ {\rm src}~\rho \ar[dr]^{r_0} && {\rm src}~\sigma \ar[dl]_{s_1} \\ & {\rm src}~r = {\rm trg}~s & } \endxy \qquad = \qquad \xy/78pt/ \xymatrix{ & X \ar[d]^{w} \ar@/ur15pt/[ddr]^{w_s'} \xtwocell[ddr]{}\omit{^{\gamma_s}} \\ & {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\ {\rm src}~\rho \ar[dr]^{r_0} && {\rm src}~\sigma \ar[dl]_{s_1} \\ & {\rm src}~r = {\rm trg}~s & } \endxy $$ \ proof - have "(s\<^sub>1 \ ?w\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>r) = \\<^sub>s \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "\\<^sub>s \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] = (\[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \ (r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ ((inv r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \ \[s\<^sub>1, p\<^sub>0, w]) \ (r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \\<^sub>s_def comp_assoc by metis also have "... = (\[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \ (r\<^sub>0 \ \\<^sub>r)" proof - have "(r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ ((inv r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \ \[s\<^sub>1, p\<^sub>0, w]) \ (r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] = r\<^sub>0 \ \\<^sub>r" proof - have "(r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ ((inv r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \ \[s\<^sub>1, p\<^sub>0, w]) \ (r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] = (r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ ((inv r\<^sub>0s\<^sub>1.\ \ w) \ (r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using r\<^sub>0s\<^sub>1.\_uniqueness(2) comp_assoc_assoc' comp_cod_arr by simp (* Used here that \ is a section. *) also have "... = (r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using r\<^sub>0s\<^sub>1.\_uniqueness(2) comp_inv_arr' \.T0.antipar(1) whisker_right [of w "inv r\<^sub>0s\<^sub>1.\" r\<^sub>0s\<^sub>1.\] comp_cod_arr by simp also have "... = r\<^sub>0 \ \\<^sub>r" proof - have "hseq r\<^sub>0 \\<^sub>r" using \\<^sub>s \\<^sub>s_def by blast thus ?thesis using comp_assoc_assoc' comp_arr_dom by (metis (no_types) \\<^sub>r \.ide_leg0 comp_assoc_assoc'(1) hcomp_simps(3) hseq_char ide_char in_homE r\<^sub>0s\<^sub>1.ide_leg1 r\<^sub>0s\<^sub>1.p\<^sub>1_simps w w\<^sub>r) qed finally show ?thesis by blast qed thus ?thesis by simp qed also have "... = \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>r)" using comp_assoc by presburger also have "... = (s\<^sub>1 \ ?w\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>r)" proof - have "(s\<^sub>1 \ ?w\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w'] = \[s\<^sub>1, p\<^sub>0, w']" using comp_cod_arr by simp thus ?thesis using comp_assoc by metis qed finally show ?thesis by simp qed also have "... = (s\<^sub>1 \ \\<^sub>s) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \\<^sub>s by simp finally show ?thesis by simp qed have 3: "\!\. \\ : w \ w'\ \ \\<^sub>s = (p\<^sub>0 \ w') \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = \\<^sub>r" using w w' w\<^sub>s' w\<^sub>r \\<^sub>r \\<^sub>s eq\<^sub>t r\<^sub>0s\<^sub>1.\_biuniversal_prop(2) [of ?w\<^sub>s' ?w\<^sub>r w w' \\<^sub>s "p\<^sub>0 \ w'" \\<^sub>r] by blast obtain \ where \: "\\ : w \ w'\ \ \\<^sub>s = (p\<^sub>0 \ w') \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = \\<^sub>r" using 3 by blast show "\!\. \\ : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \ \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \)" proof - have "\\. \\ : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \ \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \)" proof - have "\ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \)" proof - have "\' \ ((s\<^sub>0 \ p\<^sub>0) \ \) = (\\<^sub>s' \ \[s\<^sub>0, p\<^sub>0, w']) \ ((s\<^sub>0 \ p\<^sub>0) \ \)" using \\<^sub>s'_def comp_arr_dom comp_assoc comp_assoc_assoc'(2) by auto also have "... = (\\<^sub>s' \ (s\<^sub>0 \ p\<^sub>0 \ \)) \ \[s\<^sub>0, p\<^sub>0, w]" using assoc_naturality [of s\<^sub>0 p\<^sub>0 \] comp_assoc by (metis \ \\<^sub>r \.leg0_simps(4-5) r\<^sub>0s\<^sub>1.leg0_simps(4-5) r\<^sub>0s\<^sub>1.leg1_simps(3) hseqE in_homE leg0_simps(2)) also have "... = \\<^sub>s \ \[s\<^sub>0, p\<^sub>0, w]" by (metis \ \\<^sub>s arrI comp_ide_arr w\<^sub>s') also have "... = \" using \\<^sub>s_def comp_assoc comp_arr_dom comp_assoc_assoc' by simp finally show ?thesis by simp qed moreover have "\ = (r\<^sub>1 \ p\<^sub>1) \ \" proof - have "\ = \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \\<^sub>r \ \[r\<^sub>1, p\<^sub>1, w]" proof - have "\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \\<^sub>r \ \[r\<^sub>1, p\<^sub>1, w] = (\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \[r\<^sub>1, p\<^sub>1, w']) \ \ \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \ \[r\<^sub>1, p\<^sub>1, w]" using \\<^sub>r_def comp_assoc by simp also have "... = \" using comp_arr_dom comp_cod_arr by (metis \.ide_leg1 r\<^sub>0s\<^sub>1.ide_leg1 comp_assoc_assoc'(2) hseqE ideD(1) uw\w'\'\.\_simps(1) uw\w'\'\.\_simps(4-5) leg1_simps(2) w w' w\<^sub>r w\<^sub>r') finally show ?thesis by simp qed also have "... = \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ (r\<^sub>1 \ \\<^sub>r) \ \[r\<^sub>1, p\<^sub>1, w]" using \\<^sub>r by simp also have "... = \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \[r\<^sub>1, p\<^sub>1, w'] \ ((r\<^sub>1 \ p\<^sub>1) \ \)" using assoc_naturality [of r\<^sub>1 p\<^sub>1 \] by (metis \ \\<^sub>r \.ide_leg1 r\<^sub>0s\<^sub>1.leg1_simps(5-6) hseqE ide_char in_homE leg1_simps(2)) also have "... = (\\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ \[r\<^sub>1, p\<^sub>1, w']) \ ((r\<^sub>1 \ p\<^sub>1) \ \)" using comp_assoc by presburger also have "... = (r\<^sub>1 \ p\<^sub>1) \ \" using comp_cod_arr by (metis \.ide_leg1 r\<^sub>0s\<^sub>1.ide_leg1 calculation comp_assoc_assoc'(2) comp_ide_arr hseqE ideD(1) ide_cod local.uw\w'\'\.\_simps(1) local.uw\w'\'\.\_simps(5) w' w\<^sub>r') finally show ?thesis by simp qed ultimately show "\\. \\ : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \ \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \)" using \ by blast qed moreover have "\\'. \\' : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \' \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \') \ \' = \" proof - fix \' assume \': "\\' : w \ w'\ \ \ = (r\<^sub>1 \ p\<^sub>1) \ \' \ \ = \' \ ((s\<^sub>0 \ p\<^sub>0) \ \')" show "\' = \" proof - let ?P\<^sub>r = "\\. \\ : ?w\<^sub>r \ ?w\<^sub>r'\ \ \\<^sub>r = \\<^sub>r' \ (r\<^sub>0 \ \) \ \\<^sub>r = r\<^sub>1 \ \" let ?P\<^sub>s = "\\. \\ : ?w\<^sub>s \ ?w\<^sub>s'\ \ \\<^sub>s = \\<^sub>s' \ (s\<^sub>0 \ \) \ \\<^sub>s = s\<^sub>1 \ \" let ?\\<^sub>r' = "p\<^sub>1 \ \'" let ?\\<^sub>s' = "p\<^sub>0 \ \'" let ?P\<^sub>t = "\\. \\ : w \ w'\ \ \\<^sub>s = (p\<^sub>0 \ w') \ (p\<^sub>0 \ \) \ p\<^sub>1 \ \ = \\<^sub>r" have "hseq p\<^sub>0 \'" proof (intro hseqI) show "\\' : src \ \ src p\<^sub>0\" using \' by (metis hseqE hseqI' in_hhom_def uw\w'\'\.\_simps(1) src_hcomp src_vcomp leg0_simps(2) leg1_simps(3) uw\w'\'\.uw\.\_simps(1) vseq_implies_hpar(1)) show "\p\<^sub>0 : src p\<^sub>0 \ src s\<^sub>0\" by simp qed hence "hseq p\<^sub>1 \'" using hseq_char by simp have "\?\\<^sub>r' : ?w\<^sub>r \ ?w\<^sub>r'\" using \' by auto moreover have "\\<^sub>r = \\<^sub>r' \ (r\<^sub>0 \ ?\\<^sub>r')" proof - text \ Note that @{term \\<^sub>r} is the composite of ``everything to the right'' of @{term "\ \ ?w\<^sub>r"}, and similarly for @{term \\<^sub>r'}. We can factor @{term \\<^sub>r} as @{term "(s \ \) \ X w"}, where @{term "X w"} is a composite of @{term \} and @{term \}. We can similarly factor @{term \\<^sub>r'} as @{term "(s \ \') \ X w'"}. Then @{term "\\<^sub>r' \ (r\<^sub>0 \ ?\\<^sub>r') = (s \ \') \ X w' \ (r\<^sub>0 \ ?\\<^sub>r')"}, which equals @{term "(s \ \') \ (s \ (s\<^sub>0 \ p\<^sub>0) \ ?\\<^sub>r') \ X w = \\<^sub>r"}. \ let ?X = "\w. (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" have "\\<^sub>r' \ (r\<^sub>0 \ ?\\<^sub>r') = (s \ \') \ ?X w' \ (r\<^sub>0 \ ?\\<^sub>r')" using \\<^sub>r'_def comp_assoc by simp also have "... = (s \ \') \ (s \ (s\<^sub>0 \ p\<^sub>0) \ \') \ ?X w" proof - have "(s \ \') \ ((s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \ (r\<^sub>0 \ p\<^sub>1 \ \') = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ p\<^sub>1 \ \')" using comp_assoc by presburger also have "... = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ \[s\<^sub>1, p\<^sub>0, w'] \ ((r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \')) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using assoc'_naturality [of r\<^sub>0 p\<^sub>1 \'] comp_assoc by (metis \' \\p\<^sub>1 \ \' : p\<^sub>1 \ w \ p\<^sub>1 \ w'\\ \.T0.antipar(1) \.leg0_in_hom(2) r\<^sub>0s\<^sub>1.leg1_simps(4-6) r\<^sub>0s\<^sub>1.base_simps(2) hcomp_in_vhomE in_homE trg_hcomp) also have "... = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ (\ \ p\<^sub>0 \ w') \ (\[s\<^sub>1, p\<^sub>0, w'] \ ((s\<^sub>1 \ p\<^sub>0) \ \')) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "(r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \') = r\<^sub>0s\<^sub>1.\ \ \'" using \' interchange [of r\<^sub>0s\<^sub>1.\ "r\<^sub>0 \ p\<^sub>1" w' \'] comp_arr_dom comp_cod_arr by auto also have "... = ((s\<^sub>1 \ p\<^sub>0) \ \') \ (r\<^sub>0s\<^sub>1.\ \ w)" using \' interchange \hseq p\<^sub>0 \'\ comp_arr_dom comp_cod_arr by (metis comp_arr_ide r\<^sub>0s\<^sub>1.\_simps(1,5) seqI' uw\w'\'\.uw\.w_in_hom(2) w) finally have "(r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \') = ((s\<^sub>1 \ p\<^sub>0) \ \') \ (r\<^sub>0s\<^sub>1.\ \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ \[s, s\<^sub>0, p\<^sub>0 \ w'] \ ((\ \ p\<^sub>0 \ w') \ (s\<^sub>1 \ p\<^sub>0 \ \')) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \' assoc_naturality [of s\<^sub>1 p\<^sub>0 \'] comp_assoc by (metis \.leg1_simps(2) \.leg1_simps(3,5-6) r\<^sub>0s\<^sub>1.leg0_simps(4-5) hcomp_in_vhomE hseqE in_homE uw\w'\'\.\_simps(1) leg0_in_hom(2) leg1_simps(3)) also have "... = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ (\[s, s\<^sub>0, p\<^sub>0 \ w'] \ ((s \ s\<^sub>0) \ p\<^sub>0 \ \')) \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "(\ \ p\<^sub>0 \ w') \ (s\<^sub>1 \ p\<^sub>0 \ \') = \ \ p\<^sub>0 \ \'" using \' interchange [of \ s\<^sub>1 "p\<^sub>0 \ w'" "p\<^sub>0 \ \'"] whisker_left \hseq p\<^sub>0 \'\comp_arr_dom comp_cod_arr by (metis \.tab_simps(1) \.tab_simps(4) hcomp_simps(4) in_homE r\<^sub>0s\<^sub>1.leg0_simps(5)) also have "... = ((s \ s\<^sub>0) \ p\<^sub>0 \ \') \ (\ \ p\<^sub>0 \ w)" using \' interchange [of "s \ s\<^sub>0" \ "p\<^sub>0 \ \'" "p\<^sub>0 \ w"] whisker_left comp_arr_dom comp_cod_arr \hseq p\<^sub>0 \'\ by (metis \.tab_simps(1) \.tab_simps(5) hcomp_simps(3) in_homE r\<^sub>0s\<^sub>1.leg0_simps(4)) finally have "(\ \ p\<^sub>0 \ w') \ (s\<^sub>1 \ p\<^sub>0 \ \') = ((s \ s\<^sub>0) \ p\<^sub>0 \ \') \ (\ \ p\<^sub>0 \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (s \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ ((s \ s\<^sub>0 \ p\<^sub>0 \ \') \ \[s, s\<^sub>0, p\<^sub>0 \ w]) \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \' assoc_naturality [of s s\<^sub>0 "p\<^sub>0 \ \'"] \hseq p\<^sub>0 \'\ by force also have "... = (s \ \') \ ((s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ (s \ s\<^sub>0 \ p\<^sub>0 \ \')) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using comp_assoc by presburger also have "... = (s \ \') \ ((s \ (s\<^sub>0 \ p\<^sub>0) \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])) \ \[s, s\<^sub>0, p\<^sub>0 \ w] \ (\ \ p\<^sub>0 \ w) \ \[s\<^sub>1, p\<^sub>0, w] \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" proof - have "(s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ (s \ s\<^sub>0 \ p\<^sub>0 \ \') = (s \ (s\<^sub>0 \ p\<^sub>0) \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])" proof - have "(s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \ (s \ s\<^sub>0 \ p\<^sub>0 \ \') = s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] \ (s\<^sub>0 \ p\<^sub>0 \ \')" proof - have "seq \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] (s\<^sub>0 \ p\<^sub>0 \ \')" proof (* It seems to be too time-consuming for auto to solve these. *) show "\s\<^sub>0 \ p\<^sub>0 \ \' : s\<^sub>0 \ p\<^sub>0 \ w \ s\<^sub>0 \ p\<^sub>0 \ w'\" using \' by (intro hcomp_in_vhom, auto) show "\\\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] : s\<^sub>0 \ p\<^sub>0 \ w' \ (s\<^sub>0 \ p\<^sub>0) \ w'\" by auto qed thus ?thesis using w w' \' whisker_left by simp qed also have "... = s \ ((s\<^sub>0 \ p\<^sub>0) \ \') \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]" using \' \hseq p\<^sub>0 \'\ assoc'_naturality [of s\<^sub>0 p\<^sub>0 \'] by fastforce also have "... = (s \ (s\<^sub>0 \ p\<^sub>0) \ \') \ (s \ \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])" proof - have "seq ((s\<^sub>0 \ p\<^sub>0) \ \') \\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]" proof (* Same here. *) show "\\\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w] : s\<^sub>0 \ p\<^sub>0 \ w \ (s\<^sub>0 \ p\<^sub>0) \ w\" by auto show "\(s\<^sub>0 \ p\<^sub>0) \ \' : (s\<^sub>0 \ p\<^sub>0) \ w \ (s\<^sub>0 \ p\<^sub>0) \ w'\" using \' by (intro hcomp_in_vhom, auto) qed thus ?thesis using w w' \' whisker_left by simp qed finally show ?thesis by blast qed thus ?thesis by simp qed also have "... = (s \ \') \ (s \ (s\<^sub>0 \ p\<^sub>0) \ \') \ ?X w" using comp_assoc by presburger finally show ?thesis by simp qed also have "... = \\<^sub>r" using \\<^sub>r_def \' uw\w'\'\.uw\.\_simps(1) whisker_left \.ide_base comp_assoc by simp finally show ?thesis by simp qed moreover have "\\<^sub>r = r\<^sub>1 \ ?\\<^sub>r'" proof - have "\\<^sub>r = \[r\<^sub>1, p\<^sub>1, w'] \ ((r\<^sub>1 \ p\<^sub>1) \ \') \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]" using \\<^sub>r_def \' by simp also have "... = \[r\<^sub>1, p\<^sub>1, w'] \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \ (r\<^sub>1 \ p\<^sub>1 \ \')" using \' assoc'_naturality by (metis \.leg1_simps(5-6) r\<^sub>0s\<^sub>1.leg1_simps(5-6) hcomp_in_vhomE hseqE in_homE uw\w'\'\.\_simps(1) leg1_in_hom(2)) also have "... = (\[r\<^sub>1, p\<^sub>1, w'] \ \\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w']) \ (r\<^sub>1 \ p\<^sub>1 \ \')" using comp_assoc by presburger also have "... = r\<^sub>1 \ p\<^sub>1 \ \'" using comp_cod_arr by (metis (no_types, lifting) \\<^sub>r \.ide_leg1 r\<^sub>0s\<^sub>1.ide_leg1 arrI calculation comp_assoc_assoc'(1) comp_ide_arr ide_hcomp hseq_char' ideD(1) seq_if_composable hcomp_simps(2) leg1_simps(2) w' w\<^sub>r') finally show ?thesis by simp qed ultimately have P\<^sub>r': "?P\<^sub>r ?\\<^sub>r'" by simp have eq\<^sub>r: "\\<^sub>r = ?\\<^sub>r'" using 1 \\<^sub>r P\<^sub>r' by blast have "\?\\<^sub>s' : ?w\<^sub>s \ ?w\<^sub>s'\" using \' by auto moreover have "\\<^sub>s = \\<^sub>s' \ (s\<^sub>0 \ ?\\<^sub>s')" using \' \hseq p\<^sub>0 \'\ \.leg0_simps(2,4-5) \.leg1_simps(3) \\<^sub>s'_def \\<^sub>s_def assoc'_naturality hseqE in_homE comp_assoc r\<^sub>0s\<^sub>1.leg0_simps(4-5) r\<^sub>0s\<^sub>1.p\<^sub>0_simps by metis moreover have "\\<^sub>s = s\<^sub>1 \ ?\\<^sub>s'" proof - have "\[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ \\<^sub>r) \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] = \[s\<^sub>1, p\<^sub>0, w'] \ (r\<^sub>0s\<^sub>1.\ \ w') \ (\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ p\<^sub>1 \ \')) \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using eq\<^sub>r comp_assoc by simp also have "... = \[s\<^sub>1, p\<^sub>0, w'] \ ((r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \')) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" proof - have "\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \ (r\<^sub>0 \ p\<^sub>1 \ \') = ((r\<^sub>0 \ p\<^sub>1) \ \') \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]" using \' assoc'_naturality \hseq p\<^sub>1 \'\ by (metis \.leg0_simps(2,4-5) \.leg1_simps(3) r\<^sub>0s\<^sub>1.leg1_simps(5-6) hseqE in_homE leg1_simps(2)) thus ?thesis using comp_assoc by presburger qed also have "... = (\[s\<^sub>1, p\<^sub>0, w'] \ ((s\<^sub>1 \ p\<^sub>0) \ \')) \ (r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w] \ (inv r\<^sub>0s\<^sub>1.\ \ w) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" proof - have "(r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \') = r\<^sub>0s\<^sub>1.\ \ \'" using \' interchange [of r\<^sub>0s\<^sub>1.\ "r\<^sub>0 \ p\<^sub>1" w' \'] comp_arr_dom comp_cod_arr by auto also have "... = ((s\<^sub>1 \ p\<^sub>0) \ \') \ (r\<^sub>0s\<^sub>1.\ \ w)" using \' interchange \hseq p\<^sub>0 \'\ comp_arr_dom comp_cod_arr by (metis in_homE r\<^sub>0s\<^sub>1.\_simps(1,5)) finally have "(r\<^sub>0s\<^sub>1.\ \ w') \ ((r\<^sub>0 \ p\<^sub>1) \ \') = ((s\<^sub>1 \ p\<^sub>0) \ \') \ (r\<^sub>0s\<^sub>1.\ \ w)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (s\<^sub>1 \ ?\\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w] \ ((r\<^sub>0s\<^sub>1.\ \ w) \ (\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w]) \ (inv r\<^sub>0s\<^sub>1.\ \ w)) \ \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" proof - have "\[s\<^sub>1, p\<^sub>0, w'] \ ((s\<^sub>1 \ p\<^sub>0) \ \') = (s\<^sub>1 \ ?\\<^sub>s') \ \[s\<^sub>1, p\<^sub>0, w]" using \' assoc_naturality [of s\<^sub>1 p\<^sub>0 \'] \hseq p\<^sub>0 \'\ by auto thus ?thesis using comp_assoc by presburger qed also have "... = s\<^sub>1 \ ?\\<^sub>s'" proof - have "\\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \ \[r\<^sub>0, p\<^sub>1, w] = cod (inv r\<^sub>0s\<^sub>1.\ \ w)" using r\<^sub>0s\<^sub>1.\_uniqueness(2) \.T0.antipar(1) comp_assoc_assoc' by simp text \Here the fact that \\\ is a retraction is used.\ moreover have "(r\<^sub>0s\<^sub>1.\ \ w) \ (inv r\<^sub>0s\<^sub>1.\ \ w) = cod \\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]" using r\<^sub>0s\<^sub>1.\_uniqueness(2) comp_arr_inv' whisker_right [of w r\<^sub>0s\<^sub>1.\ "inv r\<^sub>0s\<^sub>1.\"] by simp moreover have "cod (inv r\<^sub>0s\<^sub>1.\ \ w) \ (inv r\<^sub>0s\<^sub>1.\ \ w) = inv r\<^sub>0s\<^sub>1.\ \ w" using \\<^sub>s_def \\<^sub>s by (meson arrI comp_cod_arr seqE) ultimately show ?thesis using \' \hseq p\<^sub>0 \'\ comp_arr_dom comp_cod_arr comp_assoc_assoc' whisker_left [of s\<^sub>1 "p\<^sub>0 \ \'" "p\<^sub>0 \ w"] whisker_left [of p\<^sub>0] by (metis \.ide_leg1 assoc'_simps(1) hseqE ideD(1) in_homE r\<^sub>0s\<^sub>1.ide_leg0 r\<^sub>0s\<^sub>1.p\<^sub>0_simps w w\<^sub>s) qed finally show ?thesis using \\<^sub>s_def by simp qed ultimately have P\<^sub>s': "?P\<^sub>s ?\\<^sub>s'" by simp have eq\<^sub>s: "\\<^sub>s = ?\\<^sub>s'" using 2 \\<^sub>s P\<^sub>s' by blast have "?P\<^sub>t \'" using \' comp_cod_arr \\p\<^sub>0 \ \' : p\<^sub>0 \ w \ p\<^sub>0 \ w'\\ eq\<^sub>r eq\<^sub>s by auto thus "\' = \" using 3 \ by blast qed qed ultimately show ?thesis by blast qed qed qed end sublocale composite_tabulation_in_maps \ tabulation V H \ \ src trg \r \ s\ tab \s\<^sub>0 \ p\<^sub>0\ \r\<^sub>1 \ p\<^sub>1\ using composite_is_tabulation by simp sublocale composite_tabulation_in_maps \ tabulation_in_maps V H \ \ src trg \r \ s\ tab \s\<^sub>0 \ p\<^sub>0\ \r\<^sub>1 \ p\<^sub>1\ using T0.is_map \.leg1_is_map \.T0.antipar(2) composable \.leg1_is_map \.T0.antipar apply unfold_locales apply simp apply (intro left_adjoints_compose) by auto subsection "The Classifying Category of Maps" text \ \sloppypar We intend to show that if \B\ is a bicategory of spans, then \B\ is biequivalent to \Span(Maps(B))\, for a specific category \Maps(B)\ derived from \B\. The category \Maps(B)\ is constructed in this section as the ``classifying category'' of maps of \B\, which has the same objects as \B\ and which has as 1-cells the isomorphism classes of maps of \B\. We show that, if \B\ is a bicategory of spans, then \Maps(B)\ has pullbacks. \ locale maps_category = B: bicategory_of_spans begin no_notation B.in_hhom ("\_ : _ \ _\") no_notation B.in_hom ("\_ : _ \\<^sub>B _\") notation B.in_hhom ("\_ : _ \\<^sub>B _\") notation B.in_hom ("\_ : _ \\<^sub>B _\") notation B.isomorphic (infix "\\<^sub>B" 50) notation B.iso_class ("\_\\<^sub>B") text \ I attempted to modularize the construction here, by refactoring ``classifying category'' out as a separate locale, but it ended up causing extra work because to apply it we first need to obtain the full sub-bicategory of 2-cells between maps, then construct its classifying category, and then we have to re-prove everything about it, to get rid of any mention of the sub-bicategory construction. So the construction is being done here as a ``one-off'' special case construction, with the necessary properties proved just once. \ text \ The ``hom-categories'' of \Maps(C)\ have as arrows the isomorphism classes of maps of \B\. \ abbreviation Hom where "Hom a b \ {F. \f. \f : a \\<^sub>B b\ \ B.is_left_adjoint f \ F = \f\\<^sub>B}" lemma in_HomD: assumes "F \ Hom a b" shows "F \ {}" and "B.is_iso_class F" and "f \ F \ B.ide f" and "f \ F \ \f : a \\<^sub>B b\" and "f \ F \ B.is_left_adjoint f" and "f \ F \ F = \f\\<^sub>B" proof - show "F \ {}" using assms B.ide_in_iso_class B.left_adjoint_is_ide B.iso_class_is_nonempty by auto show 1: "B.is_iso_class F" using assms B.is_iso_classI B.left_adjoint_is_ide by fastforce show "f \ F \ B.ide f" using assms 1 B.iso_class_memb_is_ide by blast obtain f' where f': "\f' : a \\<^sub>B b\ \ B.is_left_adjoint f' \ F = \f'\\<^sub>B" using assms by auto show "f \ F \ \f : a \\<^sub>B b\" using assms f' B.iso_class_def B.isomorphic_implies_hpar by auto show "f \ F \ B.is_left_adjoint f" using assms f' B.iso_class_def B.left_adjoint_preserved_by_iso [of f'] by auto show "f \ F \ F = \f\\<^sub>B" by (metis B.adjoint_pair_antipar(1) f' B.ide_in_iso_class B.is_iso_classI B.iso_class_elems_isomorphic B.iso_class_eqI) qed definition Comp where "Comp G F \ {h. B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ g \ G \ g \ f \\<^sub>B h)}" lemma in_CompI [intro]: assumes "B.is_iso_class F" and "B.is_iso_class G" and "f \ F" and "g \ G" and "g \ f \\<^sub>B h" shows "h \ Comp G F" unfolding Comp_def using assms by auto lemma in_CompE [elim]: assumes "h \ Comp G F" and "\f g. \ B.is_iso_class F; B.is_iso_class G; f \ F; g \ G; g \ f \\<^sub>B h \ \ T" shows T using assms Comp_def by auto lemma is_iso_class_Comp: assumes "Comp G F \ {}" shows "B.is_iso_class (Comp G F)" proof - obtain h where h: "h \ Comp G F" using assms by auto have ide_h: "B.ide h" using h Comp_def B.isomorphic_implies_hpar(2) by auto obtain f g where fg: "B.is_iso_class F \ B.is_iso_class G \ f \ F \ g \ G \ g \ f \\<^sub>B h" using h Comp_def by auto have "Comp G F = \g \ f\\<^sub>B \ B.ide (g \ f)" proof (intro conjI) show "B.ide (g \ f)" using fg B.iso_class_memb_is_ide B.isomorphic_implies_ide(1) by auto show "Comp G F = \g \ f\\<^sub>B" proof show "\g \ f\\<^sub>B \ Comp G F" unfolding Comp_def B.iso_class_def using fg by auto show "Comp G F \ \g \ f\\<^sub>B" proof fix h' assume h': "h' \ Comp G F" obtain f' g' where f'g': "f' \ F \ g' \ G \ g' \ f' \\<^sub>B h'" using h' Comp_def by auto have 1: "f' \\<^sub>B f \ g' \\<^sub>B g" using f'g' fg B.iso_class_elems_isomorphic by auto moreover have "B.ide f \ B.ide f' \ B.ide g \ B.ide g'" using 1 B.isomorphic_implies_hpar by auto ultimately have "g' \ f' \\<^sub>B g \ f" using f'g' fg B.hcomp_isomorphic_ide B.hcomp_ide_isomorphic B.isomorphic_transitive B.isomorphic_implies_hpar by (meson B.hseqE B.ideD(1)) hence "h' \\<^sub>B g \ f" using f'g' B.isomorphic_symmetric B.isomorphic_transitive by blast thus "h' \ B.iso_class (g \ f)" using B.iso_class_def B.isomorphic_symmetric by simp qed qed qed thus ?thesis using assms B.is_iso_class_def B.ide_in_iso_class by auto qed lemma Comp_is_extensional: assumes "Comp G F \ {}" shows "B.is_iso_class F" and "B.is_iso_class G" and "F \ {}" and "G \ {}" using assms Comp_def by auto lemma Comp_eqI [intro]: assumes "h \ Comp G F" and "h' \ Comp G' F'" and "h \\<^sub>B h'" shows "Comp G F = Comp G' F'" proof - obtain f g where fg: "f \ F \ g \ G \ g \ f \\<^sub>B h" using assms comp_def by auto obtain f' g' where f'g': "f' \ F' \ g' \ G' \ g' \ f' \\<^sub>B h'" using assms by auto have "h \ Comp G F \ Comp G' F'" by (meson IntI assms in_CompE in_CompI B.isomorphic_symmetric B.isomorphic_transitive) hence "Comp G F \ Comp G' F' \ {}" by auto thus ?thesis using assms is_iso_class_Comp by (metis empty_iff B.iso_class_eq) qed lemma Comp_eq_iso_class_memb: assumes "h \ Comp G F" shows "Comp G F = \h\\<^sub>B" proof show "Comp G F \ \h\\<^sub>B" proof fix h' assume h': "h' \ Comp G F" obtain f g where fg: "f \ F \ g \ G \ g \ f \\<^sub>B h" using assms by auto obtain f' g' where f'g': "f' \ F \ g' \ G \ g' \ f' \\<^sub>B h'" using h' by auto have "f \\<^sub>B f' \ g \\<^sub>B g'" using assms fg f'g' in_HomD(6) B.iso_class_elems_isomorphic by auto moreover have "B.ide f \ B.ide f' \ B.ide g \ B.ide g'" using assms fg f'g' in_HomD [of F] in_HomD [of G] by (meson calculation B.isomorphic_implies_ide(1) B.isomorphic_implies_ide(2)) moreover have "src g = trg f \ src g = trg f' \ src g' = trg f \ src g' = trg f'" using fg f'g' by (metis B.seq_if_composable calculation(1) B.ideD(1) B.isomorphic_implies_hpar(1) B.isomorphic_implies_hpar(3) B.not_arr_null) ultimately have "g \ f \\<^sub>B g' \ f'" using fg f'g' B.hcomp_ide_isomorphic B.hcomp_isomorphic_ide B.isomorphic_transitive by metis thus "h' \ \h\\<^sub>B" using fg f'g' B.isomorphic_symmetric B.isomorphic_transitive B.iso_class_def [of h] by blast qed show "\h\\<^sub>B \ Comp G F" proof (unfold B.iso_class_def Comp_def) obtain f g where 1: "f \ F \ g \ G \ g \ f \\<^sub>B h" using assms in_HomD Comp_def by (meson in_CompE) show "{h'. B.isomorphic h h'} \ {h. B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ g \ G \ g \ f \\<^sub>B h)}" using assms 1 B.isomorphic_transitive by blast qed qed interpretation concrete_category \Collect B.obj\ Hom B.iso_class \\_ _ _. Comp\ proof show "\a. a \ Collect B.obj \ \a\\<^sub>B \ Hom a a" by (metis (mono_tags, lifting) B.ide_in_hom(1) mem_Collect_eq B.objE B.obj_is_self_adjoint(1)) show "\a b c F G. \ a \ Collect B.obj; b \ Collect B.obj; c \ Collect B.obj; F \ Hom a b; G \ Hom b c \ \ Comp G F \ Hom a c" proof - fix a b c F G assume a: "a \ Collect B.obj" and b: "b \ Collect B.obj" and c: "c \ Collect B.obj" and F: "F \ Hom a b" and G: "G \ Hom b c" obtain f where f: "\f : a \\<^sub>B b\ \ B.is_left_adjoint f \ F = \f\\<^sub>B" using F by blast obtain g where g: "\g : b \\<^sub>B c\ \ B.is_left_adjoint g \ G = \g\\<^sub>B" using G by blast have "{h. B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ \f : a \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>B h)} = \g \ f\\<^sub>B" proof show "{h. B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ \f : a \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>B h)} \ \g \ f\\<^sub>B" proof fix h assume "h \ {h. B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ \f : a \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>B h)}" hence h: "B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ \f : a \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>B h)" by simp show "h \ \g \ f\\<^sub>B" proof - obtain f' g' where f'g': "g' \ G \ f' \ F \ g' \ f' \\<^sub>B h" using h by auto obtain \ where \: "\\ : f \\<^sub>B f'\ \ B.iso \" using f f'g' F B.iso_class_def by auto obtain \ where \: "\\ : g \\<^sub>B g'\ \ B.iso \" using g f'g' G B.iso_class_def by auto have 1: "\\ \ \ : g \ f \\<^sub>B g' \ f'\" using f g \ \ by auto moreover have "B.iso (\ \ \)" using f g \ \ 1 B.iso_hcomp by auto ultimately show ?thesis using f'g' B.iso_class_def B.isomorphic_def by auto qed qed show "\g \ f\\<^sub>B \ {h. B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ \f : a \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>B h)}" using f g B.iso_class_def B.isomorphic_reflexive B.left_adjoint_is_ide B.is_iso_classI by blast qed hence 1: "\gf. gf \ B.iso_class (g \ f) \ B.is_iso_class F \ B.is_iso_class G \ (\f g. f \ F \ \f : a \\<^sub>B b\ \ g \ G \ \g : b \\<^sub>B c\ \ g \ f \\<^sub>B gf)" by blast show "Comp G F \ Hom a c" proof - have gf: "B.is_left_adjoint (g \ f)" by (meson f g B.hseqE B.hseqI B.left_adjoints_compose) obtain gf' where gf': "B.adjoint_pair (g \ f) gf'" using gf by blast hence "\g \ f\\<^sub>B = Comp G F" using 1 Comp_eq_iso_class_memb B.ide_in_iso_class B.left_adjoint_is_ide by blast thus ?thesis using f g gf' by blast qed qed show "\a b F. \ a \ Collect B.obj; F \ Hom a b \ \ Comp F \a\\<^sub>B = F" proof - fix a b F assume a: "a \ Collect B.obj" assume F: "F \ Hom a b" obtain f where f: "\f : a \\<^sub>B b\ \ B.is_left_adjoint f \ F = \f\\<^sub>B" using F by auto have *: "\f'. f' \ F \ \f' : a \\<^sub>B b\ \ B.ide f' \ f \\<^sub>B f'" using f B.iso_class_def by force show "Comp F \a\\<^sub>B = F" proof show "Comp F \a\\<^sub>B \ F" proof fix h assume "h \ Comp F \a\\<^sub>B" hence h: "\f' a'. f' \ F \ a' \ \a\\<^sub>B \ f' \ a' \\<^sub>B h" unfolding Comp_def by auto obtain f' a' where f'a': "f' \ F \ a' \ \a\\<^sub>B \ f' \ a' \\<^sub>B h" using h by auto have "B.isomorphic f h" proof - have "B.isomorphic f (f \ a)" using f B.iso_runit' [of f] B.isomorphic_def B.left_adjoint_is_ide by blast also have "f \ a \\<^sub>B f' \ a" using f f'a' B.iso_class_def B.hcomp_isomorphic_ide apply (elim conjE B.in_hhomE) by auto also have "f' \ a \\<^sub>B f' \ a'" using f'a' * [of f'] B.iso_class_def B.hcomp_ide_isomorphic by auto also have "f' \ a' \\<^sub>B h" using f'a' by simp finally show ?thesis by blast qed thus "h \ F" using f B.iso_class_def by simp qed show "F \ Comp F \a\\<^sub>B" proof fix h assume h: "h \ F" have "f \ F" using f B.iso_class_def B.right_adjoint_determines_left_up_to_iso by auto moreover have "a \ B.iso_class a" using a B.iso_class_def B.isomorphic_reflexive by auto moreover have "f \ a \\<^sub>B h" proof - have "f \ a \\<^sub>B f" using f B.iso_runit [of f] B.isomorphic_def B.left_adjoint_is_ide by blast also have "f \\<^sub>B h" using h * by simp finally show ?thesis by blast qed ultimately show "h \ Comp F \a\\<^sub>B" unfolding Comp_def using a f F B.is_iso_classI B.left_adjoint_is_ide by auto qed qed qed show "\a b F. \ b \ Collect B.obj; F \ Hom a b \ \ Comp \b\\<^sub>B F = F" proof - fix a b F assume b: "b \ Collect B.obj" assume F: "F \ Hom a b" obtain f where f: "\f : a \\<^sub>B b\ \ B.is_left_adjoint f \ F = \f\\<^sub>B" using F by auto have *: "\f'. f' \ F \ \f' : a \\<^sub>B b\ \ B.ide f' \ f \\<^sub>B f'" using f B.iso_class_def by force show "Comp \b\\<^sub>B F = F" proof show "Comp \b\\<^sub>B F \ F" proof fix h assume "h \ Comp \b\\<^sub>B F" hence h: "\b' f'. b' \ \b\\<^sub>B \ f' \ F \ b' \ f' \\<^sub>B h" unfolding Comp_def by auto obtain b' f' where b'f': "b' \ \b\\<^sub>B \ f' \ F \ b' \ f' \\<^sub>B h" using h by auto have "f \\<^sub>B h" proof - have "f \\<^sub>B b \ f" using f B.iso_lunit' [of f] B.isomorphic_def B.left_adjoint_is_ide by blast also have "... \\<^sub>B b \ f'" using f b'f' B.iso_class_def B.hcomp_ide_isomorphic by (elim conjE B.in_hhomE, auto) also have "... \\<^sub>B b' \ f'" using b'f' * [of f'] B.iso_class_def B.hcomp_isomorphic_ide by auto also have "... \\<^sub>B h" using b'f' by simp finally show ?thesis by blast qed thus "h \ F" using f B.iso_class_def by simp qed show "F \ Comp \b\\<^sub>B F" proof fix h assume h: "h \ F" have "f \ F" using f B.iso_class_def B.right_adjoint_determines_left_up_to_iso by auto moreover have "b \ B.iso_class b" using b B.iso_class_def B.isomorphic_reflexive by auto moreover have "b \ f \\<^sub>B h" proof - have "b \ f \\<^sub>B f" using f B.iso_lunit [of f] B.isomorphic_def B.left_adjoint_is_ide by blast also have "f \\<^sub>B h" using h * by simp finally show ?thesis by blast qed ultimately show "h \ Comp \b\\<^sub>B F" unfolding Comp_def using b f F B.is_iso_classI B.left_adjoint_is_ide by auto qed qed qed show "\a b c d F G H. \ a \ Collect B.obj; b \ Collect B.obj; c \ Collect B.obj; d \ Collect B.obj; F \ Hom a b; G \ Hom b c; H \ Hom c d \ \ Comp H (Comp G F) = Comp (Comp H G) F" proof - fix a b c d F G H assume F: "F \ Hom a b" and G: "G \ Hom b c" and H: "H \ Hom c d" show "Comp H (Comp G F) = Comp (Comp H G) F" proof show "Comp H (Comp G F) \ Comp (Comp H G) F" proof fix x assume x: "x \ Comp H (Comp G F)" obtain f g h gf where 1: "B.is_iso_class F \ B.is_iso_class G \ B.is_iso_class H \ h \ H \ g \ G \ f \ F \ gf \ Comp G F \ g \ f \\<^sub>B gf \ h \ gf \\<^sub>B x" using x unfolding Comp_def by blast have hgf: "B.ide f \ B.ide g \ B.ide h" using 1 F G H B.isomorphic_implies_hpar in_HomD B.left_adjoint_is_ide by (metis (mono_tags, lifting)) have "h \ g \ f \\<^sub>B x" proof - have "h \ g \ f \\<^sub>B h \ gf" using 1 hgf B.hcomp_ide_isomorphic by (metis (full_types) B.isomorphic_implies_hpar(1) B.isomorphic_reflexive B.isomorphic_symmetric B.seq_if_composable) also have "h \ gf \\<^sub>B x" using 1 by simp finally show ?thesis by blast qed moreover have "(h \ g) \ f \\<^sub>B h \ g \ f" using 1 hgf B.iso_assoc B.assoc_in_hom B.isomorphic_def by (metis B.hseq_char B.ideD(1-3) B.isomorphic_implies_hpar(1) B.trg_hcomp calculation) moreover have hg: "\h \ g : b \\<^sub>B d\" using G H 1 in_HomD(4) by blast moreover have "h \ g \ Comp H G" unfolding Comp_def using 1 hgf F G H in_HomD [of F a b] in_HomD [of G b c] in_HomD [of H c d] B.isomorphic_reflexive B.hcomp_ide_isomorphic B.hseqI' by (metis (no_types, lifting) B.hseqE B.hseqI mem_Collect_eq) ultimately show "x \ Comp (Comp H G) F" using 1 F G H hgf B.is_iso_class_def is_iso_class_Comp [of H G] B.isomorphic_reflexive [of "h \ g"] apply (intro in_CompI) apply auto[7] apply blast apply simp by (meson B.isomorphic_symmetric B.isomorphic_transitive) qed show "Comp (Comp H G) F \ Comp H (Comp G F)" proof fix x assume x: "x \ Comp (Comp H G) F" obtain f g h hg where 1: "B.is_iso_class F \ B.is_iso_class G \ B.is_iso_class H \ h \ H \ g \ G \ f \ F \ hg \ Comp H G \ h \ g \\<^sub>B hg \ hg \ f \\<^sub>B x" using x unfolding Comp_def by blast have hgf: "B.ide f \ B.ide g \ B.ide h \ src h = trg g \ src g = trg f" using 1 F G H in_HomD B.left_adjoint_is_ide by (metis (no_types, lifting) B.hseq_char' B.ideD(1) B.ide_dom B.in_homE B.isomorphic_def B.isomorphic_symmetric B.seqI' B.seq_if_composable B.src_dom B.src_hcomp B.vseq_implies_hpar(1)) have 2: "(h \ g) \ f \\<^sub>B x" proof - have "(h \ g) \ f \\<^sub>B hg \ f" using 1 F G H hgf by (simp add: B.hcomp_isomorphic_ide) also have "hg \ f \\<^sub>B x" using 1 by simp finally show ?thesis by blast qed moreover have "(h \ g) \ f \\<^sub>B h \ g \ f" using hgf B.iso_assoc B.assoc_in_hom B.isomorphic_def by auto moreover have "g \ f \ Comp G F" using 1 F G hgf B.isomorphic_reflexive by blast ultimately show "x \ Comp H (Comp G F)" using 1 hgf is_iso_class_Comp [of G F] B.isomorphic_reflexive [of "g \ f"] apply (intro in_CompI) apply auto[6] apply simp apply auto[1] by (meson B.isomorphic_symmetric B.isomorphic_transitive) qed qed qed qed lemma is_concrete_category: shows "concrete_category (Collect B.obj) Hom B.iso_class (\_ _ _. Comp)" .. sublocale concrete_category \Collect B.obj\ Hom B.iso_class \\_ _ _. Comp\ using is_concrete_category by simp abbreviation comp (infixr "\" 55) where "comp \ COMP" notation in_hom ("\_ : _ \ _\") no_notation B.in_hom ("\_ : _ \\<^sub>B _\") lemma Map_memb_in_hhom: assumes "\F : A \ B\" and "f \ Map F" shows "\f : Dom A \\<^sub>B Dom B\" proof - have "\f : Dom F \\<^sub>B Cod F\" using assms arr_char [of F] in_HomD [of "Map F" "Dom F" "Cod F"] by blast moreover have "Dom F = Dom A" using assms by auto moreover have "Cod F = Dom B" using assms by auto ultimately show ?thesis by simp qed lemma MkArr_in_hom': assumes "B.is_left_adjoint f" and "\f : a \\<^sub>B b\" shows "\MkArr a b \f\\<^sub>B : MkIde a \ MkIde b\" using assms MkArr_in_hom by blast text \ The isomorphisms in \Maps(B)\ are the isomorphism classes of equivalence maps in \B\. \ lemma iso_char: shows "iso F \ arr F \ (\f. f \ Map F \ B.equivalence_map f)" proof assume F: "iso F" have "\f. f \ Map F \ B.equivalence_map f" proof - fix f assume f: "f \ Map F" obtain G where G: "inverse_arrows F G" using F by auto obtain g where g: "g \ Map G" using G arr_char [of G] in_HomD(1) by blast have f: "f \ Map F \ \f : Dom F \\<^sub>B Cod F\ \ B.ide f \ B.is_left_adjoint f" by (metis (mono_tags, lifting) F iso_is_arr is_concrete_category concrete_category.arr_char f in_HomD(2,4-5) B.iso_class_memb_is_ide) have g: "g \ Map G \ \g : Cod F \\<^sub>B Dom F\ \ B.ide g \ B.is_left_adjoint g" by (metis (no_types, lifting) F G Cod_cod Cod_dom arr_inv cod_inv dom_inv inverse_unique iso_is_arr is_concrete_category concrete_category.Map_in_Hom g in_HomD(2,4-5) B.iso_class_memb_is_ide) have "src (g \ f) \\<^sub>B g \ f" proof - have "g \ f \ Map (G \ F)" proof - have 1: "f \ Map F \ g \ Map G \ g \ f \\<^sub>B g \ f" using F G f g B.isomorphic_reflexive by force have 2: "Dom G = Cod F \ Cod G = Dom F" using F G arr_char by (metis (no_types, lifting) Cod.simps(1) Cod_dom arr_inv cod_char comp_inv_arr dom_inv inverse_unique iso_is_arr is_concrete_category concrete_category.Cod_comp) have "g \ f \ Comp (Map G) (Map F)" using 1 F iso_is_arr Map_in_Hom [of F] in_HomD(2) apply (intro in_CompI, auto) proof - show "B.is_iso_class (Map G)" using G iso_is_arr Map_in_Hom [of G] in_HomD(2) [of "Map G"] by blast qed thus ?thesis using F G f g comp_char [of G F] by auto qed moreover have "Dom F \ Map (G \ F)" by (metis (no_types, lifting) F G Map_dom comp_inv_arr iso_is_arr g B.ide_in_iso_class B.in_hhomE B.objE) moreover have "Map (G \ F) = \Dom F\\<^sub>B" by (simp add: F G comp_inv_arr iso_is_arr) moreover have "Dom F = src (g \ f)" using f g by fastforce ultimately show ?thesis using f g B.iso_class_elems_isomorphic B.is_iso_classI by (metis B.hseqI B.ide_src) qed moreover have "f \ g \\<^sub>B trg f" proof - have "f \ g \ Map (F \ G)" proof - have 1: "f \ Map F \ g \ Map G \ f \ g \\<^sub>B f \ g" using F G f g B.isomorphic_reflexive by force have 2: "Dom G = Cod F \ Cod G = Dom F" using F G arr_char by (metis (no_types, lifting) Cod.simps(1) Cod_dom arr_inv cod_char comp_inv_arr dom_inv inverse_unique iso_is_arr is_concrete_category concrete_category.Cod_comp) hence "f \ g \ Comp (Map F) (Map G)" using 1 F iso_is_arr Map_in_Hom [of F] in_HomD(2) apply (intro in_CompI, auto) proof - show "B.is_iso_class (Map G)" using G iso_is_arr Map_in_Hom [of G] in_HomD(2) [of "Map G"] by blast qed thus ?thesis using F G f g comp_char [of F G] by auto qed moreover have "Cod F \ Map (F \ G)" by (metis (no_types, lifting) F G Map_cod comp_arr_inv dom_inv inverse_unique iso_is_arr g B.ide_in_iso_class B.in_hhomE B.src.preserves_ide) moreover have "Map (F \ G) = \Cod F\\<^sub>B" by (metis (no_types, lifting) F G Map_cod comp_arr_inv dom_inv inverse_unique iso_is_arr) moreover have "Cod F = trg (f \ g)" using f g by fastforce ultimately show ?thesis using B.iso_class_elems_isomorphic by (metis f g B.is_iso_classI B.in_hhomE B.src.preserves_ide) qed ultimately show "B.equivalence_map f" using f g B.equivalence_mapI B.quasi_inversesI [of f g] by fastforce qed thus "arr F \ (\f. f \ Map F \ B.equivalence_map f)" using F by blast next assume F: "arr F \ (\f. f \ Map F \ B.equivalence_map f)" show "iso F" proof - obtain f where f: "f \ Map F" using F arr_char in_HomD(1) by blast have f_in_hhom: "\f : Dom F \\<^sub>B Cod F\" using f F arr_char in_HomD(4) [of "Map F" "Dom F" "Cod F" f] by simp have "Map F = B.iso_class f" using f F arr_char in_HomD(6) [of "Map F" "Dom F" "Cod F" f] by simp obtain g \ \' where \': "equivalence_in_bicategory V H \ \ src trg f g \ \'" using f F B.equivalence_map_def by auto interpret \': equivalence_in_bicategory V H \ \ src trg f g \ \' using \' by auto obtain \ where \: "adjoint_equivalence_in_bicategory V H \ \ src trg f g \ \" using f F \'.ide_right \'.antipar \'.unit_in_hom \'.unit_is_iso B.equivalence_map_def B.equivalence_refines_to_adjoint_equivalence [of f g \] by auto interpret \: adjoint_equivalence_in_bicategory V H \ \ src trg f g \ \ using \ by simp have g_in_hhom: "\g : Cod F \\<^sub>B Dom F\" using \.ide_right \.antipar f_in_hhom by auto have fg: "B.quasi_inverses f g" using B.quasi_inverses_def \.equivalence_in_bicategory_axioms by auto have g: "\g : Cod F \\<^sub>B Dom F\ \ B.is_left_adjoint g \ \g\\<^sub>B = \g\\<^sub>B" using \'.dual_equivalence B.equivalence_is_left_adjoint B.equivalence_map_def g_in_hhom by blast have F_as_MkArr: "F = MkArr (Dom F) (Cod F) \f\\<^sub>B" using F MkArr_Map \Map F = B.iso_class f\ by fastforce have F_in_hom: "in_hom F (MkIde (Dom F)) (MkIde (Cod F))" using F arr_char dom_char cod_char by (intro in_homI, auto) let ?G = "MkArr (Cod F) (Dom F) \g\\<^sub>B" have "arr ?G" using MkArr_in_hom' g by blast hence G_in_hom: "\?G : MkIde (Cod F) \ MkIde (Dom F)\" using arr_char MkArr_in_hom by simp have "inverse_arrows F ?G" proof show "ide (?G \ F)" proof - have "?G \ F = dom F" proof (intro arr_eqI) show 1: "seq ?G F" using F_in_hom G_in_hom by blast show "arr (dom F)" using F_in_hom ide_dom by fastforce show "Dom (?G \ F) = Dom (dom F)" using F_in_hom G_in_hom 1 comp_char by auto show "Cod (?G \ F) = Cod (dom F)" using F_in_hom G_in_hom 1 comp_char by auto show "Map (?G \ F) = Map (dom F)" proof - have "Map (?G \ F) = Comp \g\\<^sub>B \f\\<^sub>B" using 1 comp_char [of ?G F] \Map F = B.iso_class f\ by simp also have "... = \g \ f\\<^sub>B" proof - have "g \ f \ Comp \g\\<^sub>B \f\\<^sub>B" by (metis \.ide_left \.ide_right \.unit_in_vhom \.unit_simps(3) B.arrI B.ide_cod B.ide_in_iso_class in_CompI B.is_iso_classI B.isomorphic_reflexive) thus ?thesis using Comp_eq_iso_class_memb F_in_hom G_in_hom arr_char arr_char \Map F = B.iso_class f\ by auto qed also have "... = \src f\\<^sub>B" using \.unit_in_hom \.unit_is_iso B.isomorphic_def B.iso_class_def B.isomorphic_symmetric by (intro B.iso_class_eqI, blast) also have "... = \Dom F\\<^sub>B" using \.ide_left f_in_hhom B.ide_in_iso_class B.in_hhomE B.src.preserves_ide B.isomorphic_reflexive by (intro B.iso_class_eqI, fastforce) also have "... = Map (dom F)" using F_in_hom dom_char by auto finally show ?thesis by blast qed qed thus ?thesis using F by simp qed show "ide (F \ ?G)" proof - have "F \ ?G = cod F" proof (intro arr_eqI) show 1: "seq F ?G" using F_in_hom G_in_hom by blast show "arr (cod F)" using F_in_hom ide_cod by fastforce show "Dom (F \ ?G) = Dom (cod F)" using F_in_hom G_in_hom 1 comp_char by auto show "Cod (F \ ?G) = Cod (cod F)" using F_in_hom G_in_hom 1 comp_char by auto show "Map (F \ ?G) = Map (cod F)" proof - have "Map (F \ ?G) = Comp \f\\<^sub>B \g\\<^sub>B" using 1 comp_char [of F ?G] \Map F = \f\\<^sub>B\ by simp also have "... = \f \ g\\<^sub>B" proof - have "f \ g \ Comp \f\\<^sub>B \g\\<^sub>B" by (metis \.antipar(1) \.ide_left \.ide_right B.ide_hcomp B.ide_in_iso_class in_CompI B.is_iso_classI B.isomorphic_reflexive) thus ?thesis using Comp_eq_iso_class_memb F_in_hom G_in_hom arr_char arr_char \Map F = \f\\<^sub>B\ by auto qed also have "... = \trg f\\<^sub>B" proof - have "trg f \ \f \ g\\<^sub>B" using \.counit_in_hom \.counit_is_iso B.isomorphic_def B.iso_class_def B.isomorphic_symmetric by blast thus ?thesis using B.iso_class_eqI by (metis \.antipar(1) \.ide_left \.ide_right B.ide_hcomp B.ide_in_iso_class B.is_iso_classI B.iso_class_elems_isomorphic) qed also have "... = \Cod F\\<^sub>B" using f_in_hhom by auto also have "... = Map (cod F)" using F_in_hom dom_char by auto finally show ?thesis by blast qed qed thus ?thesis using F by simp qed qed thus ?thesis by auto qed qed lemma iso_char': shows "iso F \ arr F \ (\f. f \ Map F \ B.equivalence_map f)" proof - have "arr F \ (\f. f \ Map F \ B.equivalence_map f) \ (\f. f \ Map F \ B.equivalence_map f)" proof assume F: "arr F" show "(\f. f \ Map F \ B.equivalence_map f) \ (\f. f \ Map F \ B.equivalence_map f)" using F arr_char in_HomD(1) by blast show "(\f. f \ Map F \ B.equivalence_map f) \ (\f. f \ Map F \ B.equivalence_map f)" by (metis (no_types, lifting) F is_concrete_category concrete_category.arr_char B.equivalence_map_preserved_by_iso in_HomD(2) B.iso_class_elems_isomorphic) qed thus ?thesis using iso_char by blast qed text \ The following mapping takes a map in \B\ to the corresponding arrow of \Maps(B)\. \ abbreviation CLS ("\\_\\") where "\\f\\ \ MkArr (src f) (trg f) \f\\<^sub>B" lemma ide_CLS_obj: assumes "B.obj a" shows "ide \\a\\" by (simp add: assms B.obj_simps) lemma CLS_in_hom: assumes "B.is_left_adjoint f" shows "\\\f\\ : \\src f\\ \ \\trg f\\\" using assms B.left_adjoint_is_ide MkArr_in_hom MkArr_in_hom' by simp lemma CLS_eqI: assumes "B.ide f" shows "\\f\\ = \\g\\ \ f \\<^sub>B g" by (metis arr.inject assms B.ide_in_iso_class B.iso_class_def B.iso_class_eqI B.isomorphic_implies_hpar(3) B.isomorphic_implies_hpar(4) B.isomorphic_symmetric mem_Collect_eq) lemma CLS_hcomp: assumes "B.ide f" and "B.ide g" and "src f = trg g" shows "\\f \ g\\ = MkArr (src g) (trg f) (Comp \f\\<^sub>B \g\\<^sub>B)" proof - have "\\f \ g\\ = MkArr (src g) (trg f) \f \ g\\<^sub>B" using assms B.left_adjoint_is_ide by simp moreover have "\f \ g\\<^sub>B = Comp \f\\<^sub>B \g\\<^sub>B" proof show "\f \ g\\<^sub>B \ Comp \f\\<^sub>B \g\\<^sub>B" by (metis assms(1-2) B.ide_in_iso_class in_CompI B.is_iso_classI B.iso_class_def mem_Collect_eq subsetI) show "Comp \f\\<^sub>B \g\\<^sub>B \ \f \ g\\<^sub>B" by (metis Comp_eq_iso_class_memb \\f \ g\\<^sub>B \ Comp \f\\<^sub>B \g\\<^sub>B\ assms(1-3) B.ide_hcomp B.ide_in_iso_class subset_iff) qed ultimately show ?thesis by simp qed lemma comp_CLS: assumes "B.is_left_adjoint f" and "B.is_left_adjoint g" and "f \ g \\<^sub>B h" shows "\\f\\ \ \\g\\ = \\h\\" proof - have hseq_fg: "B.hseq f g" using assms B.isomorphic_implies_hpar(1) by simp have "seq \\f\\ \\g\\" using assms hseq_fg CLS_in_hom [of f] CLS_in_hom [of g] apply (elim B.hseqE) by auto hence "\\f\\ \ \\g\\ = MkArr (src g) (trg f) (Comp \f\\<^sub>B \g\\<^sub>B)" using comp_char [of "CLS f" "CLS g"] by simp also have "... = \\f \ g\\" using assms hseq_fg CLS_hcomp apply (elim B.hseqE) using B.adjoint_pair_antipar(1) by auto also have "... = \\h\\" using assms B.isomorphic_symmetric by (simp add: assms(3) B.iso_class_eqI B.isomorphic_implies_hpar(3) B.isomorphic_implies_hpar(4)) finally show ?thesis by simp qed text \ The following mapping that takes an arrow of \Maps(B)\ and chooses some representative map in \B\. \ definition REP where "REP F \ if arr F then SOME f. f \ Map F else B.null" lemma REP_in_Map: assumes "arr A" shows "REP A \ Map A" proof - have "Map A \ {}" using assms arr_char in_HomD(1) by blast thus ?thesis using assms REP_def someI_ex [of "\f. f \ Map A"] by auto qed lemma REP_in_hhom [intro]: assumes "in_hom F A B" shows "\REP F : src (REP A) \\<^sub>B src (REP B)\" and "B.is_left_adjoint (REP F)" proof - have "Map F \ {}" using assms in_hom_char arr_char null_char in_HomD(1) [of "Map F" "Dom F" "Cod F"] by blast hence 1: "REP F \ Map F" using assms REP_def someI_ex [of "\f. f \ Map F"] by auto hence 2: "B.arr (REP F)" using assms 1 in_hom_char [of F A B] B.iso_class_def Map_memb_in_hhom B.in_hhom_def by blast hence "\REP F : Dom A \\<^sub>B Dom B\" using assms 1 in_hom_char [of F A B] Map_memb_in_hhom by auto thus "\REP F : src (REP A) \\<^sub>B src (REP B)\" using assms by (metis (no_types, lifting) Map_memb_in_hhom ideD(1) in_homI in_hom_char REP_in_Map B.in_hhom_def) have "REP F \ \REP F\\<^sub>B" using assms 1 2 arr_char [of F] in_HomD(6) B.ide_in_iso_class B.iso_class_memb_is_ide in_hom_char by blast thus "B.is_left_adjoint (REP F)" using assms 1 2 arr_char in_HomD(5) [of "Map F" "Dom F" "Cod F" "REP F"] by auto qed lemma ide_REP: assumes "arr A" shows "B.ide (REP A)" using assms REP_in_hhom(2) B.adjoint_pair_antipar(1) by blast lemma REP_simps [simp]: assumes "arr A" shows "B.ide (REP A)" and "src (REP A) = Dom A" and "trg (REP A) = Cod A" and "B.dom (REP A) = REP A" and "B.cod (REP A) = REP A" proof - show "B.ide (REP A)" using assms ide_REP by blast show "src (REP A) = Dom A" using assms REP_in_hhom by (metis (no_types, lifting) Map_memb_in_hhom Dom_dom in_homI REP_in_Map B.in_hhom_def) show "trg (REP A) = Cod A" using assms REP_in_hhom by (metis (no_types, lifting) Map_memb_in_hhom Dom_cod in_homI REP_in_Map B.in_hhom_def) show "B.dom (REP A) = REP A" using assms in_homI REP_in_hhom(2) B.adjoint_pair_antipar(1) B.ideD(2) by blast show "B.cod (REP A) = REP A" using assms in_homI REP_in_hhom(2) B.adjoint_pair_antipar(1) B.ideD(3) by blast qed lemma isomorphic_REP_src: assumes "ide A" shows "REP A \\<^sub>B src (REP A)" using assms by (metis (no_types, lifting) ideD(1) ide_char REP_in_Map ide_REP REP_simps(2) B.is_iso_classI B.ide_in_iso_class B.iso_class_elems_isomorphic B.src.preserves_ide) lemma isomorphic_REP_trg: assumes "ide A" shows "REP A \\<^sub>B trg (REP A)" using assms ide_char isomorphic_REP_src by auto lemma CLS_REP: assumes "arr F" shows "\\REP F\\ = F" by (metis (mono_tags, lifting) MkArr_Map is_concrete_category REP_in_Map REP_simps(2) REP_simps(3) assms concrete_category.Map_in_Hom in_HomD(6)) lemma REP_CLS: assumes "B.is_left_adjoint f" shows "REP \\f\\ \\<^sub>B f" by (metis (mono_tags, lifting) CLS_in_hom Map.simps(1) in_homE REP_in_Map assms B.iso_class_def B.isomorphic_symmetric mem_Collect_eq) lemma isomorphic_hcomp_REP: assumes "seq F G" shows "REP F \ REP G \\<^sub>B REP (F \ G)" proof - have 1: "Dom F = Cod G" using assms seq_char by simp have 2: "Map F = B.iso_class (REP F)" using assms seq_char arr_char REP_in_Map in_HomD(6) by meson have 3: "Map G = B.iso_class (REP G)" using assms seq_char arr_char REP_in_Map in_HomD(6) [of "Map G" "Dom G" "Cod G" "REP G"] by auto have "Map (F \ G) = Comp \REP F\\<^sub>B \REP G\\<^sub>B" using assms seq_char null_char by (metis (no_types, lifting) CLS_REP Map.simps(1) Map_comp) have "Comp \REP F\\<^sub>B \REP G\\<^sub>B = \REP F \ REP G\\<^sub>B" proof - have "REP F \ REP G \ Comp \REP F\\<^sub>B \REP G\\<^sub>B" proof - have "REP F \ Map F \ REP G \ Map G" using assms seq_char REP_in_Map by auto moreover have "REP F \ REP G \\<^sub>B REP F \ REP G" using assms seq_char 2 B.isomorphic_reflexive by auto ultimately show ?thesis using assms 1 2 3 B.iso_class_def B.is_iso_class_def by (intro in_CompI, auto) qed moreover have "\REP F\\<^sub>B \ Hom (Cod G) (Cod F)" using assms 1 2 arr_char [of F] by auto moreover have "\REP G\\<^sub>B \ Hom (Dom G) (Cod G)" using assms 1 3 arr_char [of G] by auto ultimately show ?thesis using Comp_eq_iso_class_memb assms arr_char arr_char REP_in_Map by (simp add: Comp_def) qed moreover have "REP (F \ G) \ Comp \REP F\\<^sub>B \REP G\\<^sub>B" proof - have "Map (F \ G) = Comp (Map F) (Map G)" using assms 1 comp_char [of F G] by simp thus ?thesis using assms 1 2 3 REP_in_Map [of "F \ G"] by simp qed ultimately show ?thesis using B.iso_class_def by simp qed text \ We now show that \Maps(B)\ has pullbacks. For this we need to exhibit functions \PRJ\<^sub>0\ and \PRJ\<^sub>1\ that produce the legs of the pullback of a cospan \(H, K)\ and verify that they have the required universal property. These are obtained by choosing representatives \h\ and \k\ of \H\ and \K\, respectively, and then taking \PRJ\<^sub>0 = CLS (tab\<^sub>0 (k\<^sup>* \ h))\ and \PRJ\<^sub>1 = CLS (tab\<^sub>1 (k\<^sup>* \ h))\. That these constitute a pullback in \Maps(B)\ follows from the fact that \tab\<^sub>0 (k\<^sup>* \ h)\ and \tab\<^sub>1 (k\<^sup>* \ h)\ form a pseudo-pullback of \(h, k)\ in the underlying bicategory. For our purposes here, it is not sufficient simply to show that \Maps(B)\ has pullbacks and then to treat it as an abstract ``category with pullbacks'' where the pullbacks are chosen arbitrarily. Instead, we have to retain the connection between a pullback in Maps and the tabulation of \k\<^sup>* \ h\ that is ultimately used to obtain it. If we don't do this, then it becomes problematic to define the compositor of a pseudofunctor from the underlying bicategory \B\ to \Span(Maps(B))\, because the components will go from the apex of a composite span (obtained by pullback) to the apex of a tabulation (chosen arbitrarily using \BS2\) and these need not be in agreement with each other. \ definition PRJ\<^sub>0 where "PRJ\<^sub>0 \ \K H. if cospan K H then \\B.tab\<^sub>0 ((REP K)\<^sup>* \ (REP H))\\ else null" definition PRJ\<^sub>1 where "PRJ\<^sub>1 \ \K H. if cospan K H then \\B.tab\<^sub>1 ((REP K)\<^sup>* \ (REP H))\\ else null" interpretation elementary_category_with_pullbacks comp PRJ\<^sub>0 PRJ\<^sub>1 proof show "\H K. \ cospan K H \ PRJ\<^sub>0 K H = null" unfolding PRJ\<^sub>0_def by auto show "\H K. \ cospan K H \ PRJ\<^sub>1 K H = null" unfolding PRJ\<^sub>1_def by auto show "\H K. cospan K H \ commutative_square K H (PRJ\<^sub>1 K H) (PRJ\<^sub>0 K H)" proof - fix H K assume HK: "cospan K H" define h where "h = REP H" define k where "k = REP K" have h: "h \ Map H" using h_def HK REP_in_Map by blast have k: "k \ Map K" using k_def HK REP_in_Map by blast have 1: "B.is_left_adjoint h \ B.is_left_adjoint k \ B.ide h \ B.ide k \ trg h = trg k" using h k h_def k_def HK arr_char cod_char B.in_hhom_def B.left_adjoint_is_ide in_HomD(5) [of "Map H" "Dom H" "Cod H" h] in_HomD(5) [of "Map K" "Dom K" "Cod K" k] apply auto by (metis (no_types, lifting) HK Dom_cod) interpret h: map_in_bicategory \(\)\ \(\)\ \ \ src trg h using 1 by unfold_locales auto interpret k: map_in_bicategory \(\)\ \(\)\ \ \ src trg k using 1 by unfold_locales auto interpret hk: cospan_of_maps_in_bicategory_of_spans \(\)\ \(\)\ \ \ src trg h k using 1 by unfold_locales auto let ?f = "B.tab\<^sub>0 (k\<^sup>* \ h)" let ?g = "B.tab\<^sub>1 (k\<^sup>* \ h)" have span: "span \\?f\\ \\?g\\" using dom_char CLS_in_hom [of ?f] CLS_in_hom [of ?g] by auto have seq: "seq H \\?f\\" using HK seq_char hk.leg0_is_map CLS_in_hom h_def hk.p\<^sub>0_simps hk.satisfies_T0 by fastforce have seq': "seq K \\?g\\" using HK k arr_char dom_char cod_char in_HomD(5) hk.leg1_is_map CLS_in_hom by (metis (no_types, lifting) Cod.simps(1) seq_char REP_simps(2) hk.p\<^sub>1_simps k_def span) show "commutative_square K H (PRJ\<^sub>1 K H) (PRJ\<^sub>0 K H)" proof show "cospan K H" using HK by simp show "dom K = cod (PRJ\<^sub>1 K H)" using seq' PRJ\<^sub>1_def HK h_def k_def by auto show "span (PRJ\<^sub>1 K H) (PRJ\<^sub>0 K H)" unfolding PRJ\<^sub>0_def PRJ\<^sub>1_def using HK span h_def k_def by simp show "K \ PRJ\<^sub>1 K H = H \ PRJ\<^sub>0 K H" proof - have iso: "h \ ?f \\<^sub>B k \ ?g" using hk.\_uniqueness B.isomorphic_symmetric B.isomorphic_def by blast have *: "Comp (Map H) \?f\\<^sub>B = Comp (Map K) \?g\\<^sub>B" proof (intro Comp_eqI) show "h \ ?f \ Comp (Map H) \B.tab\<^sub>0 (k\<^sup>* \ h)\\<^sub>B" proof (unfold Comp_def) have "B.is_iso_class \?f\\<^sub>B" by (simp add: B.is_iso_classI) moreover have "B.is_iso_class (Map H)" using CLS_REP HK Map.simps(1) B.is_iso_classI h.ide_left h_def by (metis (no_types, lifting)) moreover have "?f \ \B.tab\<^sub>0 (k\<^sup>* \ h)\\<^sub>B" by (simp add: B.ide_in_iso_class(1)) moreover have "\?f : src (B.tab\<^sub>0 (k\<^sup>* \ h)) \\<^sub>B Dom H\" using seq seq_char by simp moreover have "h \ Map H" by fact moreover have "\h : Dom H \\<^sub>B Cod H\" by (simp add: HK h_def) moreover have "h \ ?f \\<^sub>B h \ ?f" using B.isomorphic_reflexive by auto ultimately show "h \ B.tab\<^sub>0 (k\<^sup>* \ h) \ {h'. B.is_iso_class \B.tab\<^sub>0 (k\<^sup>* \ h)\\<^sub>B \ B.is_iso_class (Map H) \ (\f g. f \ \B.tab\<^sub>0 (k\<^sup>* \ h)\\<^sub>B \ g \ Map H \ g \ f \\<^sub>B h')}" by auto qed show "k \ ?g \ Comp (Map K) \B.tab\<^sub>1 (k\<^sup>* \ h)\\<^sub>B" proof (unfold Comp_def) have "B.is_iso_class \?g\\<^sub>B" by (simp add: B.is_iso_classI) moreover have "B.is_iso_class (Map K)" by (metis (no_types, lifting) CLS_REP HK Map.simps(1) B.is_iso_classI k.ide_left k_def) moreover have "?g \ \B.tab\<^sub>1 (k\<^sup>* \ h)\\<^sub>B" by (simp add: B.ide_in_iso_class(1)) moreover have "\?g : src (B.tab\<^sub>1 (k\<^sup>* \ h)) \\<^sub>B Dom K\" using seq seq_char B.in_hhom_def seq' by auto moreover have "k \ Map K" by fact moreover have "\k : Dom K \\<^sub>B Cod K\" by (simp add: HK k_def) moreover have "k \ ?g \\<^sub>B k \ ?g" using B.isomorphic_reflexive iso B.isomorphic_implies_hpar(2) by auto ultimately show "k \ B.tab\<^sub>1 (k\<^sup>* \ h) \ {h'. B.is_iso_class \B.tab\<^sub>1 (k\<^sup>* \ h)\\<^sub>B \ B.is_iso_class (Map K) \ (\f g. f \ \B.tab\<^sub>1 (k\<^sup>* \ h)\\<^sub>B \ g \ Map K \ g \ f \\<^sub>B h')}" by auto qed show "h \ ?f \\<^sub>B k \ ?g" using iso by simp qed have "K \ PRJ\<^sub>1 K H = K \ \\?g\\" unfolding PRJ\<^sub>1_def using HK h_def k_def by simp also have "... = MkArr (src ?g) (Cod K) (Comp (Map K) \?g\\<^sub>B)" using seq' comp_char [of K "\\?g\\"] by simp also have "... = MkArr (src ?f) (Cod H) (Comp (Map H) \?f\\<^sub>B)" using * HK cod_char by auto also have "... = comp H \\?f\\" using seq comp_char [of H "\\?f\\"] by simp also have "... = comp H (PRJ\<^sub>0 K H)" unfolding PRJ\<^sub>0_def using HK h_def k_def by simp finally show ?thesis by simp qed qed qed show "\H K U V. commutative_square K H V U \ \!E. comp (PRJ\<^sub>1 K H) E = V \ comp (PRJ\<^sub>0 K H) E = U" proof - fix H K U V assume cs: "commutative_square K H V U" have HK: "cospan K H" using cs by auto (* TODO: Is there any way to avoid this repetition? *) define h where "h = REP H" define k where "k = REP K" have h: "h \ Map H" using h_def HK REP_in_Map by blast have k: "k \ Map K" using k_def HK REP_in_Map by blast have 1: "B.is_left_adjoint h \ B.is_left_adjoint k \ B.ide h \ B.ide k \ trg h = trg k" using h k h_def k_def HK arr_char cod_char B.in_hhom_def B.left_adjoint_is_ide in_HomD(5) [of "Map H" "Dom H" "Cod H" h] in_HomD(5) [of "Map K" "Dom K" "Cod K" k] apply auto by (metis (no_types, lifting) HK Dom_cod) interpret h: map_in_bicategory \(\)\ \(\)\ \ \ src trg h using 1 by unfold_locales auto interpret k: map_in_bicategory \(\)\ \(\)\ \ \ src trg k using 1 by unfold_locales auto interpret hk: cospan_of_maps_in_bicategory_of_spans \(\)\ \(\)\ \ \ src trg h k using 1 by unfold_locales auto let ?f = "B.tab\<^sub>0 (k\<^sup>* \ h)" let ?g = "B.tab\<^sub>1 (k\<^sup>* \ h)" have seq_HU: "seq H U" using cs by auto have seq_KV: "seq K V" using cs by auto let ?u = "REP U" let ?v = "REP V" have u: "B.ide ?u" using ide_REP seq_HU by auto have v: "B.ide ?v" using ide_REP seq_KV by auto have u_is_map: "B.is_left_adjoint ?u" using u seq_HU REP_in_Map arr_char [of U] in_HomD(5) [of "Map U" "Dom U" "Cod U" ?u] by auto have v_is_map: "B.is_left_adjoint ?v" using v seq_KV REP_in_Map arr_char [of V] in_HomD(5) [of "Map V" "Dom V" "Cod V" ?v] by auto have *: "h \ ?u \\<^sub>B k \ ?v" proof - have "h \ ?u \\<^sub>B REP (H \ U)" proof - have "h \ ?u \\<^sub>B REP H \ ?u" proof - have "h \\<^sub>B REP H" using h h_def HK arr_char REP_in_Map B.iso_class_elems_isomorphic in_HomD(5) [of "Map H" "Dom H" "Cod H" h] B.isomorphic_reflexive by auto thus ?thesis using h_def seq_HU B.isomorphic_implies_hpar(1) B.isomorphic_reflexive by (simp add: seq_char) qed also have "... \\<^sub>B REP (H \ U)" using seq_HU isomorphic_hcomp_REP isomorphic_def by blast finally show ?thesis by blast qed also have "... \\<^sub>B REP (K \ V)" using seq_HU cs B.isomorphic_reflexive by auto also have "... \\<^sub>B (k \ ?v)" proof - have "... \\<^sub>B REP K \ ?v" using seq_KV isomorphic_hcomp_REP B.isomorphic_def B.isomorphic_symmetric by blast also have "... \\<^sub>B k \ ?v" proof - have "k \\<^sub>B REP K" using k k_def HK arr_char REP_in_Map B.iso_class_elems_isomorphic in_HomD(5) [of "Map K" "Dom K" "Cod K" k] B.isomorphic_reflexive by auto thus ?thesis using k_def seq_KV B.isomorphic_implies_hpar(1) B.isomorphic_reflexive by (simp add: seq_char) qed finally show ?thesis by blast qed finally show ?thesis by blast qed have hseq_hu: "src h = trg ?u" using * B.isomorphic_implies_hpar by (meson B.hseqE B.ideD(1)) have hseq_kv: "src k = trg ?v" using * B.isomorphic_implies_hpar by (meson B.hseqE B.ideD(1)) obtain w where w: "B.is_left_adjoint w \ ?f \ w \\<^sub>B ?u \ ?v \\<^sub>B (?g \ w)" using * u_is_map v_is_map hk.has_pseudo_pullback [of ?u ?v] B.isomorphic_symmetric by blast have w_in_hom: "\w : src ?u \\<^sub>B src ?f\ \ B.ide w" using w B.left_adjoint_is_ide B.src_cod B.trg_cod B.isomorphic_def by (metis B.hseqE B.ideD(1) B.in_hhom_def B.isomorphic_implies_hpar(3) B.isomorphic_implies_ide(1) B.src_hcomp) let ?W = "CLS w" have W: "\?W : dom U \ dom \\?f\\\" proof show "arr ?W" using w CLS_in_hom by blast thus "dom ?W = dom U" using w_in_hom dom_char REP_in_hhom(1) CLS_in_hom by (metis (no_types, lifting) Dom.simps(1) commutative_squareE dom_char REP_simps(2) cs B.in_hhomE) show "cod ?W = dom \\?f\\" proof - have "src ?f = trg w" by (metis (lifting) B.in_hhomE w_in_hom) thus ?thesis using CLS_in_hom [of ?f] CLS_in_hom [of w] hk.satisfies_T0 w by fastforce qed qed show "\!E. PRJ\<^sub>1 K H \ E = V \ PRJ\<^sub>0 K H \ E = U" proof - have "PRJ\<^sub>1 K H \ ?W = V \ PRJ\<^sub>0 K H \ ?W = U" proof - have "\\?f\\ \ ?W = U" using w w_in_hom u CLS_in_hom comp_CLS B.isomorphic_symmetric CLS_REP hk.leg0_is_map by (metis (mono_tags, lifting) commutative_square_def cs) moreover have "\\?g\\ \ ?W = V" using w w_in_hom v CLS_in_hom comp_CLS B.isomorphic_symmetric CLS_REP hk.leg1_is_map by (metis (mono_tags, lifting) commutative_square_def cs) ultimately show ?thesis using HK h_def k_def PRJ\<^sub>0_def PRJ\<^sub>1_def by auto qed moreover have "\W'. PRJ\<^sub>1 K H \ W' = V \ PRJ\<^sub>0 K H \ W' = U \ W' = ?W" proof - fix W' assume "PRJ\<^sub>1 K H \ W' = V \ PRJ\<^sub>0 K H \ W' = U" hence W': "\W' : dom U \ dom \\?f\\\ \ \\?f\\ \ W' = U \ \\?g\\ \ W' = V" using PRJ\<^sub>0_def PRJ\<^sub>1_def HK h_def k_def apply simp using cs arr_iff_in_hom by blast let ?w' = "REP W'" have w': "B.ide ?w'" using W' ide_REP by auto have fw'_iso_u: "?f \ ?w' \\<^sub>B ?u" proof - have "?f \ ?w' \\<^sub>B REP \\?f\\ \ ?w'" by (metis (no_types, lifting) Cod.simps(1) in_hom_char REP_CLS REP_simps(3) W W' B.hcomp_isomorphic_ide hk.satisfies_T0 B.in_hhomE B.isomorphic_symmetric w' w_in_hom) also have "REP \\?f\\ \ ?w' \\<^sub>B ?u" using W' isomorphic_hcomp_REP cs by blast finally show ?thesis by blast qed have gw'_iso_v: "?g \ ?w' \\<^sub>B ?v" proof - have "?g \ ?w' \\<^sub>B REP \\?g\\ \ ?w'" proof - have "?g \\<^sub>B REP \\?g\\" using REP_CLS B.isomorphic_symmetric hk.leg1_is_map by blast moreover have "B.ide (REP W')" using W' by auto moreover have "src ?f = trg ?w'" using w_in_hom W W' in_hom_char arr_char B.in_hhom_def by (meson fw'_iso_u B.hseqE B.ideD(1) B.isomorphic_implies_ide(1)) ultimately show ?thesis using B.hcomp_isomorphic_ide by simp qed also have "... \\<^sub>B ?v" using W' isomorphic_hcomp_REP cs by blast finally show ?thesis by blast qed show "W' = ?W" proof - have "W' = \\?w'\\" using w' W' CLS_REP by auto also have "... = ?W" proof - have "?w' \\<^sub>B w" using * w W' hk.has_pseudo_pullback(2) u_is_map v_is_map B.isomorphic_symmetric fw'_iso_u gw'_iso_v by blast thus ?thesis using CLS_eqI B.iso_class_eqI w' by blast qed finally show ?thesis by blast qed qed ultimately show ?thesis by auto qed qed qed lemma is_elementary_category_with_pullbacks: shows "elementary_category_with_pullbacks comp PRJ\<^sub>0 PRJ\<^sub>1" .. lemma is_category_with_pullbacks: shows "category_with_pullbacks comp" .. sublocale elementary_category_with_pullbacks comp PRJ\<^sub>0 PRJ\<^sub>1 using is_elementary_category_with_pullbacks by simp end text \ Here we relate the projections of the chosen pullbacks in \Maps(B)\ to the projections associated with the chosen tabulations in \B\. \ context composite_tabulation_in_maps begin interpretation Maps: maps_category V H \ \ src trg .. interpretation r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg s\<^sub>1 r\<^sub>0 using \.leg0_is_map \.leg1_is_map composable by unfold_locales auto lemma prj_char: shows "Maps.PRJ\<^sub>0 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>0 s\<^sub>1 r\<^sub>0\\" and "Maps.PRJ\<^sub>1 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>1 s\<^sub>1 r\<^sub>0\\" proof - have "Maps.arr (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\)" using \.leg1_in_hom Maps.CLS_in_hom \.leg1_is_map Maps.arr_char by auto moreover have "Maps.arr (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\)" using Maps.CLS_in_hom composable r\<^sub>0s\<^sub>1.k_is_map by fastforce moreover have "Maps.cod (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\) = Maps.cod (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\)" unfolding Maps.arr_char using \.leg1_in_hom \.leg0_in_hom by (simp add: Maps.cod_char calculation(1) calculation(2)) ultimately have "Maps.PRJ\<^sub>0 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\tab\<^sub>0 ((Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\))\<^sup>* \ Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\))\\ \ Maps.PRJ\<^sub>1 (Maps.CLS r\<^sub>0) (Maps.CLS s\<^sub>1) = \\tab\<^sub>1 ((Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\))\<^sup>* \ Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\))\\" unfolding Maps.PRJ\<^sub>0_def Maps.PRJ\<^sub>1_def using Maps.CLS_in_hom \.leg1_is_map \.leg0_is_map composable by simp moreover have "r\<^sub>0\<^sup>* \ s\<^sub>1 \ (Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\))\<^sup>* \ Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\)" proof - have "r\<^sub>0 \ Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\)" using Maps.REP_CLS composable isomorphic_symmetric r\<^sub>0s\<^sub>1.k_is_map by fastforce hence 3: "isomorphic r\<^sub>0\<^sup>* (Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\))\<^sup>*" using \.leg0_is_map by (simp add: isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint) moreover have 4: "s\<^sub>1 \ Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \s\<^sub>1\)" using Maps.REP_CLS isomorphic_symmetric r\<^sub>0s\<^sub>1.h_is_map by fastforce ultimately show ?thesis proof - have 1: "src r\<^sub>0\<^sup>* = trg s\<^sub>1" using \.T0.antipar(2) r\<^sub>0s\<^sub>1.cospan by argo have 2: "ide s\<^sub>1" by simp have "src (Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \r\<^sub>0\))\<^sup>* = trg s\<^sub>1" by (metis 3 \.T0.antipar(2) isomorphic_implies_hpar(3) r\<^sub>0s\<^sub>1.cospan) thus ?thesis using 1 2 by (meson 3 4 hcomp_ide_isomorphic hcomp_isomorphic_ide isomorphic_implies_ide(2) isomorphic_transitive) qed qed ultimately have 1: "Maps.PRJ\<^sub>0 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>0 s\<^sub>1 r\<^sub>0\\ \ Maps.PRJ\<^sub>1 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>1 s\<^sub>1 r\<^sub>0\\" using r\<^sub>0s\<^sub>1.isomorphic_implies_same_tab by simp show "Maps.PRJ\<^sub>0 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>0 s\<^sub>1 r\<^sub>0\\" using 1 by simp show "Maps.PRJ\<^sub>1 \\r\<^sub>0\\ \\s\<^sub>1\\ = \\prj\<^sub>1 s\<^sub>1 r\<^sub>0\\" using 1 by simp qed end context identity_in_bicategory_of_spans begin interpretation Maps: maps_category V H \ \ src trg .. interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. notation isomorphic (infix "\" 50) text \ A 1-cell \r\ in a bicategory of spans is a map if and only if the ``input leg'' \tab\<^sub>0 r\ of the chosen tabulation of \r\ is an equivalence map. Since a tabulation of \r\ is unique up to equivalence, and equivalence maps compose, the result actually holds if ``chosen tabulation'' is replaced by ``any tabulation''. \ lemma is_map_iff_tab\<^sub>0_is_equivalence: shows "is_left_adjoint r \ equivalence_map (tab\<^sub>0 r)" proof assume 1: "equivalence_map (tab\<^sub>0 r)" have 2: "quasi_inverses (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>*" proof - obtain g' \' \' where \'\': "equivalence_in_bicategory V H \ \ src trg (tab\<^sub>0 r) g' \' \'" using 1 equivalence_map_def by auto have "adjoint_pair (tab\<^sub>0 r) g'" using \'\' quasi_inverses_def quasi_inverses_are_adjoint_pair by blast moreover have "adjoint_pair (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>*" using T0.adjunction_in_bicategory_axioms adjoint_pair_def by auto ultimately have "g' \ (tab\<^sub>0 r)\<^sup>*" using left_adjoint_determines_right_up_to_iso by simp thus ?thesis using \'\' quasi_inverses_def quasi_inverses_isomorphic_right by blast qed obtain \' \' where \'\': "equivalence_in_bicategory V H \ \ src trg (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>* \' \'" using 2 quasi_inverses_def by auto interpret \'\': equivalence_in_bicategory V H \ \ src trg \tab\<^sub>0 r\ \(tab\<^sub>0 r)\<^sup>*\ \' \' using \'\' by auto have "is_left_adjoint (tab\<^sub>0 r)\<^sup>*" using 2 quasi_inverses_are_adjoint_pair quasi_inverses_symmetric by blast hence "is_left_adjoint (tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*)" using left_adjoints_compose by simp thus "is_left_adjoint r" using yields_isomorphic_representation isomorphic_def left_adjoint_preserved_by_iso' by meson next assume 1: "is_left_adjoint r" have 2: "is_left_adjoint (tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*)" using 1 yields_isomorphic_representation left_adjoint_preserved_by_iso' isomorphic_symmetric isomorphic_def by meson hence "is_left_adjoint (tab\<^sub>0 r)\<^sup>*" using is_ide BS4 [of "tab\<^sub>1 r" "(tab\<^sub>0 r)\<^sup>*"] by auto hence "is_left_adjoint ((tab\<^sub>0 r)\<^sup>* \ tab\<^sub>0 r) \ is_left_adjoint (tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*)" using left_adjoints_compose T0.antipar by simp hence 3: "iso \ \ iso \" using BS3 [of "src (tab\<^sub>0 r)" "(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>0 r" \ \] BS3 [of "tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*" "trg (tab\<^sub>0 r)" \ \] T0.unit_in_hom T0.counit_in_hom obj_is_self_adjoint by auto hence "equivalence_in_bicategory V H \ \ src trg (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>* \ \" apply unfold_locales by auto thus "equivalence_map (tab\<^sub>0 r)" using equivalence_map_def by blast qed text \ The chosen tabulation (and indeed, any other tabulation, which is equivalent) of an object is symmetric in the sense that its two legs are isomorphic. \ lemma obj_has_symmetric_tab: assumes "obj r" shows "tab\<^sub>0 r \ tab\<^sub>1 r" proof - have "tab\<^sub>0 r \ r \ tab\<^sub>0 r" proof - have "trg (tab\<^sub>0 r) = r" using assms by auto moreover have "\\\<^sup>-\<^sup>1[tab\<^sub>0 r] : tab\<^sub>0 r \ trg (tab\<^sub>0 r) \ tab\<^sub>0 r\ \ iso \\<^sup>-\<^sup>1[tab\<^sub>0 r]" using assms by simp ultimately show ?thesis unfolding isomorphic_def by metis qed also have "... \ tab\<^sub>1 r" proof - have "\tab : tab\<^sub>1 r \ r \ tab\<^sub>0 r\" using tab_in_hom by simp moreover have "is_left_adjoint (r \ tab\<^sub>0 r)" using assms left_adjoints_compose obj_is_self_adjoint by simp ultimately show ?thesis using BS3 [of "tab\<^sub>1 r" "r \ tab\<^sub>0 r" tab tab] isomorphic_symmetric isomorphic_def by auto qed finally show ?thesis by simp qed text \ The chosen tabulation of \r\ determines a span in \Maps(B)\. \ lemma determines_span: assumes "ide r" shows "span_in_category Maps.comp \Leg0 = \\tab\<^sub>0 r\\, Leg1 = \\tab\<^sub>1 r\\\" using assms Maps.CLS_in_hom [of "tab\<^sub>0 r"] Maps.CLS_in_hom [of "tab\<^sub>1 r"] tab\<^sub>0_in_hom tab\<^sub>1_in_hom apply unfold_locales by fastforce end subsection "Arrows of Tabulations in Maps" text \ Here we consider the situation of two tabulations: a tabulation \\\ of \r\ and a tabulation \\\ of \s\, both ``legs'' of each tabulation being maps, together with an arbitrary 2-cell \\\ : r \ s\\. The 2-cell \\\ at the base composes with the tabulation \\\ to yield a 2-cell \\ = (\ \ r\<^sub>0) \ \\ ``over'' s. By property \T1\ of tabulation \\\, this induces a map from the apex of \\\ to the apex of \\\, which together with the other data forms a triangular prism whose sides commute up to (unique) isomorphism. \ text \ $$ \xymatrix{ && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^<-1>\sigma} & \\ &{\rm trg}~s && {\rm src}~s \ar[ll]^{s} \\ & \rrtwocell\omit{^\mu} &&\\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \ar@ {.>}[uuur]^<>(0.3){{\rm chine}} \dtwocell\omit{^\rho}& \\ {\rm trg}~r \ar@ {=}[uuur] && {\rm src}~r \ar[ll]^{r} \ar@ {=}[uuur] } $$ \ locale arrow_of_tabulations_in_maps = bicategory_of_spans V H \ \ src trg + \: tabulation_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 + \: tabulation_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and r :: 'a and \ :: 'a and r\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s :: 'a and \ :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a and \ :: 'a + assumes in_hom: "\\ : r \ s\" begin abbreviation (input) \ where "\ \ (\ \ r\<^sub>0) \ \" lemma \_in_hom [intro]: shows "\\ : src \ \ trg \\" and "\\ : r\<^sub>1 \ s \ r\<^sub>0\" proof - show "\\ : r\<^sub>1 \ s \ r\<^sub>0\" using in_hom \.leg0_in_hom(2) \.tab_in_vhom' by auto thus "\\ : src \ \ trg \\" by (metis \.tab_simps(3) \.base_in_hom(2) \.tab_simps(3) \.base_in_hom(2) arrI in_hom seqI' vcomp_in_hhom vseq_implies_hpar(1-2)) qed lemma \_simps [simp]: shows "arr \" and "src \ = src \" and "trg \ = trg \" and "dom \ = r\<^sub>1" and "cod \ = s \ r\<^sub>0" using \_in_hom by auto abbreviation is_induced_map where "is_induced_map w \ \.is_induced_by_cell w r\<^sub>0 \" text \ The following is an equivalent restatement, in elementary terms, of the conditions for being an induced map. \ abbreviation (input) is_induced_map' where "is_induced_map' w \ ide w \ (\\ \. \\ : s\<^sub>0 \ w \ r\<^sub>0\ \ \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ iso \ \ \ = (s \ \) \ \[s, s\<^sub>0, w] \ (\ \ w) \ \)" lemma is_induced_map_iff: shows "is_induced_map w \ is_induced_map' w" proof assume w: "is_induced_map' w" show "is_induced_map w" proof have 1: "dom \ = r\<^sub>1" by auto interpret w: arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 \dom \\ s \ s\<^sub>0 s\<^sub>1 w proof - have "arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 w" using w apply unfold_locales by auto thus "arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 (dom \) s \ s\<^sub>0 s\<^sub>1 w" using 1 by simp qed show "arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 (dom \) s\<^sub>0 s\<^sub>1 w" using w.arrow_of_spans_of_maps_axioms by auto show "\.composite_cell w w.the_\ \ w.the_\ = \" proof - obtain \ \ where \\: "\\ : s\<^sub>0 \ w \ r\<^sub>0\ \ \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ iso \ \ \ = (s \ \) \ \[s, s\<^sub>0, w] \ (\ \ w) \ \" using w w.the_\_props(1) by auto have "(s \ \) \ \[s, s\<^sub>0, w] \ (\ \ w) \ \ = \" using \\ by argo moreover have "\ = w.the_\ \ \ = w.the_\" using \\ 1 w.the_\_props(1) w.leg0_uniquely_isomorphic w.leg1_uniquely_isomorphic by auto ultimately show ?thesis using comp_assoc by simp qed qed next assume w: "is_induced_map w" show "is_induced_map' w" proof - interpret w: arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 w using w in_hom by auto interpret w: arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 w .. have "dom \ = r\<^sub>1" by auto thus ?thesis using w comp_assoc w.the_\_props(1) w.the_\_props(2) w.uw\ by metis qed qed lemma exists_induced_map: shows "\w. is_induced_map w" proof - obtain w \ \ where w\\: "ide w \ \\ : s\<^sub>0 \ w \ r\<^sub>0\ \ \\ : r\<^sub>1 \ s\<^sub>1 \ w\ \ iso \ \ \ = (s \ \) \ \[s, s\<^sub>0, w] \ (\ \ w) \ \" using \_in_hom \.ide_leg0 \.T1 comp_assoc by (metis in_homE) thus ?thesis using is_induced_map_iff by blast qed lemma induced_map_unique: assumes "is_induced_map w" and "is_induced_map w'" shows "w \ w'" using assms \.induced_map_unique by blast definition chine where "chine \ SOME w. is_induced_map w" lemma chine_is_induced_map: shows "is_induced_map chine" unfolding chine_def using exists_induced_map someI_ex [of is_induced_map] by simp lemma chine_in_hom [intro]: shows "\chine : src r\<^sub>0 \ src s\<^sub>0\" and "\chine: chine \ chine\" proof - show "\chine : src r\<^sub>0 \ src s\<^sub>0\" using chine_is_induced_map by (metis \_simps(1) \_simps(4) \.leg1_simps(3) \.ide_base \.ide_leg0 \.leg0_simps(3) \.tab_simps(2) arrow_of_spans_of_maps.is_ide arrow_of_spans_of_maps.the_\_simps(2) assoc_simps(2) hseqE in_hhom_def seqE src_vcomp vseq_implies_hpar(1)) show "\chine: chine \ chine\" using chine_is_induced_map by (meson arrow_of_spans_of_maps.is_ide ide_in_hom(2)) qed lemma chine_simps [simp]: shows "arr chine" and "ide chine" and "src chine = src r\<^sub>0" and "trg chine = src s\<^sub>0" and "dom chine = chine" and "cod chine = chine" using chine_in_hom apply auto by (meson arrow_of_spans_of_maps.is_ide chine_is_induced_map) end sublocale arrow_of_tabulations_in_maps \ arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 chine using chine_is_induced_map is_induced_map_iff by unfold_locales auto sublocale arrow_of_tabulations_in_maps \ arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 chine .. context arrow_of_tabulations_in_maps begin text \ The two factorizations of the composite 2-cell \\\ amount to a naturality condition. \ lemma \_naturality: shows "(\ \ r\<^sub>0) \ \ = (s \ the_\) \ \[s, s\<^sub>0, chine] \ (\ \ chine) \ the_\" using chine_is_induced_map is_induced_map_iff by (metis leg0_uniquely_isomorphic(2) leg1_uniquely_isomorphic(2) the_\_props(1) uw\) lemma induced_map_preserved_by_iso: assumes "is_induced_map w" and "isomorphic w w'" shows "is_induced_map w'" proof - interpret w: arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 w using assms in_hom by auto interpret w: arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 w .. obtain \ where \: "\\ : w \ w'\ \ iso \" using assms(2) isomorphic_def by auto show ?thesis proof interpret w': arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 \dom \\ s\<^sub>0 s\<^sub>1 w' proof show "is_left_adjoint r\<^sub>0" by (simp add: \.satisfies_T0) show "is_left_adjoint (dom \)" by (simp add: \.leg1_is_map) show "ide w'" using assms by force show "\\. \\ : s\<^sub>0 \ w' \ r\<^sub>0\" using \ w.the_\_props \.leg0_in_hom(2) assms(2) comp_in_homI hcomp_in_vhom inv_in_hom isomorphic_implies_hpar(4) w.the_\_simps(4) w.w_simps(4) by metis have "\(s\<^sub>1 \ \) \ w.the_\ : r\<^sub>1 \ s\<^sub>1 \ w'\ \ iso ((s\<^sub>1 \ \) \ w.the_\)" proof (intro conjI) show "\(s\<^sub>1 \ \) \ w.the_\ : r\<^sub>1 \ s\<^sub>1 \ w'\" using \ w.the_\_props by (intro comp_in_homI, auto) thus "iso ((s\<^sub>1 \ \) \ w.the_\)" using \ w.the_\_props by (meson \.ide_leg1 arrI iso_hcomp hseqE ide_is_iso isos_compose seqE) qed hence "\\. \\ : r\<^sub>1 \ s\<^sub>1 \ w'\ \ iso \" by auto thus "\\. \\ : dom \ \ s\<^sub>1 \ w'\ \ iso \" by auto qed interpret w': arrow_of_spans_of_maps_to_tabulation V H \ \ src trg r\<^sub>0 \dom \\ s \ s\<^sub>0 s\<^sub>1 w' .. show "arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 (dom \) s\<^sub>0 s\<^sub>1 w'" using w'.arrow_of_spans_of_maps_axioms by auto show "\.composite_cell w' w'.the_\ \ w'.the_\ = \" proof - have 1: "w'.the_\ = w.the_\ \ (s\<^sub>0 \ inv \)" proof - have "\w.the_\ \ (s\<^sub>0 \ inv \) : s\<^sub>0 \ w' \ r\<^sub>0\" using w.the_\_props \ by (intro comp_in_homI, auto) moreover have "\w'.the_\ : s\<^sub>0 \ w' \ r\<^sub>0\" using w'.the_\_props by simp ultimately show ?thesis using w'.leg0_uniquely_isomorphic(2) by blast qed moreover have "w'.the_\ = (s\<^sub>1 \ \) \ w.the_\" proof - have "\(s\<^sub>1 \ \) \ w.the_\ : dom \ \ s\<^sub>1 \ w'\" using w.the_\_props \ by (intro comp_in_homI, auto) moreover have "iso ((s\<^sub>1 \ \) \ w.the_\)" using w.the_\_props \ iso_hcomp by (meson \.ide_leg1 arrI calculation hseqE ide_is_iso isos_compose seqE) ultimately show ?thesis using w'.the_\_props w'.leg1_uniquely_isomorphic(2) by blast qed ultimately have "\.composite_cell w' w'.the_\ \ w'.the_\ = \.composite_cell w' (w.the_\ \ (s\<^sub>0 \ inv \)) \ (s\<^sub>1 \ \) \ w.the_\" by simp also have "... = (s \ w.the_\ \ (s\<^sub>0 \ inv \)) \ \[s, s\<^sub>0, w'] \ (\ \ w') \ (s\<^sub>1 \ \) \ w.the_\" using comp_assoc by presburger also have "... = (s \ w.the_\) \ ((s \ s\<^sub>0 \ inv \) \ \[s, s\<^sub>0, w'] \ (\ \ w') \ (s\<^sub>1 \ \)) \ w.the_\" using 1 comp_assoc w'.the_\_simps(1) whisker_left by auto also have "... = (s \ w.the_\) \ (\[s, s\<^sub>0, w] \ (\ \ w)) \ w.the_\" proof - have "(s \ s\<^sub>0 \ inv \) \ \[s, s\<^sub>0, w'] \ (\ \ w') \ (s\<^sub>1 \ \) = \[s, s\<^sub>0, w] \ (\ \ w)" proof - have "(s \ s\<^sub>0 \ inv \) \ \[s, s\<^sub>0, w'] \ (\ \ w') \ (s\<^sub>1 \ \) = \[s, s\<^sub>0, w] \ ((s \ s\<^sub>0) \ inv \) \ (\ \ w') \ (s\<^sub>1 \ \)" proof - have "(s \ s\<^sub>0 \ inv \) \ \[s, s\<^sub>0, w'] = \[s, s\<^sub>0, w] \ ((s \ s\<^sub>0) \ inv \)" using assms \ assoc_naturality [of s s\<^sub>0 "inv \"] w.w_simps(4) by (metis \.leg0_simps(2-5) \.base_simps(2-4) arr_inv cod_inv dom_inv in_homE trg_cod) thus ?thesis using comp_assoc by metis qed also have "... = \[s, s\<^sub>0, w] \ (\ \ w) \ (s\<^sub>1 \ inv \) \ (s\<^sub>1 \ \)" proof - have "((s \ s\<^sub>0) \ inv \) \ (\ \ w') = (\ \ w) \ (s\<^sub>1 \ inv \)" using \ comp_arr_dom comp_cod_arr in_hhom_def interchange [of "s \ s\<^sub>0" \ "inv \" w'] interchange [of \ s\<^sub>1 w "inv \"] by auto thus ?thesis using comp_assoc by metis qed also have "... = \[s, s\<^sub>0, w] \ (\ \ w)" proof - have "(\ \ w) \ (s\<^sub>1 \ inv \) \ (s\<^sub>1 \ \) = \ \ w" proof - have "(\ \ w) \ (s\<^sub>1 \ inv \) \ (s\<^sub>1 \ \) = (\ \ w) \ (s\<^sub>1 \ inv \ \ \)" using \ whisker_left in_hhom_def by auto also have "... = (\ \ w) \ (s\<^sub>1 \ w)" using \ comp_inv_arr' by auto also have "... = \ \ w" using whisker_right [of w \ s\<^sub>1] comp_arr_dom in_hhom_def by auto finally show ?thesis by blast qed thus ?thesis by simp qed finally show ?thesis by simp qed thus ?thesis by simp qed also have "... = \" using assms(1) comp_assoc w.is_ide w.the_\_props(1) w.the_\_props(1) by simp finally show ?thesis using comp_assoc by auto qed qed qed end text \ In the special case that \\\ is an identity 2-cell, the induced map from the apex of \\\ to the apex of \\\ is an equivalence map. \ locale identity_arrow_of_tabulations_in_maps = arrow_of_tabulations_in_maps + assumes is_ide: "ide \" begin lemma r_eq_s: shows "r = s" using is_ide by (metis ide_char in_hom in_homE) lemma \_eq_\: shows "\ = \" by (meson \_simps(1) comp_ide_arr ide_hcomp hseq_char' ide_u is_ide seqE seq_if_composable) lemma chine_is_equivalence: shows "equivalence_map chine" proof - obtain w w' \ \ \ \ \' \' where e: "equivalence_in_bicategory (\) (\) \ \ src trg w' w \ \ \ \w : src s\<^sub>0 \ src r\<^sub>0\ \ \w' : src r\<^sub>0 \ src s\<^sub>0\ \ \\ : r\<^sub>0 \ w \ s\<^sub>0\ \ \\ : s\<^sub>1 \ r\<^sub>1 \ w\ \ iso \ \ \ = (s \ \) \ \[s, r\<^sub>0, w] \ (\ \ w) \ \ \ \\' : s\<^sub>0 \ w' \ r\<^sub>0\ \ \\' : r\<^sub>1 \ s\<^sub>1 \ w'\ \ iso \' \ \ = (s \ \') \ \[s, s\<^sub>0, w'] \ (\ \ w') \ \'" using r_eq_s \.apex_unique_up_to_equivalence [of \ r\<^sub>0 r\<^sub>1] \.tabulation_axioms by blast have w': "equivalence_map w'" using e equivalence_map_def by auto hence "is_induced_map w'" using e r_eq_s \_eq_\ is_induced_map_iff comp_assoc equivalence_map_is_ide by metis hence "isomorphic chine w'" using induced_map_unique chine_is_induced_map by simp thus ?thesis using w' equivalence_map_preserved_by_iso isomorphic_symmetric by blast qed end text \ The following gives an interpretation of @{locale arrow_of_tabulations_in_maps} in the special case that the tabulations are those that we have chosen for the domain and codomain of the underlying 2-cell \\\ : r \ s\\. In this case, we can recover \\\ from \\\ via adjoint transpose. \ locale arrow_in_bicategory_of_spans = bicategory_of_spans V H \ \ src trg + r: identity_in_bicategory_of_spans V H \ \ src trg r + s: identity_in_bicategory_of_spans V H \ \ src trg s for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and r :: 'a and s :: 'a and \ :: 'a + assumes in_hom: "\\ : r \ s\" begin abbreviation (input) r\<^sub>0 where "r\<^sub>0 \ tab\<^sub>0 r" abbreviation (input) r\<^sub>1 where "r\<^sub>1 \ tab\<^sub>1 r" abbreviation (input) s\<^sub>0 where "s\<^sub>0 \ tab\<^sub>0 s" abbreviation (input) s\<^sub>1 where "s\<^sub>1 \ tab\<^sub>1 s" lemma is_arrow_of_tabulations_in_maps: shows "arrow_of_tabulations_in_maps V H \ \ src trg r r.tab r\<^sub>0 r\<^sub>1 s s.tab s\<^sub>0 s\<^sub>1 \" using in_hom by unfold_locales auto end sublocale identity_in_bicategory_of_spans \ arrow_in_bicategory_of_spans V H \ \ src trg r r r apply unfold_locales using is_ide by auto context arrow_in_bicategory_of_spans begin interpretation arrow_of_tabulations_in_maps V H \ \ src trg r r.tab r\<^sub>0 r\<^sub>1 s s.tab s\<^sub>0 s\<^sub>1 \ using is_arrow_of_tabulations_in_maps by simp interpretation r: arrow_of_tabulations_in_maps V H \ \ src trg r r.tab r\<^sub>0 r\<^sub>1 r r.tab r\<^sub>0 r\<^sub>1 r using r.is_arrow_of_tabulations_in_maps by simp lemma \_in_terms_of_\: shows "\ = r.T0.trnr\<^sub>\ (cod \) \ \ inv (r.T0.trnr\<^sub>\ r r.tab)" proof - have \: "arr \" using in_hom by auto have "\ \ r.T0.trnr\<^sub>\ r r.tab = r.T0.trnr\<^sub>\ s \" proof - have "\ \ r.T0.trnr\<^sub>\ r r.tab = (\ \ \[r]) \ (r \ r.\) \ \[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \ (r.tab \ (tab\<^sub>0 r)\<^sup>*)" unfolding r.T0.trnr\<^sub>\_def using comp_assoc by presburger also have "... = \[s] \ ((\ \ src \) \ (r \ r.\)) \ \[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \ (r.tab \ (tab\<^sub>0 r)\<^sup>*)" using \ runit_naturality comp_assoc by (metis in_hom in_homE) also have "... = \[s] \ (s \ r.\) \ ((\ \ tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*) \ \[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*]) \ (r.tab \ (tab\<^sub>0 r)\<^sup>*)" proof - have "(\ \ src \) \ (r \ r.\) = \ \ r.\" using \ interchange comp_arr_dom comp_cod_arr by (metis in_hom in_homE r.T0.counit_simps(1) r.T0.counit_simps(3) r.u_simps(3) src_dom) also have "... = (s \ r.\) \ (\ \ tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*)" using in_hom interchange [of s \ r.\ "tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*"] comp_arr_dom comp_cod_arr r.T0.counit_simps(1) r.T0.counit_simps(2) by auto finally have "(\ \ src \) \ (r \ r.\) = (s \ r.\) \ (\ \ tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*)" by blast thus ?thesis using comp_assoc by presburger qed also have "... = \[s] \ (s \ r.\) \ \[s, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \ ((\ \ tab\<^sub>0 r) \ (tab\<^sub>0 r)\<^sup>*) \ (r.tab \ (tab\<^sub>0 r)\<^sup>*)" proof - have "(\ \ tab\<^sub>0 r \ (tab\<^sub>0 r)\<^sup>*) \ \[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] = \[s, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \ ((\ \ tab\<^sub>0 r) \ (tab\<^sub>0 r)\<^sup>*)" using \ assoc_naturality [of \ "tab\<^sub>0 r" "(tab\<^sub>0 r)\<^sup>*"] by (metis ide_char in_hom in_homE r.T0.antipar(1) r.T0.ide_right r.u_simps(3) src_dom u_simps(2) u_simps(4-5)) thus ?thesis using comp_assoc by presburger qed also have "... = \[s] \ (s \ r.\) \ \[s, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \ ((\ \ tab\<^sub>0 r) \ r.tab \ (tab\<^sub>0 r)\<^sup>*)" using \ whisker_right \_simps(1) by auto also have "... = r.T0.trnr\<^sub>\ s \" unfolding r.T0.trnr\<^sub>\_def by simp finally show ?thesis by blast qed thus ?thesis using \ r.yields_isomorphic_representation invert_side_of_triangle(2) by (metis in_hom in_homE seqI') qed end subsubsection "Vertical Composite" locale vertical_composite_of_arrows_of_tabulations_in_maps = bicategory_of_spans V H \ \ src trg + \: tabulation_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 + \: tabulation_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 + \: tabulation_in_maps V H \ \ src trg t \ t\<^sub>0 t\<^sub>1 + \: arrow_of_tabulations_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 \ + \: arrow_of_tabulations_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 t \ t\<^sub>0 t\<^sub>1 \ for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and r :: 'a and \ :: 'a and r\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s :: 'a and \ :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a and t :: 'a and \ :: 'a and t\<^sub>0 :: 'a and t\<^sub>1 :: 'a and \ :: 'a and \ :: 'a begin text \ $$ \xymatrix{ &&& {\rm src}~\tau \ar[dl]_{t_1} \ar[dr]^{t_0} \dtwocell\omit{^<-1>\tau} & \\ &&{\rm trg}~t && {\rm src}~t \ar[ll]^{s} \\ && \rrtwocell\omit{^\pi} && \\ && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \ar[uuur]^<>(0.3){\pi.{\rm chine}} \dtwocell\omit{^<-1>\sigma} & \\ &{\rm trg}~s \ar@ {=}[uuur] && {\rm src}~s \ar[ll]^{s} \ar@ {=}[uuur] \\ & \rrtwocell\omit{^\mu} &&\\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \ar[uuur]^<>(0.3){\mu.{\rm chine}} \dtwocell\omit{^\rho} & \\ {\rm trg}~r \ar@ {=}[uuur] && {\rm src}~r \ar[ll]^{r} \ar@ {=}[uuur] } $$ \ notation isomorphic (infix "\" 50) interpretation arrow_of_tabulations_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 t \ t\<^sub>0 t\<^sub>1 \\ \ \\ using \.in_hom \.in_hom by (unfold_locales, blast) lemma is_arrow_of_tabulations_in_maps: shows "arrow_of_tabulations_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 t \ t\<^sub>0 t\<^sub>1 (\ \ \)" .. lemma chine_char: shows "chine \ \.chine \ \.chine" proof - have "is_induced_map (\.chine \ \.chine)" proof - let ?f = "\.chine" have f: "\?f : src \ \ src \\ \ is_left_adjoint ?f \ ide ?f \ \.is_induced_map ?f" using \.chine_is_induced_map \.is_map by auto let ?g = "\.chine" have g: "\?g : src \ \ src \\ \ is_left_adjoint ?g \ ide ?g \ \.is_induced_map ?g" using \.chine_is_induced_map \.is_map by auto let ?\ = "\.the_\ \ (\.the_\ \ ?f) \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]" let ?\ = "\[t\<^sub>1, ?g, ?f] \ (\.the_\ \ ?f) \ \.the_\" have \: "\?\ : t\<^sub>0 \ ?g \ ?f \ r\<^sub>0\" using f g \.the_\_props \.the_\_props by (intro comp_in_homI hcomp_in_vhom, auto+) have \: "\?\ : r\<^sub>1 \ t\<^sub>1 \ ?g \ ?f\" using f g \.the_\_props \.the_\_props by (intro comp_in_homI hcomp_in_vhom, auto) interpret gf: arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 r\<^sub>1 t\<^sub>0 t\<^sub>1 \?g \ ?f\ proof show "ide (?g \ ?f)" by simp show "\\. \\ : t\<^sub>0 \ ?g \ ?f \ r\<^sub>0\" using \ by blast show "\\. \\ : r\<^sub>1 \ t\<^sub>1 \ ?g \ ?f\ \ iso \" using \ \.the_\_props \.the_\_props \.the_\_props \.the_\_props isos_compose [of "\.the_\" "\.the_\"] \.is_ide \ \ide (\.chine \ \.chine)\ \.uw\ \.w_simps(4) \.ide_leg1 \.leg1_simps(3) arrI hseq_char ideD(1) ide_is_iso iso_assoc iso_hcomp isos_compose seqE by metis qed show ?thesis proof (intro conjI) have \_eq: "?\ = gf.the_\" using \ gf.the_\_props gf.leg0_uniquely_isomorphic by auto have \_eq: "?\ = gf.the_\" using \ gf.the_\_props gf.leg1_uniquely_isomorphic by auto have A: "src ?g = trg ?f" using f g by fastforce show "arrow_of_spans_of_maps V H \ \ src trg r\<^sub>0 (dom \) t\<^sub>0 t\<^sub>1 (?g \ ?f)" using gf.arrow_of_spans_of_maps_axioms by simp have "((t \ gf.the_\) \ \[t, t\<^sub>0, ?g \ ?f] \ (\ \ ?g \ ?f)) \ gf.the_\ = \" proof - have "\ = (\ \ r\<^sub>0) \ (\ \ r\<^sub>0) \ \" using whisker_right comp_assoc by (metis \_simps(1) hseqE ide_u seqE) also have "... = ((\ \ r\<^sub>0) \ (s \ \.the_\)) \ \[s, s\<^sub>0, ?f] \ (\ \ ?f) \ \.the_\" using \.\_naturality comp_assoc by simp also have "... = (t \ \.the_\) \ ((\ \ s\<^sub>0 \ ?f) \ \[s, s\<^sub>0, ?f]) \ (\ \ ?f) \ \.the_\" proof - have "(\ \ r\<^sub>0) \ (s \ \.the_\) = \ \ \.the_\" using f comp_arr_dom comp_cod_arr \.the_\_props \.in_hom interchange [of \ s r\<^sub>0 \.the_\] by (metis in_homE) also have "... = (t \ \.the_\) \ (\ \ s\<^sub>0 \ ?f)" using f comp_arr_dom comp_cod_arr \.the_\_props \.in_hom interchange [of t \ \.the_\ "s\<^sub>0 \ ?f"] by (metis in_homE) finally have "(\ \ r\<^sub>0) \ (s \ \.the_\) = (t \ \.the_\) \ (\ \ s\<^sub>0 \ ?f)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (t \ \.the_\) \ \[t, s\<^sub>0, ?f] \ (((\ \ s\<^sub>0) \ ?f) \ (\ \ ?f)) \ \.the_\" proof - have "(\ \ s\<^sub>0 \ ?f) \ \[s, s\<^sub>0, ?f] = \[t, s\<^sub>0, ?f] \ ((\ \ s\<^sub>0) \ ?f)" using f assoc_naturality [of \ s\<^sub>0 ?f] \.in_hom by auto thus ?thesis using comp_assoc by presburger qed also have "... = (t \ \.the_\) \ \[t, s\<^sub>0, ?f] \ (\.\ \ ?f) \ \.the_\" using whisker_right comp_assoc by simp also have "... = (t \ \.the_\) \ \[t, s\<^sub>0, ?f] \ ((t \ \.the_\) \ \[t, t\<^sub>0, ?g] \ (\ \ ?g) \ \.the_\ \ ?f) \ \.the_\" using \.\_naturality by simp also have "... = (t \ \.the_\) \ \[t, s\<^sub>0, ?f] \ (((t \ \.the_\) \ ?f) \ (\[t, t\<^sub>0, ?g] \ ?f) \ ((\ \ ?g) \ ?f) \ (\.the_\ \ ?f)) \ \.the_\" using f g \.the_\_props \.the_\_props whisker_right by (metis \.\_simps(1) \.\_naturality seqE) also have "... = (t \ \.the_\) \ (\[t, s\<^sub>0, ?f] \ ((t \ \.the_\) \ ?f)) \ (\[t, t\<^sub>0, ?g] \ ?f) \ ((\ \ ?g) \ ?f) \ (\.the_\ \ ?f) \ \.the_\" using comp_assoc by presburger also have "... = (t \ \.the_\) \ (t \ \.the_\ \ ?f) \ (\[t, t\<^sub>0 \ ?g, ?f] \ (\[t, t\<^sub>0, ?g] \ ?f)) \ ((\ \ ?g) \ ?f) \ (\.the_\ \ ?f) \ \.the_\" using f \.the_\_props assoc_naturality [of t "\.the_\" ?f] \.\_simps(3) comp_assoc by auto also have "... = (t \ \.the_\) \ (t \ \.the_\ \ ?f) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \ \[t, t\<^sub>0, ?g \ ?f] \ (\[t \ t\<^sub>0, ?g, ?f] \ ((\ \ ?g) \ ?f)) \ (\.the_\ \ ?f) \ \.the_\" proof - have "seq \[t, t\<^sub>0, ?g \ ?f] \[t \ t\<^sub>0, ?g, ?f]" using f g by fastforce moreover have "inv (t \ \[t\<^sub>0, ?g, ?f]) = t \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]" using f g by simp moreover have "iso (t \ \[t\<^sub>0, ?g, ?f])" using f g by simp have "\[t, t\<^sub>0 \ ?g, ?f] \ (\[t, t\<^sub>0, ?g] \ ?f) = (t \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \ \[t, t\<^sub>0, ?g \ ?f] \ \[t \ t\<^sub>0, ?g, ?f]" proof - have "seq \[t, t\<^sub>0, ?g \ ?f] \[t \ t\<^sub>0, ?g, ?f]" using f g by fastforce moreover have "inv (t \ \[t\<^sub>0, ?g, ?f]) = t \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]" using f g by simp moreover have "iso (t \ \[t\<^sub>0, ?g, ?f])" using f g by simp ultimately show ?thesis using A f g pentagon invert_side_of_triangle(1) by (metis \.w_simps(4) \.ide_base \.ide_leg0 \.leg0_simps(3)) qed thus ?thesis using comp_assoc by presburger qed also have "... = ((t \ \.the_\) \ (t \ \.the_\ \ ?f)) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \ \[t, t\<^sub>0, ?g \ ?f] \ (\ \ ?g \ ?f) \ \[t\<^sub>1, ?g, ?f] \ (\.the_\ \ ?f) \ \.the_\" using f g assoc_naturality [of \ ?g ?f] comp_assoc by simp also have "... = (t \ \.the_\ \ (\.the_\ \ ?f) \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \ \[t, t\<^sub>0, ?g \ ?f] \ (\ \ ?g \ ?f) \ \[t\<^sub>1, ?g, ?f] \ (\.the_\ \ ?f) \ \.the_\" proof - have 1: "seq \.the_\ ((\.the_\ \ ?f) \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f])" using \_eq by auto hence "t \ (\.the_\ \ ?f) \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f] = (t \ \.the_\ \ ?f) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f])" using whisker_left \.ide_base by blast thus ?thesis using 1 whisker_left \.ide_base comp_assoc by presburger qed also have "... = ((t \ gf.the_\) \ \[t, t\<^sub>0, ?g \ ?f] \ (\ \ ?g \ ?f)) \ gf.the_\" using \_eq \_eq by (simp add: comp_assoc) finally show ?thesis using comp_assoc by presburger qed thus "((t \ gf.the_\) \ \[t, t\<^sub>0, ?g \ ?f] \ (\ \ ?g \ ?f)) \ arrow_of_spans_of_maps.the_\ (\) (\) (dom ((\ \ \ \ r\<^sub>0) \ \)) t\<^sub>1 (?g \ ?f) = \" by simp qed qed thus ?thesis using chine_is_induced_map induced_map_unique by simp qed end sublocale vertical_composite_of_arrows_of_tabulations_in_maps \ arrow_of_tabulations_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 t \ t\<^sub>0 t\<^sub>1 "\ \ \" using is_arrow_of_tabulations_in_maps by simp subsubsection "Horizontal Composite" locale horizontal_composite_of_arrows_of_tabulations_in_maps = bicategory_of_spans V H \ \ src trg + \: tabulation_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 + \: tabulation_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 + \: tabulation_in_maps V H \ \ src trg t \ t\<^sub>0 t\<^sub>1 + \: tabulation_in_maps V H \ \ src trg u \ u\<^sub>0 u\<^sub>1 + \\: composite_tabulation_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 s \ s\<^sub>0 s\<^sub>1 + \\: composite_tabulation_in_maps V H \ \ src trg t \ t\<^sub>0 t\<^sub>1 u \ u\<^sub>0 u\<^sub>1 + \: arrow_of_tabulations_in_maps V H \ \ src trg r \ r\<^sub>0 r\<^sub>1 t \ t\<^sub>0 t\<^sub>1 \ + \: arrow_of_tabulations_in_maps V H \ \ src trg s \ s\<^sub>0 s\<^sub>1 u \ u\<^sub>0 u\<^sub>1 \ for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and r :: 'a and \ :: 'a and r\<^sub>0 :: 'a and r\<^sub>1 :: 'a and s :: 'a and \ :: 'a and s\<^sub>0 :: 'a and s\<^sub>1 :: 'a and t :: 'a and \ :: 'a and t\<^sub>0 :: 'a and t\<^sub>1 :: 'a and u :: 'a and \ :: 'a and u\<^sub>0 :: 'a and u\<^sub>1 :: 'a and \ :: 'a and \ :: 'a begin text \ $$ \xymatrix{ &&& {\rm src}~t_0u_1.\phi \ar[dl]_{\tau\mu.p_1} \ar[dr]^{\tau\mu.p_0} \ddtwocell\omit{^{t_0u_1.\phi}} \\ && {\rm src}~\tau \ar[dl]_{t_1} \ar[dr]^<>(0.4){t_0} \dtwocell\omit{^<-1>\tau} && {\rm src}~\mu \ar[dl]_{u_1} \ar[dr]^{u_0} \dtwocell\omit{^<-1>\mu} & \\ & {\rm trg}~t && {\rm src}~t = {\rm trg}~u \ar[ll]^{t} && {\rm src}~u \ar[ll]^{u} \\ & \xtwocell[r]{}\omit{^\omega} & {\rm src}~r_0s_1.\phi \ar[uuur]_<>(0.2){{\rm chine}} \ar[dl]^{\rho\sigma.p_1} \ar[dr]_{\rho\sigma.p_0\hspace{20pt}} \ddtwocell\omit{^{r_0s_1.\phi}} & \rrtwocell\omit{^\chi} && \\ & {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \ar[uuur]^<>(0.4){\omega.{\rm chine}} \dtwocell\omit{^\rho} && {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \ar[uuur]^<>(0.4){\chi.{\rm chine}} \dtwocell\omit{^<-1>\sigma} & \\ {\rm trg}~r \ar@ {=}[uuur] && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} \ar@ {=}[uuur] && {\rm src}~s \ar[ll]^{s} \ar@ {=}[uuur] \\ } $$ \ notation isomorphic (infix "\" 50) interpretation arrow_of_tabulations_in_maps V H \ \ src trg \r \ s\ \\.tab \s\<^sub>0 \ \\.p\<^sub>0\ \r\<^sub>1 \ \\.p\<^sub>1\ \t \ u\ \\.tab \u\<^sub>0 \ \\.p\<^sub>0\ \t\<^sub>1 \ \\.p\<^sub>1\ \\ \ \\ using \\.composable \.in_hom \.in_hom by unfold_locales auto lemma is_arrow_of_tabulations_in_maps: shows "arrow_of_tabulations_in_maps V H \ \ src trg (r \ s) \\.tab (s\<^sub>0 \ \\.p\<^sub>0) (r\<^sub>1 \ \\.p\<^sub>1) (t \ u) \\.tab (u\<^sub>0 \ \\.p\<^sub>0) (t\<^sub>1 \ \\.p\<^sub>1) (\ \ \)" .. sublocale arrow_of_tabulations_in_maps V H \ \ src trg \r \ s\ \\.tab \s\<^sub>0 \ \\.p\<^sub>0\ \r\<^sub>1 \ \\.p\<^sub>1\ \t \ u\ \\.tab \u\<^sub>0 \ \\.p\<^sub>0\ \t\<^sub>1 \ \\.p\<^sub>1\ \\ \ \\ using is_arrow_of_tabulations_in_maps by simp interpretation Maps: maps_category V H \ \ src trg .. notation Maps.comp (infixr "\" 55) interpretation r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans \(\)\ \(\)\ \ \ src trg s\<^sub>1 r\<^sub>0 using \.leg0_is_map \.leg1_is_map \\.composable apply unfold_locales by auto interpretation r\<^sub>0s\<^sub>1: arrow_of_tabulations_in_maps \(\)\ \(\)\ \ \ src trg \r\<^sub>0\<^sup>* \ s\<^sub>1\ r\<^sub>0s\<^sub>1.tab r\<^sub>0s\<^sub>1.p\<^sub>0 r\<^sub>0s\<^sub>1.p\<^sub>1 \r\<^sub>0\<^sup>* \ s\<^sub>1\ r\<^sub>0s\<^sub>1.tab r\<^sub>0s\<^sub>1.p\<^sub>0 r\<^sub>0s\<^sub>1.p\<^sub>1 \r\<^sub>0\<^sup>* \ s\<^sub>1\ using r\<^sub>0s\<^sub>1.is_arrow_of_tabulations_in_maps by simp interpretation t\<^sub>0u\<^sub>1: cospan_of_maps_in_bicategory_of_spans \(\)\ \(\)\ \ \ src trg u\<^sub>1 t\<^sub>0 using \.leg0_is_map \.leg1_is_map \\.composable apply unfold_locales by auto interpretation t\<^sub>0u\<^sub>1: arrow_of_tabulations_in_maps \(\)\ \(\)\ \ \ src trg \t\<^sub>0\<^sup>* \ u\<^sub>1\ t\<^sub>0u\<^sub>1.tab t\<^sub>0u\<^sub>1.p\<^sub>0 t\<^sub>0u\<^sub>1.p\<^sub>1 \t\<^sub>0\<^sup>* \ u\<^sub>1\ t\<^sub>0u\<^sub>1.tab t\<^sub>0u\<^sub>1.p\<^sub>0 t\<^sub>0u\<^sub>1.p\<^sub>1 \t\<^sub>0\<^sup>* \ u\<^sub>1\ using t\<^sub>0u\<^sub>1.is_arrow_of_tabulations_in_maps by simp interpretation E: self_evaluation_map V H \ \ src trg .. notation E.eval ("\_\") no_notation in_hom ("\_ : _ \ _\") text \ The following lemma states that the rectangular faces of the ``top prism'' commute up to isomorphism. This was not already proved in @{locale composite_tabulation_in_maps}, because there we did not consider any composite structure of the ``source'' 2-cell. There are common elements, though to the proof that the composite of tabulations is a tabulation and the present lemma. The proof idea is to use property \T2\ of the ``base'' tabulations to establish the existence of the desired isomorphisms. The proofs have to be carried out in sequence, starting from the ``output'' side, because the arrow \\\ required in the hypotheses of \T2\ depends, for the ``input'' tabulation, on the isomorphism constructed for the ``output'' tabulation. \ lemma prj_chine: shows "\\.p\<^sub>0 \ chine \ \.chine \ \\.p\<^sub>0" and "\\.p\<^sub>1 \ chine \ \.chine \ \\.p\<^sub>1" proof - have 1: "arrow_of_spans_of_maps V H \ \ src trg (s\<^sub>0 \ \\.p\<^sub>0) (r\<^sub>1 \ \\.p\<^sub>1) (u\<^sub>0 \ \\.p\<^sub>0) (t\<^sub>1 \ \\.p\<^sub>1) chine \ (((t \ u) \ the_\) \ \[t \ u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\\.tab \ chine)) \ the_\ = ((\ \ \) \ s\<^sub>0 \ \\.p\<^sub>0) \ \\.tab" using chine_is_induced_map by simp let ?u\<^sub>\ = "u \ s\<^sub>0 \ \\.p\<^sub>0" let ?w\<^sub>\ = "\.chine \ \\.p\<^sub>1" let ?w\<^sub>\' = "\\.p\<^sub>1 \ chine" have u\<^sub>\: "ide ?u\<^sub>\" using \.u_simps(3) by auto have w\<^sub>\: "ide ?w\<^sub>\ \ is_left_adjoint ?w\<^sub>\" by (simp add: \.is_map \.T0.antipar(1) left_adjoints_compose) have w\<^sub>\': "ide ?w\<^sub>\' \ is_left_adjoint ?w\<^sub>\'" by (simp add: is_map left_adjoints_compose) let ?\\<^sub>\ = "\[u, s\<^sub>0, \\.p\<^sub>0] \ ((\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]" let ?\\<^sub>\' = "(u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]" let ?\\<^sub>\ = "\[t\<^sub>1, \\.p\<^sub>1, chine] \ the_\ \ (inv \.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>1, \.chine, \\.p\<^sub>1]" have \\<^sub>\: "\?\\<^sub>\ : t\<^sub>0 \ ?w\<^sub>\ \ ?u\<^sub>\\" using \.T0.antipar(1) \.the_\_in_hom \.u_simps(3) by (intro comp_in_homI, auto) have \\<^sub>\': "\?\\<^sub>\' : t\<^sub>0 \ ?w\<^sub>\' \ ?u\<^sub>\\" proof (intro comp_in_homI) show "\\\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine] : t\<^sub>0 \ \\.p\<^sub>1 \ chine \ (t\<^sub>0 \ \\.p\<^sub>1) \ chine\" using t\<^sub>0u\<^sub>1.p\<^sub>1_simps assoc'_in_hom by simp show "\t\<^sub>0u\<^sub>1.\ \ chine : (t\<^sub>0 \ \\.p\<^sub>1) \ chine \ (u\<^sub>1 \ \\.p\<^sub>0) \ chine\" using \.T0.antipar(1) by (intro hcomp_in_vhom, auto) show "\(\ \ \\.p\<^sub>0) \ chine : (u\<^sub>1 \ \\.p\<^sub>0) \ chine \ ((u \ u\<^sub>0) \ \\.p\<^sub>0) \ chine\" by (intro hcomp_in_vhom, auto) show "\\[u, u\<^sub>0, \\.p\<^sub>0] \ chine : ((u \ u\<^sub>0) \ \\.p\<^sub>0) \ chine \ (u \ u\<^sub>0 \ \\.p\<^sub>0) \ chine\" using assoc_in_hom by auto show "\\[u, u\<^sub>0 \ \\.p\<^sub>0, chine] : (u \ u\<^sub>0 \ \\.p\<^sub>0) \ chine \ u \ (u\<^sub>0 \ \\.p\<^sub>0) \ chine\" by auto show "\u \ the_\ : u \ (u\<^sub>0 \ \\.p\<^sub>0) \ chine \ ?u\<^sub>\\" by (intro hcomp_in_vhom, auto) qed have \\<^sub>\: "\?\\<^sub>\ : t\<^sub>1 \ ?w\<^sub>\ \ t\<^sub>1 \ ?w\<^sub>\'\" proof (intro comp_in_homI) show "\\\<^sup>-\<^sup>1[t\<^sub>1, \.chine, \\.p\<^sub>1] : t\<^sub>1 \ ?w\<^sub>\ \ (t\<^sub>1 \ \.chine) \ \\.p\<^sub>1\" using \.T0.antipar(1) by auto show "\inv \.the_\ \ \\.p\<^sub>1 : (t\<^sub>1 \ \.chine) \ \\.p\<^sub>1 \ r\<^sub>1 \ \\.p\<^sub>1\" using \.the_\_props \.T0.antipar(1) by (intro hcomp_in_vhom, auto) show "\the_\ : r\<^sub>1 \ \\.p\<^sub>1 \ (t\<^sub>1 \ \\.p\<^sub>1) \ chine\" using the_\_in_hom(2) by simp show "\\[t\<^sub>1, \\.p\<^sub>1, chine] : (t\<^sub>1 \ \\.p\<^sub>1) \ chine \ t\<^sub>1 \ ?w\<^sub>\'\" using t\<^sub>0u\<^sub>1.p\<^sub>1_simps assoc_in_hom by simp qed define LHS where "LHS = (t \ ?\\<^sub>\) \ \[t, t\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)" have LHS: "\LHS : t\<^sub>1 \ ?w\<^sub>\ \ t \ ?u\<^sub>\\" proof (unfold LHS_def, intro comp_in_homI) show "\\ \ ?w\<^sub>\ : t\<^sub>1 \ \.chine \ \\.p\<^sub>1 \ (t \ t\<^sub>0) \ ?w\<^sub>\\" using \.T0.antipar(1) by (intro hcomp_in_vhom, auto) show "\\[t, t\<^sub>0, ?w\<^sub>\] : (t \ t\<^sub>0) \ ?w\<^sub>\ \ t \ t\<^sub>0 \ ?w\<^sub>\\" using \.T0.antipar(1) by auto show "\t \ ?\\<^sub>\ : t \ t\<^sub>0 \ ?w\<^sub>\ \ t \ ?u\<^sub>\\" proof - have "src t = trg (t\<^sub>0 \ \.chine \ r\<^sub>0s\<^sub>1.p\<^sub>1)" by (metis \.u_simps(3) \.ide_base \.ide_leg0 \.leg1_simps(3) \\.composable \\<^sub>\ arrI assoc_simps(3) r\<^sub>0s\<^sub>1.ide_u r\<^sub>0s\<^sub>1.p\<^sub>0_simps trg_vcomp vconn_implies_hpar(2)) thus ?thesis using \\<^sub>\ by blast qed qed define RHS where "RHS = ((t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) \ ?\\<^sub>\" have RHS: "\RHS : t\<^sub>1 \ ?w\<^sub>\ \ t \ ?u\<^sub>\\" unfolding RHS_def proof show "\?\\<^sub>\ : t\<^sub>1 \ ?w\<^sub>\ \ t\<^sub>1 \ ?w\<^sub>\'\" using \\<^sub>\ by simp show "\(t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\') : t\<^sub>1 \ ?w\<^sub>\' \ t \ ?u\<^sub>\\" proof show "\\[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\') : t\<^sub>1 \ ?w\<^sub>\' \ t \ t\<^sub>0 \ ?w\<^sub>\'\" using \.T0.antipar(1) by fastforce show "\t \ ?\\<^sub>\' : t \ t\<^sub>0 \ ?w\<^sub>\' \ t \ ?u\<^sub>\\" using w\<^sub>\' \\<^sub>\' \.leg0_simps(2) \.leg0_simps(3) hseqI' ideD(1) t\<^sub>0u\<^sub>1.p\<^sub>1_simps trg_hcomp \.base_in_hom(2) hcomp_in_vhom by presburger qed qed have eq: "LHS = RHS" proof - have "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ LHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1) = \" proof - text \ Here we use \\.\_naturality\ to replace @{term \.chine} in favor of @{term \}. We have to bring @{term \.the_\}, @{term \}, and @{term \.the_\} together, with @{term \\.p\<^sub>1} on the right. \ have "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ LHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1) = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0] \ ((\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]) \ \[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \ (\ \ \.chine \ \\.p\<^sub>1) \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1)" unfolding LHS_def using comp_assoc by presburger also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ (t \ \.the_\ \ \\.p\<^sub>1) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]) \ \[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \ ((\ \ \.chine \ \\.p\<^sub>1) \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ (\.the_\ \ \\.p\<^sub>1)" proof - have "t \ \[u, s\<^sub>0, \\.p\<^sub>0] \ ((\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] = (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ (t \ \.the_\ \ \\.p\<^sub>1) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1])" using whisker_left \.ide_base \\<^sub>\ arrI seqE by (metis (full_types)) thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ (t \ \.the_\ \ \\.p\<^sub>1) \ ((t \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]) \ \[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \ \[t \ t\<^sub>0, \.chine, \\.p\<^sub>1]) \ ((\ \ \.chine) \ \\.p\<^sub>1) \ (\.the_\ \ \\.p\<^sub>1)" proof - have "(\ \ \.chine \ \\.p\<^sub>1) \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] = \[t \ t\<^sub>0, \.chine, \\.p\<^sub>1] \ ((\ \ \.chine) \ \\.p\<^sub>1)" using assoc_naturality by (metis \.w_simps(2-6) \.leg1_simps(3) \\.leg1_simps(2) \.tab_simps(1) \.tab_simps(2,4-5) hseqE r\<^sub>0s\<^sub>1.leg1_simps(5) r\<^sub>0s\<^sub>1.leg1_simps(6)) thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ ((t \ \.the_\ \ \\.p\<^sub>1) \ \[t, t\<^sub>0 \ \.chine, \\.p\<^sub>1]) \ (\[t, t\<^sub>0, \.chine] \ \\.p\<^sub>1) \ ((\ \ \.chine) \ \\.p\<^sub>1) \ (\.the_\ \ \\.p\<^sub>1)" proof - have "(t \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]) \ \[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \ \[t \ t\<^sub>0, \.chine, \\.p\<^sub>1] = \[t, t\<^sub>0 \ \.chine, \\.p\<^sub>1] \ (\[t, t\<^sub>0, \.chine] \ \\.p\<^sub>1)" proof - have "seq \[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \[t \ t\<^sub>0, \.chine, \\.p\<^sub>1]" by (simp add: \.T0.antipar(1)) moreover have "inv (t \ \[t\<^sub>0, \.chine, \\.p\<^sub>1]) = t \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]" using \.T0.antipar(1) by simp ultimately show ?thesis using pentagon \.T0.antipar(1) iso_hcomp invert_side_of_triangle(1) [of "\[t, t\<^sub>0, \.chine \ \\.p\<^sub>1] \ \[t \ t\<^sub>0, \.chine, \\.p\<^sub>1]" "t \ \[t\<^sub>0, \.chine, \\.p\<^sub>1]" "\[t, t\<^sub>0 \ \.chine, \\.p\<^sub>1] \ (\[t, t\<^sub>0, \.chine] \ \\.p\<^sub>1)"] by simp qed thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ \[t, r\<^sub>0, \\.p\<^sub>1] \ (((t \ \.the_\) \ \\.p\<^sub>1) \ (\[t, t\<^sub>0, \.chine] \ \\.p\<^sub>1) \ ((\ \ \.chine) \ \\.p\<^sub>1)) \ (\.the_\ \ \\.p\<^sub>1)" proof - have "(t \ \.the_\ \ \\.p\<^sub>1) \ \[t, t\<^sub>0 \ \.chine, \\.p\<^sub>1] = \[t, r\<^sub>0, \\.p\<^sub>1] \ ((t \ \.the_\) \ \\.p\<^sub>1)" using assoc_naturality [of t \.the_\ \\.p\<^sub>1] \.\_simps(3) \\.leg1_simps(2) hseq_char by auto thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ \[t, r\<^sub>0, \\.p\<^sub>1] \ ((\ \ r\<^sub>0) \ \ \ \\.p\<^sub>1)" using whisker_right \.T0.antipar(1) \.\_simps(1) \.\_naturality comp_assoc by fastforce also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ ((t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \\.p\<^sub>0)) \ (t \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ \[t, r\<^sub>0, \\.p\<^sub>1] \ ((\ \ r\<^sub>0) \ \ \ \\.p\<^sub>1)" proof - have "t \ (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0 = (t \ (\ \ s\<^sub>0) \ \\.p\<^sub>0) \ (t \ \ \ \\.p\<^sub>0)" using whisker_left whisker_right \.T0.antipar(1) by (metis (full_types) \.\_simps(1) \.ide_base \\<^sub>\ arrI r\<^sub>0s\<^sub>1.ide_u seqE) thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (t \ \[s, s\<^sub>0, \\.p\<^sub>0]) \ (t \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ \[t, r\<^sub>0, \\.p\<^sub>1] \ ((\ \ r\<^sub>0) \ \ \ \\.p\<^sub>1)" proof - have "(t \ \[u, s\<^sub>0, \\.p\<^sub>0]) \ (t \ (\ \ s\<^sub>0) \ \\.p\<^sub>0) = t \ \[u, s\<^sub>0, \\.p\<^sub>0] \ ((\ \ s\<^sub>0) \ \\.p\<^sub>0)" using \.in_hom whisker_left by auto also have "... = t \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0]" using assoc_naturality [of \ s\<^sub>0 \\.p\<^sub>0] \.in_hom by auto also have "... = (t \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (t \ \[s, s\<^sub>0, \\.p\<^sub>0])" proof - have "seq (\ \ s\<^sub>0 \ \\.p\<^sub>0) \[s, s\<^sub>0, \\.p\<^sub>0]" using \.in_hom apply (intro seqI hseqI) apply auto proof - show "\\ : src u \ trg \\" by (metis \.\_simps(1) \.u_simps(3) hseqE in_hhom_def seqE) show "dom (\ \ s\<^sub>0 \ \\.p\<^sub>0) = s \ s\<^sub>0 \ \\.p\<^sub>0" by (metis \_simps(1) \.in_hom hcomp_simps(1,3) hseq_char in_homE seqE u_simps(4)) qed thus ?thesis using whisker_left by simp qed finally show ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (t \ \[s, s\<^sub>0, \\.p\<^sub>0]) \ (t \ \ \ \\.p\<^sub>0) \ (t \ r\<^sub>0s\<^sub>1.\) \ (\[t, r\<^sub>0, \\.p\<^sub>1] \ ((\ \ r\<^sub>0) \ \\.p\<^sub>1)) \ (\ \ \\.p\<^sub>1)" using whisker_right comp_assoc by simp also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (t \ \[s, s\<^sub>0, \\.p\<^sub>0]) \ (t \ \ \ \\.p\<^sub>0) \ ((t \ r\<^sub>0s\<^sub>1.\) \ (\ \ r\<^sub>0 \ \\.p\<^sub>1)) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" using assoc_naturality [of \ r\<^sub>0 \\.p\<^sub>1] \.in_hom \.T0.antipar(1) comp_assoc by fastforce also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ ((t \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (t \ \[s, s\<^sub>0, \\.p\<^sub>0]) \ (t \ \ \ \\.p\<^sub>0)) \ (\ \ s\<^sub>1 \ \\.p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" proof - have "(t \ r\<^sub>0s\<^sub>1.\) \ (\ \ r\<^sub>0 \ \\.p\<^sub>1) = \ \ r\<^sub>0s\<^sub>1.\" using comp_cod_arr comp_arr_dom \.in_hom interchange comp_ide_arr by (metis \.base_in_hom(2) \.ide_base r\<^sub>0s\<^sub>1.\_simps(1) r\<^sub>0s\<^sub>1.\_simps(4) seqI') also have "... = (\ \ s\<^sub>1 \ \\.p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\)" using r\<^sub>0s\<^sub>1.\_in_hom comp_cod_arr comp_arr_dom \.in_hom interchange by (metis in_homE) finally have "(t \ r\<^sub>0s\<^sub>1.\) \ (\ \ r\<^sub>0 \ \\.p\<^sub>1) = (\ \ s\<^sub>1 \ \\.p\<^sub>0) \ (r \ r\<^sub>0s\<^sub>1.\)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ ((t \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)) \ (\ \ s\<^sub>1 \ \\.p\<^sub>0)) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" using whisker_left \.T0.antipar(1) \\.composable \.in_hom comp_assoc by auto also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (\ \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" proof - have "(t \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)) \ (\ \ s\<^sub>1 \ \\.p\<^sub>0) = \ \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)" proof - have "\(\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0) : s\<^sub>1 \ \\.p\<^sub>0 \ u \ s\<^sub>0 \ \\.p\<^sub>0\" using \.in_hom \.in_hom by force thus ?thesis by (metis (no_types) \.in_hom comp_arr_dom comp_cod_arr in_homE interchange) qed thus ?thesis using comp_assoc by presburger qed also have "... = (\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (\ \ \ \ s\<^sub>0 \ \\.p\<^sub>0)) \ (r \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" proof - have "\ \ (\ \ s\<^sub>0 \ \\.p\<^sub>0) \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0) = (\ \ \ \ s\<^sub>0 \ \\.p\<^sub>0) \ (r \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0))" proof - have "seq (\ \ s\<^sub>0 \ \\.p\<^sub>0) (\[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0))" using \.in_hom by force thus ?thesis using comp_arr_dom comp_cod_arr \.in_hom \.in_hom interchange by (metis in_homE) qed thus ?thesis using comp_assoc by presburger qed also have "... = ((\ \ \) \ s\<^sub>0 \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ \\.p\<^sub>0] \ (r \ \[s, s\<^sub>0, \\.p\<^sub>0] \ (\ \ \\.p\<^sub>0)) \ (r \ r\<^sub>0s\<^sub>1.\) \ \[r, r\<^sub>0, \\.p\<^sub>1] \ (\ \ \\.p\<^sub>1)" proof - have "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (\ \ \ \ s\<^sub>0 \ \\.p\<^sub>0) = ((\ \ \) \ s\<^sub>0 \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[r, s, s\<^sub>0 \ \\.p\<^sub>0]" using assoc_naturality \.in_hom \.in_hom by (metis \\.leg0_simps(3) assoc'_naturality hcomp_in_vhomE in_hom in_homE u_simps(2) u_simps(4) u_simps(5)) thus ?thesis using comp_assoc by presburger qed also have "... = \" using whisker_left \\.tab_def comp_assoc by simp finally show ?thesis by auto qed also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ RHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1)" proof - text \Now cancel @{term \.the_\} and its inverse.\ have "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ RHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1) = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ (u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ (\ \ \\.p\<^sub>1 \ chine) \ \[t\<^sub>1, \\.p\<^sub>1, chine] \ the_\ \ (inv \.the_\ \ \\.p\<^sub>1) \ ((\\<^sup>-\<^sup>1[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ (\.the_\ \ \\.p\<^sub>1)" unfolding RHS_def using comp_assoc by presburger also have "... = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ (u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ (\ \ \\.p\<^sub>1 \ chine) \ \[t\<^sub>1, \\.p\<^sub>1, chine] \ the_\" proof - have "the_\ \ (inv \.the_\ \ \\.p\<^sub>1) \ ((\\<^sup>-\<^sup>1[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ (\.the_\ \ \\.p\<^sub>1) = the_\ \ (inv \.the_\ \ \\.p\<^sub>1) \ ((t\<^sub>1 \ \.chine) \ \\.p\<^sub>1) \ (\.the_\ \ \\.p\<^sub>1)" using comp_inv_arr \.T0.antipar(1) comp_assoc_assoc' by simp also have "... = the_\ \ (inv \.the_\ \ \\.p\<^sub>1) \ (\.the_\ \ \\.p\<^sub>1)" using comp_cod_arr \.T0.antipar(1) by simp also have "... = the_\ \ (r\<^sub>1 \ \\.p\<^sub>1)" using whisker_right [of \\.p\<^sub>1] r\<^sub>0s\<^sub>1.ide_leg1 \.the_\_props(2) \.the_\_simps(4) \.leg1_simps(2) comp_inv_arr' by metis also have "... = the_\" using comp_arr_dom by simp finally show ?thesis using comp_assoc by simp qed text \ Now reassociate to move @{term the_\} to the left and get other terms composed with @{term chine}, where they can be reduced to @{term \\.tab}. \ also have "... = (\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ (t \ u \ the_\)) \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ (t \ (\ \ \\.p\<^sub>0) \ chine) \ (t \ t\<^sub>0u\<^sub>1.\ \ chine) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ (\ \ \\.p\<^sub>1 \ chine) \ \[t\<^sub>1, \\.p\<^sub>1, chine] \ the_\" proof - have "arr ((u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using \\<^sub>\' by blast moreover have "arr (\[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using calculation by blast moreover have "arr ((\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using calculation by blast moreover have "arr (((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using calculation by blast moreover have "arr ((t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using calculation by blast ultimately have "t \ (u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ (t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine] = (t \ u \ the_\) \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ (t \ (\ \ \\.p\<^sub>0) \ chine) \ (t \ t\<^sub>0u\<^sub>1.\ \ chine) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine])" using whisker_left \.T0.antipar(1) \.ide_base by presburger thus ?thesis using comp_assoc by presburger qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ (t \ (\ \ \\.p\<^sub>0) \ chine) \ (t \ t\<^sub>0u\<^sub>1.\ \ chine) \ (t \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ ((\ \ \\.p\<^sub>1 \ chine) \ \[t\<^sub>1, \\.p\<^sub>1, chine]) \ the_\" using assoc'_naturality [of t u the_\] \\.composable \_simps(3) comp_assoc by auto also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ (t \ (\ \ \\.p\<^sub>0) \ chine) \ (t \ t\<^sub>0u\<^sub>1.\ \ chine) \ ((t \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ \[t \ t\<^sub>0, \\.p\<^sub>1, chine]) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "(\ \ \\.p\<^sub>1 \ chine) \ \[t\<^sub>1, \\.p\<^sub>1, chine] = \[t \ t\<^sub>0, \\.p\<^sub>1, chine] \ ((\ \ \\.p\<^sub>1) \ chine)" using assoc_naturality by (metis \.leg1_simps(3) \.tab_simps(1,2,4) \.tab_simps(5) \\.leg0_simps(2) \\.leg1_simps(2) hseqE src_hcomp t\<^sub>0u\<^sub>1.leg1_simps(3,5-6) w_simps(2) w_simps(4-6)) thus ?thesis using comp_assoc by presburger qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ (t \ (\ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\ \ chine) \ \[t, t\<^sub>0 \ \\.p\<^sub>1, chine]) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "(t \ \\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ \[t \ t\<^sub>0, \\.p\<^sub>1, chine] = \[t, t\<^sub>0 \ \\.p\<^sub>1, chine] \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine)" using pentagon t\<^sub>0u\<^sub>1.p\<^sub>1_simps uw\ \.T0.antipar(1) iso_hcomp comp_assoc_assoc' invert_side_of_triangle(1) [of "\[t, t\<^sub>0, \\.p\<^sub>1 \ chine] \ \[t \ t\<^sub>0, \\.p\<^sub>1, chine]" "t \ \[t\<^sub>0, \\.p\<^sub>1, chine]" "\[t, t\<^sub>0 \ \\.p\<^sub>1, chine] \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine)"] by auto thus ?thesis using comp_assoc by presburger qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((t \ (\ \ \\.p\<^sub>0) \ chine) \ \[t, u\<^sub>1 \ \\.p\<^sub>0, chine]) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "(t \ t\<^sub>0u\<^sub>1.\ \ chine) \ \[t, t\<^sub>0 \ \\.p\<^sub>1, chine] = \[t, u\<^sub>1 \ \\.p\<^sub>0, chine] \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine)" using assoc_naturality [of t t\<^sub>0u\<^sub>1.\ chine] t\<^sub>0u\<^sub>1.cospan by auto thus ?thesis using comp_assoc by presburger qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "(t \ (\ \ \\.p\<^sub>0) \ chine) \ \[t, u\<^sub>1 \ \\.p\<^sub>0, chine] = \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ (\ \ \\.p\<^sub>0)) \ chine)" using assoc_naturality [of t "\ \ \\.p\<^sub>0" chine] by (simp add: \\.composable) thus ?thesis using comp_assoc by presburger qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "(((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine)) \ ((t \ \ \ \\.p\<^sub>0) \ chine) = ((t \ ((u \ u\<^sub>0) \ \\.p\<^sub>0)) \ chine) \ ((t \ \ \ \\.p\<^sub>0) \ chine)" using whisker_right whisker_left [of t "\\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]" "\[u, u\<^sub>0, \\.p\<^sub>0]"] \\.composable comp_assoc_assoc' by simp also have "... = (t \ \ \ \\.p\<^sub>0) \ chine" using comp_cod_arr \\.composable by simp finally have "(((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine)) \ ((t \ \ \ \\.p\<^sub>0) \ chine) = (t \ \ \ \\.p\<^sub>0) \ chine" by simp thus ?thesis using comp_assoc by metis qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ (((\[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) \ (\\<^sup>-\<^sup>1[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine)) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine)) \ ((t \ \ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine) \ the_\" proof - have "((\[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) \ (\\<^sup>-\<^sup>1[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine)) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) = ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine)" using comp_inv_arr' comp_cod_arr \\.composable comp_assoc_assoc' whisker_right [of chine "\[t, u, u\<^sub>0 \ \\.p\<^sub>0]" "\\<^sup>-\<^sup>1[t, u, u\<^sub>0 \ \\.p\<^sub>0]"] by simp thus ?thesis using comp_assoc by simp qed also have "... = ((t \ u) \ the_\) \ \\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ (\[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) \ ((\\<^sup>-\<^sup>1[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine)) \ the_\" using comp_assoc by presburger also have "... = ((t \ u) \ the_\) \ (\\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ (\[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine)) \ (\\.tab \ chine) \ the_\" proof - have "(\\<^sup>-\<^sup>1[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) \ ((t \ \[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ ((t \ \ \ \\.p\<^sub>0) \ chine) \ ((t \ t\<^sub>0u\<^sub>1.\) \ chine) \ (\[t, t\<^sub>0, \\.p\<^sub>1] \ chine) \ ((\ \ \\.p\<^sub>1) \ chine) = \\.tab \ chine" using uw\ whisker_right [of chine] by (metis \\.tab_def \\.tab_in_vhom' arrI seqE) thus ?thesis using comp_assoc by presburger qed also have "... = ((t \ u) \ the_\) \ \[t \ u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\\.tab \ chine) \ the_\" proof - have "\\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \ \\.p\<^sub>0) \ chine] \ (t \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine]) \ (t \ \[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \[t, (u \ u\<^sub>0) \ \\.p\<^sub>0, chine] \ ((t \ \\<^sup>-\<^sup>1[u, u\<^sub>0, \\.p\<^sub>0]) \ chine) \ (\[t, u, u\<^sub>0 \ \\.p\<^sub>0] \ chine) = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\t\<^bold>\, \<^bold>\u\<^bold>\, (\<^bold>\u\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\\\.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\u\<^bold>\, \<^bold>\u\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\\\.p\<^sub>0\<^bold>\, \<^bold>\chine\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\u\<^bold>\, \<^bold>\u\<^sub>0\<^bold>\, \<^bold>\\\.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\chine\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\t\<^bold>\, (\<^bold>\u\<^bold>\ \<^bold>\ \<^bold>\u\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\\\.p\<^sub>0\<^bold>\, \<^bold>\chine\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\u\<^bold>\, \<^bold>\u\<^sub>0\<^bold>\, \<^bold>\\\.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\chine\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\t\<^bold>\, \<^bold>\u\<^bold>\, \<^bold>\u\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\\\.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\chine\<^bold>\)\" using \'_def \_def \\.composable by simp also have "... = \\<^bold>\\<^bold>[\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\, \<^bold>\u\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\\\.p\<^sub>0\<^bold>\, \<^bold>\chine\<^bold>\\<^bold>]\" using \\.composable apply (intro E.eval_eqI) by simp_all also have "... = \[t \ u, u\<^sub>0 \ \\.p\<^sub>0, chine]" using \'_def \_def \\.composable by simp finally show ?thesis by simp qed also have "... = \" using \_naturality by simp finally show ?thesis by simp qed finally have "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ LHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1) = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0] \ RHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] \ (\.the_\ \ \\.p\<^sub>1)" by blast (* * TODO: This is common enough that there should be "cancel_iso_left" and * "cancel_iso_right" rules for doing it. *) hence "(LHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ (\.the_\ \ \\.p\<^sub>1) = (RHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]) \ (\.the_\ \ \\.p\<^sub>1)" using u\<^sub>\ r\<^sub>0s\<^sub>1.ide_u LHS RHS iso_is_section [of "\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ \\.p\<^sub>0]"] section_is_mono monoE \\.composable comp_assoc by (metis (no_types, lifting) \_simps(1) \.ide_base \\\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ r\<^sub>0s\<^sub>1.p\<^sub>0] \ LHS \ \[t\<^sub>1, \.chine, r\<^sub>0s\<^sub>1.p\<^sub>1] \ (\.the_\ \ r\<^sub>0s\<^sub>1.p\<^sub>1) = ((\ \ \) \ s\<^sub>0 \ r\<^sub>0s\<^sub>1.p\<^sub>0) \ \\.tab\ \.ide_base hseq_char ideD(1) ide_u iso_assoc') hence 1: "LHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1] = RHS \ \[t\<^sub>1, \.chine, \\.p\<^sub>1]" using epiE LHS RHS iso_is_retraction retraction_is_epi \\.composable \.the_\_props iso_hcomp by (metis \_simps(1) \.the_\_simps(2) \((\ \ \) \ s\<^sub>0 \ r\<^sub>0s\<^sub>1.p\<^sub>0) \ \\.tab = \\<^sup>-\<^sup>1[t, u, s\<^sub>0 \ r\<^sub>0s\<^sub>1.p\<^sub>0] \ RHS \ \[t\<^sub>1, \.chine, r\<^sub>0s\<^sub>1.p\<^sub>1] \ (\.the_\ \ r\<^sub>0s\<^sub>1.p\<^sub>1)\ \.leg1_simps(3) ide_is_iso local.comp_assoc r\<^sub>0s\<^sub>1.ide_leg1 r\<^sub>0s\<^sub>1.p\<^sub>1_simps seqE) show "LHS = RHS" proof - have "epi \[t\<^sub>1, \.chine, \\.p\<^sub>1]" using iso_is_retraction retraction_is_epi \.T0.antipar(1) by simp moreover have "seq LHS \[t\<^sub>1, \.chine, \\.p\<^sub>1]" using LHS \.T0.antipar(1) by auto moreover have "seq RHS \[t\<^sub>1, \.chine, \\.p\<^sub>1]" using RHS \.T0.antipar(1) by auto ultimately show ?thesis using epiE 1 by blast qed qed have 1: "\!\. \\ : ?w\<^sub>\ \ ?w\<^sub>\'\ \ ?\\<^sub>\ = t\<^sub>1 \ \ \ ?\\<^sub>\ = ?\\<^sub>\' \ (t\<^sub>0 \ \)" using LHS_def RHS_def u\<^sub>\ w\<^sub>\ w\<^sub>\' \\<^sub>\ \\<^sub>\ \\<^sub>\' eq \.T2 [of ?w\<^sub>\ ?w\<^sub>\' ?\\<^sub>\ ?u\<^sub>\ ?\\<^sub>\' ?\\<^sub>\] by fastforce obtain \\<^sub>\ where \\<^sub>\: "\\\<^sub>\ : ?w\<^sub>\ \ ?w\<^sub>\'\ \ ?\\<^sub>\ = t\<^sub>1 \ \\<^sub>\ \ ?\\<^sub>\ = ?\\<^sub>\' \ (t\<^sub>0 \ \\<^sub>\)" using 1 by auto text \ At this point we could show that @{term \\<^sub>\} is invertible using \BS3\, but we want to avoid using \BS3\ if possible and we also want to establish a characterization of @{term "inv \\<^sub>\"}. So we show the invertibility of @{term \\<^sub>\} directly, using a few more applications of \T2\. \ have iso_\\<^sub>\: "iso ?\\<^sub>\" using uw\ \\<^sub>\ the_\_props \.the_\_props hseqI' iso_assoc' \.hseq_leg\<^sub>0 apply (intro isos_compose) apply (metis \.is_ide \\.leg1_simps(2) \.ide_leg1 \.leg1_simps(2) \.leg1_simps(3) hseqE r\<^sub>0s\<^sub>1.ide_leg1 hcomp_simps(1) vconn_implies_hpar(3)) apply (metis \\.leg1_simps(2) hseqE ide_is_iso r\<^sub>0s\<^sub>1.ide_leg1 src_inv iso_inv_iso iso_hcomp vconn_implies_hpar(1)) apply blast apply blast apply blast apply (metis \.ide_leg1 \.leg1_simps(3) hseqE ide_char iso_assoc t\<^sub>0u\<^sub>1.ide_leg1 t\<^sub>0u\<^sub>1.p\<^sub>1_simps w\<^sub>\') by blast hence eq': "((t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) = ((t \ ?\\<^sub>\) \ \[t, t\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)) \ inv ?\\<^sub>\" proof - have "seq ((t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) ?\\<^sub>\" using LHS RHS_def eq by blast hence "(t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\') = (((t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) \ ?\\<^sub>\) \ inv ?\\<^sub>\" by (meson invert_side_of_triangle(2) iso_\\<^sub>\) thus ?thesis using LHS_def RHS_def eq by argo qed have 2: "\!\. \\ : ?w\<^sub>\' \ ?w\<^sub>\\ \ inv ?\\<^sub>\ = t\<^sub>1 \ \ \ ?\\<^sub>\' = ?\\<^sub>\ \ (t\<^sub>0 \ \)" using u\<^sub>\ w\<^sub>\ w\<^sub>\' \\<^sub>\ \\<^sub>\ \\<^sub>\' eq' \.T2 [of ?w\<^sub>\' ?w\<^sub>\ ?\\<^sub>\'?u\<^sub>\ ?\\<^sub>\ "inv ?\\<^sub>\"] iso_\\<^sub>\ comp_assoc by blast obtain \\<^sub>\' where \\<^sub>\': "\\\<^sub>\' : ?w\<^sub>\' \ ?w\<^sub>\\ \ inv ?\\<^sub>\ = t\<^sub>1 \ \\<^sub>\' \ ?\\<^sub>\' = ?\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\')" using 2 by auto have "inverse_arrows \\<^sub>\ \\<^sub>\'" proof have "\\\<^sub>\' \ \\<^sub>\ : ?w\<^sub>\ \ ?w\<^sub>\\" using \\<^sub>\ \\<^sub>\' by auto moreover have "t\<^sub>1 \ \\<^sub>\' \ \\<^sub>\ = t\<^sub>1 \ ?w\<^sub>\" using \\<^sub>\ \\<^sub>\' whisker_left \\<^sub>\ iso_\\<^sub>\ comp_inv_arr' by (metis (no_types, lifting) \.ide_leg1 calculation in_homE) moreover have "?\\<^sub>\ = ?\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\' \ \\<^sub>\)" proof - have "?\\<^sub>\ = ?\\<^sub>\' \ (t\<^sub>0 \ \\<^sub>\)" using \\<^sub>\ by simp also have "... = ?\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\') \ (t\<^sub>0 \ \\<^sub>\)" using \\<^sub>\' comp_assoc by simp also have "... = ?\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\' \ \\<^sub>\)" using \\<^sub>\ \\<^sub>\' whisker_left by (metis (full_types) \.ide_leg0 seqI') finally show ?thesis by simp qed moreover have "\\. \\ : ?w\<^sub>\ \ ?w\<^sub>\\ \ t\<^sub>1 \ \ = t\<^sub>1 \ ?w\<^sub>\ \ ?\\<^sub>\ = ?\\<^sub>\ \ (t\<^sub>0 \ \) \ \ = ?w\<^sub>\" proof - have "\!\. \\ : ?w\<^sub>\ \ ?w\<^sub>\\ \ t\<^sub>1 \ ?w\<^sub>\ = t\<^sub>1 \ \ \ ?\\<^sub>\ = ?\\<^sub>\ \ (t\<^sub>0 \ \)" proof - have "(t \ ?\\<^sub>\) \ \[t, t\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\) = ((t \ ?\\<^sub>\) \ \[t, t\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)) \ (t\<^sub>1 \ ?w\<^sub>\)" by (metis LHS LHS_def comp_arr_dom in_homE) thus ?thesis using w\<^sub>\ \\<^sub>\ \.w_simps(4) \.leg1_in_hom(2) \.leg1_simps(3) hcomp_in_vhom ideD(1) trg_hcomp ide_in_hom(2) \.T2 by presburger qed thus "\\. \\ : ?w\<^sub>\ \ ?w\<^sub>\\ \ t\<^sub>1 \ \ = t\<^sub>1 \ ?w\<^sub>\ \ ?\\<^sub>\ = ?\\<^sub>\ \ (t\<^sub>0 \ \) \ \ = ?w\<^sub>\" by (metis \\<^sub>\ comp_arr_dom ide_in_hom(2) in_homE w\<^sub>\) qed ultimately have "\\<^sub>\' \ \\<^sub>\ = ?w\<^sub>\" by simp thus "ide (\\<^sub>\' \ \\<^sub>\)" using w\<^sub>\ by simp have "\\\<^sub>\ \ \\<^sub>\' : ?w\<^sub>\' \ ?w\<^sub>\'\" using \\<^sub>\ \\<^sub>\' by auto moreover have "t\<^sub>1 \ \\<^sub>\ \ \\<^sub>\' = t\<^sub>1 \ ?w\<^sub>\'" by (metis \\<^sub>\ \\<^sub>\ \\<^sub>\' \.ide_leg1 calculation comp_arr_inv' in_homE iso_\\<^sub>\ whisker_left) moreover have "?\\<^sub>\' = ?\\<^sub>\' \ (t\<^sub>0 \ \\<^sub>\ \ \\<^sub>\')" proof - have "?\\<^sub>\' = ?\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\')" using \\<^sub>\' by simp also have "... = ?\\<^sub>\' \ (t\<^sub>0 \ \\<^sub>\) \ (t\<^sub>0 \ \\<^sub>\')" using \\<^sub>\ comp_assoc by simp also have "... = ?\\<^sub>\' \ (t\<^sub>0 \ \\<^sub>\ \ \\<^sub>\')" using \\<^sub>\ \\<^sub>\' whisker_left \.ide_leg0 by fastforce finally show ?thesis by simp qed moreover have "\\. \\ : ?w\<^sub>\' \ ?w\<^sub>\'\ \ t\<^sub>1 \ \ = t\<^sub>1 \ ?w\<^sub>\' \ ?\\<^sub>\' = ?\\<^sub>\' \ (t\<^sub>0 \ \) \ \ = ?w\<^sub>\'" proof - have "\!\. \\ : ?w\<^sub>\' \ ?w\<^sub>\'\ \ t\<^sub>1 \ ?w\<^sub>\' = t\<^sub>1 \ \ \ ?\\<^sub>\' = ?\\<^sub>\' \ (t\<^sub>0 \ \)" proof - have "(t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\') = ((t \ ?\\<^sub>\') \ \[t, t\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) \ (t\<^sub>1 \ ?w\<^sub>\')" proof - have 1: "t\<^sub>1 \ \\<^sub>\ \ \\<^sub>\' = (t\<^sub>1 \ \\<^sub>\) \ (t\<^sub>1 \ \\<^sub>\')" by (meson \\<^sub>\ \\<^sub>\' \.ide_leg1 seqI' whisker_left) have "((LHS \ inv ?\\<^sub>\) \ (t\<^sub>1 \ \\<^sub>\)) \ (t\<^sub>1 \ \\<^sub>\') = LHS \ inv ?\\<^sub>\" using LHS_def RHS_def \\<^sub>\ \\<^sub>\' eq eq' by argo thus ?thesis unfolding LHS_def using 1 by (simp add: calculation(2) eq' comp_assoc) qed thus ?thesis using w\<^sub>\' \\<^sub>\' \.w_simps(4) \.leg1_in_hom(2) \.leg1_simps(3) hcomp_in_vhom ideD(1) trg_hcomp ide_in_hom(2) \.T2 \.T0.antipar(1) t\<^sub>0u\<^sub>1.base_simps(2) t\<^sub>0u\<^sub>1.leg1_simps(4) by presburger qed thus "\\. \\ : ?w\<^sub>\' \ ?w\<^sub>\'\ \ t\<^sub>1 \ \ = t\<^sub>1 \ ?w\<^sub>\' \ ?\\<^sub>\' = ?\\<^sub>\' \ (t\<^sub>0 \ \) \ \ = ?w\<^sub>\'" by (metis \\<^sub>\' comp_arr_dom ide_in_hom(2) in_homE w\<^sub>\') qed ultimately have "\\<^sub>\ \ \\<^sub>\' = ?w\<^sub>\'" by simp thus "ide (\\<^sub>\ \ \\<^sub>\')" using w\<^sub>\' by simp qed thus "\\.p\<^sub>1 \ chine \ \.chine \ \\.p\<^sub>1" using w\<^sub>\ w\<^sub>\' \\<^sub>\ isomorphic_symmetric isomorphic_def by blast have iso_\\<^sub>\: "iso \\<^sub>\" using \inverse_arrows \\<^sub>\ \\<^sub>\'\ by auto have \\<^sub>\'_eq: "\\<^sub>\' = inv \\<^sub>\" using \inverse_arrows \\<^sub>\ \\<^sub>\'\ inverse_unique by blast let ?w\<^sub>\ = "\\.p\<^sub>0 \ chine" let ?w\<^sub>\' = "\.chine \ \\.p\<^sub>0" let ?u\<^sub>\ = "s\<^sub>0 \ \\.p\<^sub>0" let ?\\<^sub>\ = "the_\ \ \\<^sup>-\<^sup>1[u\<^sub>0, \\.p\<^sub>0, chine]" let ?\\<^sub>\' = "(\.the_\ \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0]" let ?\\<^sub>\ = "\[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" have w\<^sub>\: "ide ?w\<^sub>\ \ is_left_adjoint ?w\<^sub>\" using is_map left_adjoints_compose by simp have w\<^sub>\': "ide ?w\<^sub>\' \ is_left_adjoint ?w\<^sub>\'" using \.is_map left_adjoints_compose by (simp add: is_map left_adjoints_compose) have 1: "\!\. \\ : ?w\<^sub>\ \ ?w\<^sub>\'\ \ ?\\<^sub>\ = u\<^sub>1 \ \ \ ?\\<^sub>\ = ?\\<^sub>\' \ (u\<^sub>0 \ \)" proof - have \\<^sub>\: "\?\\<^sub>\ : u\<^sub>0 \ ?w\<^sub>\ \ ?u\<^sub>\\" by auto have \\<^sub>\': "\?\\<^sub>\' : u\<^sub>0 \ ?w\<^sub>\' \ ?u\<^sub>\\" by fastforce have \\<^sub>\: "\?\\<^sub>\ : u\<^sub>1 \ ?w\<^sub>\ \ u\<^sub>1 \ ?w\<^sub>\'\" proof (intro comp_in_homI) show "\\\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine] : u\<^sub>1 \ \\.p\<^sub>0 \ chine \ (u\<^sub>1 \ \\.p\<^sub>0) \ chine\" by auto show "\inv t\<^sub>0u\<^sub>1.\ \ chine : (u\<^sub>1 \ \\.p\<^sub>0) \ chine \ (t\<^sub>0 \ \\.p\<^sub>1) \ chine\" using t\<^sub>0u\<^sub>1.\_uniqueness(2) hcomp_in_vhom by (simp add: t\<^sub>0u\<^sub>1.\_in_hom(2) w_in_hom(2)) show "\\[t\<^sub>0, \\.p\<^sub>1, chine] : (t\<^sub>0 \ \\.p\<^sub>1) \ chine \ t\<^sub>0 \ \\.p\<^sub>1 \ chine\" using \.T0.antipar(1) by auto show "\t\<^sub>0 \ inv \\<^sub>\ : t\<^sub>0 \ \\.p\<^sub>1 \ chine \ t\<^sub>0 \ \.chine \ \\.p\<^sub>1\" using \\<^sub>\ iso_\\<^sub>\ using \.T0.antipar(1) by auto show "\\\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] : t\<^sub>0 \ \.chine \ \\.p\<^sub>1 \ (t\<^sub>0 \ \.chine) \ \\.p\<^sub>1\" using \.T0.antipar(1) by auto show "\\.the_\ \ \\.p\<^sub>1 : (t\<^sub>0 \ \.chine) \ \\.p\<^sub>1 \ r\<^sub>0 \ \\.p\<^sub>1\" using \.T0.antipar(1) by auto show "\r\<^sub>0s\<^sub>1.\ : r\<^sub>0 \ \\.p\<^sub>1 \ s\<^sub>1 \ \\.p\<^sub>0\" by auto show "\\.the_\ \ \\.p\<^sub>0 : s\<^sub>1 \ \\.p\<^sub>0 \ (u\<^sub>1 \ \.chine) \ \\.p\<^sub>0\" by auto show "\\[u\<^sub>1, \.chine, \\.p\<^sub>0] : (u\<^sub>1 \ \.chine) \ \\.p\<^sub>0 \ u\<^sub>1 \ \.chine \ \\.p\<^sub>0\" by auto qed text \ The proof of the equation below needs to make use of the equation \\\<^sub>\' = \\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\')\ from the previous section. So the overall strategy is to work toward an expression of the form \\\<^sub>\ \ (t\<^sub>0 \ \\<^sub>\')\ and perform the replacement to eliminate \\\<^sub>\'\. \ have eq\<^sub>\: "(u \ ?\\<^sub>\) \ \[u, u\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\) = ((u \ ?\\<^sub>\') \ \[u, u\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) \ ?\\<^sub>\" proof - let ?LHS = "(u \ ?\\<^sub>\) \ \[u, u\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)" let ?RHS = "((u \ ?\\<^sub>\') \ \[u, u\<^sub>0, ?w\<^sub>\'] \ (\ \ ?w\<^sub>\')) \ ?\\<^sub>\" have "?RHS = (u \ (\.the_\ \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0]) \ \[u, u\<^sub>0, \.chine \ \\.p\<^sub>0] \ (\ \ \.chine \ \\.p\<^sub>0) \ \[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_assoc by presburger also have "... = (u \ \.the_\ \ \\.p\<^sub>0) \ ((u \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0]) \ \[u, u\<^sub>0, \.chine \ \\.p\<^sub>0]) \ (\ \ \.chine \ \\.p\<^sub>0) \ \[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "u \ (\.the_\ \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0] = (u \ \.the_\ \ \\.p\<^sub>0) \ (u \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0])" using whisker_left \.ide_base \\<^sub>\' by blast thus ?thesis using comp_assoc by presburger qed also have "... = ((u \ \.the_\ \ \\.p\<^sub>0) \ \[u, u\<^sub>0 \ \.chine, \\.p\<^sub>0]) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u \ u\<^sub>0, \.chine, \\.p\<^sub>0] \ (\ \ \.chine \ \\.p\<^sub>0) \ \[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "seq (u \ \[u\<^sub>0, \.chine, \\.p\<^sub>0]) (\[u, u\<^sub>0 \ \.chine, \\.p\<^sub>0] \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0))" by auto moreover have "src u = trg \[u\<^sub>0, \.chine, \\.p\<^sub>0]" by simp moreover have "inv (u \ \[u\<^sub>0, \.chine, \\.p\<^sub>0]) = u \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0]" by simp moreover have "iso (u \ \[u\<^sub>0, \.chine, \\.p\<^sub>0])" by simp moreover have "iso \[u \ u\<^sub>0, \.chine, \\.p\<^sub>0]" by simp ultimately have "(u \ \\<^sup>-\<^sup>1[u\<^sub>0, \.chine, \\.p\<^sub>0]) \ \[u, u\<^sub>0, \.chine \ \\.p\<^sub>0] = \[u, u\<^sub>0 \ \.chine, \\.p\<^sub>0] \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u \ u\<^sub>0, \.chine, \\.p\<^sub>0]" using pentagon hseqI' comp_assoc invert_opposite_sides_of_square [of "u \ \[u\<^sub>0, \.chine, \\.p\<^sub>0]" "\[u, u\<^sub>0 \ \.chine, \\.p\<^sub>0] \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0)" "\[u, u\<^sub>0, \.chine \ \\.p\<^sub>0]" "\[u \ u\<^sub>0, \.chine, \\.p\<^sub>0]"] inv_hcomp \.is_ide \.w_simps(3) \.w_simps(4) \.base_simps(2) \.ide_base \.ide_leg0 \.leg0_simps(2) \.leg0_simps(3) \.leg1_simps(3) assoc'_eq_inv_assoc ide_hcomp r\<^sub>0s\<^sub>1.ide_u r\<^sub>0s\<^sub>1.p\<^sub>0_simps hcomp_simps(1) by presburger thus ?thesis using comp_assoc by presburger qed also have "... = \[u, s\<^sub>0, \\.p\<^sub>0] \ ((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ (\\<^sup>-\<^sup>1[u \ u\<^sub>0, \.chine, \\.p\<^sub>0] \ (\ \ \.chine \ \\.p\<^sub>0)) \ \[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "(u \ \.the_\ \ \\.p\<^sub>0) \ \[u, u\<^sub>0 \ \.chine, \\.p\<^sub>0] = \[u, s\<^sub>0, \\.p\<^sub>0] \ ((u \ \.the_\) \ \\.p\<^sub>0)" using assoc_naturality [of u \.the_\ \\.p\<^sub>0] \.\_simps(3) by auto thus ?thesis using comp_assoc by presburger qed also have "... = \[u, s\<^sub>0, \\.p\<^sub>0] \ ((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ ((\ \ \.chine) \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u\<^sub>1, \.chine, \\.p\<^sub>0] \ \[u\<^sub>1, \.chine, \\.p\<^sub>0] \ (\.the_\ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "\\<^sup>-\<^sup>1[u \ u\<^sub>0, \.chine, \\.p\<^sub>0] \ (\ \ \.chine \ \\.p\<^sub>0) = ((\ \ \.chine) \ \\.p\<^sub>0) \ \\<^sup>-\<^sup>1[u\<^sub>1, \.chine, \\.p\<^sub>0]" using assoc'_naturality [of \ \.chine \\.p\<^sub>0] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[u, s\<^sub>0, \\.p\<^sub>0] \ ((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ ((\ \ \.chine) \ \\.p\<^sub>0) \ ((\\<^sup>-\<^sup>1[u\<^sub>1, \.chine, \\.p\<^sub>0] \ \[u\<^sub>1, \.chine, \\.p\<^sub>0]) \ (\.the_\ \ \\.p\<^sub>0)) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_assoc by metis also have "... = \[u, s\<^sub>0, \\.p\<^sub>0] \ (((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ ((\ \ \.chine) \ \\.p\<^sub>0) \ (\.the_\ \ \\.p\<^sub>0)) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1] \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "(\\<^sup>-\<^sup>1[u\<^sub>1, \.chine, \\.p\<^sub>0] \ \[u\<^sub>1, \.chine, \\.p\<^sub>0]) \ (\.the_\ \ \\.p\<^sub>0) = \.the_\ \ \\.p\<^sub>0" using comp_inv_arr' comp_cod_arr by auto thus ?thesis using comp_assoc by simp qed also have "... = (\[u, s\<^sub>0, \\.p\<^sub>0] \ ((\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0) \ r\<^sub>0s\<^sub>1.\ \ (\.the_\ \ \\.p\<^sub>1) \ \\<^sup>-\<^sup>1[t\<^sub>0, \.chine, \\.p\<^sub>1]) \ (t\<^sub>0 \ inv \\<^sub>\) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "arr ((u \ \.the_\) \ \[u, u\<^sub>0, \.chine] \ (\ \ \.chine) \ \.the_\)" using \.\_simps(3) by simp hence "((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ ((\ \ \.chine) \ \\.p\<^sub>0) \ (\.the_\ \ \\.p\<^sub>0) = (u \ \.the_\) \ \[u, u\<^sub>0, \.chine] \ (\ \ \.chine) \ \.the_\ \ \\.p\<^sub>0" using whisker_right by simp also have "... = (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0" using \.\_naturality by simp finally have "((u \ \.the_\) \ \\.p\<^sub>0) \ (\[u, u\<^sub>0, \.chine] \ \\.p\<^sub>0) \ ((\ \ \.chine) \ \\.p\<^sub>0) \ (\.the_\ \ \\.p\<^sub>0) = (\ \ s\<^sub>0) \ \ \ \\.p\<^sub>0" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (?\\<^sub>\ \ (t\<^sub>0 \ inv \\<^sub>\)) \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_assoc by presburger also have "... = ?\\<^sub>\' \ \[t\<^sub>0, \\.p\<^sub>1, chine] \ (inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using \\<^sub>\' \\<^sub>\'_eq by simp also have "... = (u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ ((t\<^sub>0u\<^sub>1.\ \ chine) \ ((\\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t\<^sub>0, \\.p\<^sub>1, chine]) \ (inv t\<^sub>0u\<^sub>1.\ \ chine)) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_assoc by presburger also have "... = (u \ the_\) \ \[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ ((\ \ \\.p\<^sub>0) \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "((t\<^sub>0u\<^sub>1.\ \ chine) \ ((\\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t\<^sub>0, \\.p\<^sub>1, chine]) \ (inv t\<^sub>0u\<^sub>1.\ \ chine)) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine] = \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" proof - have "((t\<^sub>0u\<^sub>1.\ \ chine) \ ((\\<^sup>-\<^sup>1[t\<^sub>0, \\.p\<^sub>1, chine]) \ \[t\<^sub>0, \\.p\<^sub>1, chine]) \ (inv t\<^sub>0u\<^sub>1.\ \ chine)) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine] = ((t\<^sub>0u\<^sub>1.\ \ chine) \ ((t\<^sub>0 \ \\.p\<^sub>1) \ chine) \ (inv t\<^sub>0u\<^sub>1.\ \ chine)) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_inv_arr' \.T0.antipar(1) by auto also have "... = ((t\<^sub>0u\<^sub>1.\ \ chine) \ (inv t\<^sub>0u\<^sub>1.\ \ chine)) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_cod_arr t\<^sub>0u\<^sub>1.\_uniqueness by simp also have "... = (t\<^sub>0u\<^sub>1.\ \ inv t\<^sub>0u\<^sub>1.\ \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using whisker_right t\<^sub>0u\<^sub>1.\_uniqueness by simp also have "... = ((u\<^sub>1 \ \\.p\<^sub>0) \ chine) \ \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_arr_inv' \.T0.antipar(1) t\<^sub>0u\<^sub>1.\_uniqueness by simp also have "... = \\<^sup>-\<^sup>1[u\<^sub>1, \\.p\<^sub>0, chine]" using comp_cod_arr \.T0.antipar(1) by simp finally show ?thesis by simp qed thus ?thesis using comp_assoc by simp qed also have "... = (u \ the_\) \ (\[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine) \ \\<^sup>-\<^sup>1[u \ u\<^sub>0, \\.p\<^sub>0, chine]) \ (\ \ ?w\<^sub>\)" using assoc'_naturality [of \ \\.p\<^sub>0 chine] comp_assoc by auto also have "... = ((u \ the_\) \ (u \ \\<^sup>-\<^sup>1[u\<^sub>0, \\.p\<^sub>0, chine])) \ \[u, u\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)" using uw\ pentagon comp_assoc invert_opposite_sides_of_square [of "u \ \[u\<^sub>0, \\.p\<^sub>0, chine]" "\[u, u\<^sub>0 \ \\.p\<^sub>0, chine] \ (\[u, u\<^sub>0, \\.p\<^sub>0] \ chine)" "\[u, u\<^sub>0, ?w\<^sub>\]" "\[u \ u\<^sub>0, \\.p\<^sub>0, chine]"] \.base_simps(2) \.ide_base \.ide_leg0 \.leg0_simps(2) assoc'_eq_inv_assoc ide_hcomp hcomp_simps(1) t\<^sub>0u\<^sub>1.ide_u by force also have "... = (u \ the_\ \ \\<^sup>-\<^sup>1[u\<^sub>0, \\.p\<^sub>0, chine]) \ \[u, u\<^sub>0, ?w\<^sub>\] \ (\ \ ?w\<^sub>\)" using whisker_left comp_assoc by simp finally show ?thesis by simp qed show "\!\. \\ : ?w\<^sub>\ \ ?w\<^sub>\'\ \ ?\\<^sub>\ = u\<^sub>1 \ \ \ ?\\<^sub>\ = ?\\<^sub>\' \ (u\<^sub>0 \ \)" using w\<^sub>\ w\<^sub>\' \\<^sub>\ \\<^sub>\' \\<^sub>\ eq\<^sub>\ \.T2 [of ?w\<^sub>\ ?w\<^sub>\' ?\\<^sub>\ ?u\<^sub>\ ?\\<^sub>\' ?\\<^sub>\] by fast qed obtain \\<^sub>\ where \\<^sub>\: "\\\<^sub>\ : ?w\<^sub>\ \ ?w\<^sub>\'\ \ ?\\<^sub>\ = u\<^sub>1 \ \\<^sub>\ \ ?\\<^sub>\ = ?\\<^sub>\' \ (u\<^sub>0 \ \\<^sub>\)" using 1 by auto show "?w\<^sub>\ \ ?w\<^sub>\'" using w\<^sub>\ w\<^sub>\' \\<^sub>\ BS3 [of ?w\<^sub>\ ?w\<^sub>\' \\<^sub>\ \\<^sub>\] isomorphic_def by auto qed lemma comp_L: shows "Maps.seq \\t\<^sub>0\\ \\\.chine \ \\.p\<^sub>1\\" and "\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\ = Maps.MkArr (src (\.chine \ \\.p\<^sub>1)) (src t) (Maps.Comp \t\<^sub>0\ \\.chine \ \\.p\<^sub>1\)" proof - show "Maps.seq \\t\<^sub>0\\ \\\.chine \ \\.p\<^sub>1\\" proof - have "is_left_adjoint (\.chine \ \\.p\<^sub>1)" using \.is_map r\<^sub>0s\<^sub>1.leg1_is_map left_adjoints_compose r\<^sub>0s\<^sub>1.p\<^sub>1_simps by auto thus ?thesis using Maps.CLS_in_hom r\<^sub>0s\<^sub>1.leg1_is_map apply (intro Maps.seqI') apply blast using Maps.CLS_in_hom [of t\<^sub>0] \.leg0_is_map \\.leg1_in_hom by auto qed thus "\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\ = Maps.MkArr (src (\.chine \ \\.p\<^sub>1)) (src t) (Maps.Comp \t\<^sub>0\ \\.chine \ \\.p\<^sub>1\)" using Maps.comp_char by auto qed lemma comp_R: shows "Maps.seq \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" and "\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\ = Maps.MkArr (src r\<^sub>0s\<^sub>1.p\<^sub>0) (trg u) (Maps.Comp \u\<^sub>1\ \\.chine \ r\<^sub>0s\<^sub>1.p\<^sub>0\)" proof - show "Maps.seq \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" proof - have "is_left_adjoint (\.chine \ \\.p\<^sub>0)" using \.is_map r\<^sub>0s\<^sub>1.leg0_is_map left_adjoints_compose [of \.chine \\.p\<^sub>0] by simp thus ?thesis using Maps.CLS_in_hom \.leg1_is_map apply (intro Maps.seqI') apply blast using Maps.CLS_in_hom [of u\<^sub>1] \.leg1_is_map by simp qed thus "\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\ = Maps.MkArr (src r\<^sub>0s\<^sub>1.p\<^sub>0) (trg u) (Maps.Comp \u\<^sub>1\ \\.chine \ r\<^sub>0s\<^sub>1.p\<^sub>0\)" using Maps.comp_char by auto qed lemma comp_L_eq_comp_R: shows "\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\ = \\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\" proof (intro Maps.arr_eqI) show "Maps.seq \\t\<^sub>0\\ \\\.chine \ \\.p\<^sub>1\\" using comp_L(1) by simp show "Maps.seq \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" using comp_R(1) by simp show "Maps.Dom (\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\) = Maps.Dom (\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\)" by (metis (no_types, lifting) Maps.Dom.simps(1) \.w_simps(2) \.w_simps(3) \.leg1_simps(3) \\.leg1_in_hom(2) comp_L(2) comp_R(2) hcomp_in_vhomE hseqI' r\<^sub>0s\<^sub>1.leg1_simps(3) hcomp_simps(1)) show "Maps.Cod (\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\) = Maps.Cod (\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\)" by (metis Maps.Cod.simps(1) \\.composable comp_L(2) comp_R(2)) have A: "Maps.Map (\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\) = Maps.Comp \t\<^sub>0\ \\.chine \ \\.p\<^sub>1\" using comp_L(1) Maps.comp_char by auto have B: "Maps.Map (\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\) = Maps.Comp \u\<^sub>1\ \\.chine \ \\.p\<^sub>0\" using comp_R(1) Maps.comp_char by auto have C: "Maps.Comp \t\<^sub>0\ \\.chine \ \\.p\<^sub>1\ = Maps.Comp \u\<^sub>1\ \\.chine \ \\.p\<^sub>0\" proof (intro Maps.Comp_eqI) show "t\<^sub>0 \ \.chine \ \\.p\<^sub>1 \ Maps.Comp \t\<^sub>0\ \\.chine \ \\.p\<^sub>1\" proof (intro Maps.in_CompI) show "is_iso_class \\.chine \ \\.p\<^sub>1\" using prj_chine(2) is_iso_classI isomorphic_implies_hpar(2) by blast show "is_iso_class \t\<^sub>0\" using is_iso_classI by auto show "\.chine \ \\.p\<^sub>1 \ \\.chine \ \\.p\<^sub>1\" using ide_in_iso_class prj_chine(2) isomorphic_implies_hpar(2) by blast show "t\<^sub>0 \ \t\<^sub>0\" using ide_in_iso_class by simp show "t\<^sub>0 \ \.chine \ \\.p\<^sub>1 \ t\<^sub>0 \ \.chine \ \\.p\<^sub>1" using isomorphic_reflexive prj_chine(2) isomorphic_implies_hpar(2) by auto qed show "u\<^sub>1 \ \.chine \ \\.p\<^sub>0 \ Maps.Comp \u\<^sub>1\ \\.chine \ \\.p\<^sub>0\" proof (intro Maps.in_CompI) show "is_iso_class \\.chine \ \\.p\<^sub>0\" using is_iso_classI by simp show "is_iso_class \u\<^sub>1\" using is_iso_classI by simp show "\.chine \ \\.p\<^sub>0 \ \\.chine \ \\.p\<^sub>0\" using ide_in_iso_class by simp show "u\<^sub>1 \ iso_class u\<^sub>1" using ide_in_iso_class by simp show "u\<^sub>1 \ \.chine \ \\.p\<^sub>0 \ u\<^sub>1 \ \.chine \ \\.p\<^sub>0" using isomorphic_reflexive isomorphic_implies_hpar(2) by auto qed show "t\<^sub>0 \ \.chine \ \\.p\<^sub>1 \ u\<^sub>1 \ \.chine \ \\.p\<^sub>0" proof - have "t\<^sub>0 \ \.chine \ \\.p\<^sub>1 \ (t\<^sub>0 \ \.chine) \ \\.p\<^sub>1" using assoc'_in_hom [of t\<^sub>0 \.chine \\.p\<^sub>1] iso_assoc' isomorphic_def r\<^sub>0s\<^sub>1.p\<^sub>1_simps by auto also have "... \ r\<^sub>0 \ \\.p\<^sub>1" using \.leg0_uniquely_isomorphic hcomp_isomorphic_ide by (simp add: \.T0.antipar(1)) also have "... \ s\<^sub>1 \ \\.p\<^sub>0" using isomorphic_def r\<^sub>0s\<^sub>1.\_uniqueness(2) by blast also have "... \ (u\<^sub>1 \ \.chine) \ \\.p\<^sub>0" using \.leg1_uniquely_isomorphic hcomp_isomorphic_ide by auto also have "... \ u\<^sub>1 \ \.chine \ \\.p\<^sub>0" using assoc_in_hom [of u\<^sub>1 \.chine \\.p\<^sub>0] iso_assoc isomorphic_def by auto finally show ?thesis by simp qed qed show "Maps.Map (\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\) = Maps.Map (\\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\)" using A B C by simp qed lemma csq: shows "Maps.commutative_square \\t\<^sub>0\\ \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" proof show "Maps.cospan \\t\<^sub>0\\ \\u\<^sub>1\\" using comp_L(1) comp_R(1) comp_L_eq_comp_R by (metis (no_types, lifting) Maps.cod_comp Maps.seq_char) show "Maps.span \\\.chine \ \\.p\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" using comp_L(1) comp_R(1) comp_L_eq_comp_R by (metis (no_types, lifting) Maps.dom_comp Maps.seq_char) show "Maps.dom \\t\<^sub>0\\ = Maps.cod \\\.chine \ \\.p\<^sub>1\\" using comp_L(1) by auto show "\\t\<^sub>0\\ \ \\\.chine \ \\.p\<^sub>1\\ = \\u\<^sub>1\\ \ \\\.chine \ \\.p\<^sub>0\\" using comp_L_eq_comp_R by simp qed lemma CLS_chine: shows "\\chine\\ = Maps.tuple \\\.chine \ \\.p\<^sub>1\\ \\t\<^sub>0\\ \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" proof - let ?T = "Maps.tuple \\\.chine \ \\.p\<^sub>1\\ \\t\<^sub>0\\ \\u\<^sub>1\\ \\\.chine \ \\.p\<^sub>0\\" have "\!l. \\t\<^sub>0u\<^sub>1.p\<^sub>1\\ \ l = \\\.chine \ \\.p\<^sub>1\\ \ \\t\<^sub>0u\<^sub>1.p\<^sub>0\\ \ l = \\\.chine \ \\.p\<^sub>0\\" using csq \\.prj_char Maps.universal [of "\\t\<^sub>0\\" "\\u\<^sub>1\\" "\\\.chine \ \\.p\<^sub>1\\" "\\\.chine \ \\.p\<^sub>0\\"] by simp moreover have "\\\\.p\<^sub>1\\ \ ?T = \\\.chine \ \\.p\<^sub>1\\ \ \\\\.p\<^sub>0\\ \ ?T = \\\.chine \ \\.p\<^sub>0\\" using csq \\.prj_char Maps.prj_tuple [of "\\t\<^sub>0\\" "\\u\<^sub>1\\" "\\\.chine \ \\.p\<^sub>1\\" "\\\.chine \ \\.p\<^sub>0\\"] by simp moreover have "\\t\<^sub>0u\<^sub>1.p\<^sub>1\\ \ \\chine\\ = \\\.chine \ \\.p\<^sub>1\\ \ \\t\<^sub>0u\<^sub>1.p\<^sub>0\\ \ \\chine\\ = \\\.chine \ \\.p\<^sub>0\\" using prj_chine \\.leg0_is_map \\.leg1_is_map is_map t\<^sub>0u\<^sub>1.leg1_is_map t\<^sub>0u\<^sub>1.satisfies_T0 Maps.comp_CLS by blast ultimately show "\\chine\\ = ?T" by auto qed end subsection "Equivalence of B and Span(Maps(B))" subsubsection "The Functor SPN" text \ We now define a function \SPN\ on arrows and will ultimately show that it extends to a biequivalence from the underlying bicategory \B\ to \Span(Maps(B))\. The idea is that \SPN\ takes \\\ : r \ s\\ to the isomorphism class of an induced arrow of spans from the chosen tabulation of \r\ to the chosen tabulation of \s\. To obtain this, we first use isomorphisms \r.tab\<^sub>1 \ r.tab\<^sub>0\<^sup>* \ r\ and \s.tab\<^sub>1 \ s.tab\<^sub>0\<^sup>* \ s\ to transform \\\ to \\\' : r.tab\<^sub>1 \ r.tab\<^sub>0\<^sup>* \ s.tab\<^sub>1 \ s.tab\<^sub>0\<^sup>*\\. We then take the adjoint transpose of \\'\ to obtain \\\ : r.tab\<^sub>1 \ (s.tab\<^sub>1 \ s.tab\<^sub>0\<^sup>*) \ r.tab\<^sub>0\\. The 2-cell \\\ induces a map \w\ which is an arrow of spans from \(r.tab\<^sub>0, r.tab\<^sub>1)\ to \(s.tab\<^sub>0, s.tab\<^sub>1)\. We take the arrow of \Span(Maps(B))\ defined by \w\ as the value of \SPN \\. Ensuring that \SPN\ is functorial is a somewhat delicate point, which requires that all the underlying definitions that have been set up are ``just so'', with no extra choices other than those that are forced, and with the tabulation assigned to each 1-cell \r\ in the proper relationship with the canonical tabulation assigned to its chosen factorization \r = g \ f\<^sup>*\. \ context bicategory_of_spans begin interpretation Maps: maps_category V H \ \ src trg .. interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. no_notation Fun.comp (infixl "\" 55) notation Span.vcomp (infixr "\" 55) notation Span.hcomp (infixr "\" 53) notation Maps.comp (infixr "\" 55) notation isomorphic (infix "\" 50) definition spn where "spn \ \ arrow_of_tabulations_in_maps.chine V H \ \ src trg (tab_of_ide (dom \)) (tab\<^sub>0 (dom \)) (cod \) (tab_of_ide (cod \)) (tab\<^sub>0 (cod \)) (tab\<^sub>1 (cod \)) \" lemma is_induced_map_spn: assumes "arr \" shows "arrow_of_tabulations_in_maps.is_induced_map V H \ \ src trg (tab_of_ide (dom \)) (tab\<^sub>0 (dom \)) (cod \) (tab_of_ide (cod \)) (tab\<^sub>0 (cod \)) (tab\<^sub>1 (cod \)) \ (spn \)" proof - interpret \: arrow_in_bicategory_of_spans V H \ \ src trg \dom \\ \cod \\ \ using assms by unfold_locales auto interpret \: arrow_of_tabulations_in_maps V H \ \ src trg \dom \\ \.r.tab \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \.s.tab \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (cod \)\ \ using \.is_arrow_of_tabulations_in_maps by simp show ?thesis unfolding spn_def using \.chine_is_induced_map by blast qed lemma spn_props: assumes "arr \" shows "\spn \ : src (tab\<^sub>0 (dom \)) \ src (tab\<^sub>0 (cod \))\" and "is_left_adjoint (spn \)" and "tab\<^sub>0 (cod \) \ spn \ \ tab\<^sub>0 (dom \)" and "tab\<^sub>1 (cod \) \ spn \ \ tab\<^sub>1 (dom \)" proof - interpret \: arrow_in_bicategory_of_spans V H \ \ src trg \dom \\ \cod \\ \ using assms by unfold_locales auto interpret \: arrow_of_tabulations_in_maps V H \ \ src trg \dom \\ \.r.tab \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \.s.tab \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (cod \)\ \ using \.is_arrow_of_tabulations_in_maps by simp show "\spn \ : src (tab\<^sub>0 (dom \)) \ src (tab\<^sub>0 (cod \))\" using spn_def by simp show "is_left_adjoint (spn \)" using spn_def by (simp add: \.is_map) show "tab\<^sub>0 (cod \) \ spn \ \ tab\<^sub>0 (dom \)" using spn_def isomorphic_def \.leg0_uniquely_isomorphic(1) by auto show "tab\<^sub>1 (cod \) \ spn \ \ tab\<^sub>1 (dom \)" using spn_def isomorphic_def isomorphic_symmetric \.leg1_uniquely_isomorphic(1) by auto qed lemma spn_in_hom [intro]: assumes "arr \" shows "\spn \ : src (tab\<^sub>0 (dom \)) \ src (tab\<^sub>0 (cod \))\" and "\spn \ : spn \ \ spn \\" using assms spn_props left_adjoint_is_ide by auto lemma spn_simps [simp]: assumes "arr \" shows "is_left_adjoint (spn \)" and "ide (spn \)" and "src (spn \) = src (tab\<^sub>0 (dom \))" and "trg (spn \) = src (tab\<^sub>0 (cod \))" using assms spn_props left_adjoint_is_ide by auto text \ We need the next result to show that \SPN\ is functorial; in particular, that it takes \\r : r \ r\\ in the underlying bicategory to a 1-cell in \Span(Maps(B))\. The 1-cells in \Span(Maps(B))\ have objects of \Maps(B)\ as their chines, and objects of \Maps(B)\ are isomorphism classes of objects in the underlying bicategory \B\. So we need the induced map associated with \r\ to be isomorphic to an object. \ lemma spn_ide: assumes "ide r" shows "spn r \ src (tab\<^sub>0 r)" proof - interpret r: identity_in_bicategory_of_spans V H \ \ src trg r using assms by unfold_locales auto interpret r: arrow_of_tabulations_in_maps V H \ \ src trg r r.tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ r r.tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ r using r.is_arrow_of_tabulations_in_maps by simp interpret tab: tabulation V H \ \ src trg r \r.tab\ \tab\<^sub>0 r\ \dom r.tab\ using assms r.tab_is_tabulation by simp interpret tab: tabulation_in_maps V H \ \ src trg r \r.tab\ \tab\<^sub>0 r\ \dom r.tab\ by (unfold_locales, simp_all) have "tab.is_induced_by_cell (spn r) (tab\<^sub>0 r) r.tab" using spn_def comp_ide_arr r.chine_is_induced_map by auto thus ?thesis using tab.induced_map_unique [of "tab\<^sub>0 r" "r.tab" "spn r" "src r.s\<^sub>0"] tab.apex_is_induced_by_cell by (simp add: comp_assoc) qed text \ The other key result we need to show that \SPN\ is functorial is to show that the induced map of a composite is isomorphic to the composite of induced maps. \ lemma spn_hcomp: assumes "seq \ \" and "g \ spn \" and "f \ spn \" shows "spn (\ \ \) \ g \ f" proof - interpret \: arrow_in_bicategory_of_spans V H \ \ src trg \dom \\ \cod \\ \ using assms by unfold_locales auto interpret \: arrow_of_tabulations_in_maps V H \ \ src trg \dom \\ \.r.tab \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \.s.tab \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (cod \)\ \ using \.is_arrow_of_tabulations_in_maps by simp interpret \: arrow_in_bicategory_of_spans V H \ \ src trg \dom \\ \dom \\ \ using assms apply unfold_locales apply auto[1] by (elim seqE, auto) interpret \: arrow_of_tabulations_in_maps V H \ \ src trg \dom \\ \.r.tab \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \dom \\ \.r.tab \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \ using \.is_arrow_of_tabulations_in_maps by simp interpret \\: vertical_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg \dom \\ \.r.tab \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \dom \\ \.r.tab \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \.s.tab \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (cod \)\ \ \ .. have "g \ \.chine" using assms(2) spn_def by auto moreover have "f \ \.chine" using assms(1) assms(3) spn_def by auto moreover have "src g = trg f" using calculation(1-2) isomorphic_implies_hpar(3-4) by auto moreover have "src g = trg \.chine" using calculation(1) isomorphic_implies_hpar(3) by auto ultimately have "g \ f \ \.chine \ \.chine" using hcomp_ide_isomorphic hcomp_isomorphic_ide isomorphic_transitive by (meson \.is_ide isomorphic_implies_ide(1)) also have "... \ spn (\ \ \)" using spn_def \\.chine_char isomorphic_symmetric by (metis \\.in_hom in_homE) finally show ?thesis using isomorphic_symmetric by simp qed abbreviation (input) SPN\<^sub>0 where "SPN\<^sub>0 r \ Span.mkIde \\tab\<^sub>0 r\\ \\tab\<^sub>1 r\\" definition SPN where "SPN \ \ if arr \ then \Chn = \\spn \\\, Dom = \Leg0 = \\tab\<^sub>0 (dom \)\\, Leg1 = \\tab\<^sub>1 (dom \)\\\, Cod = \Leg0 = \\tab\<^sub>0 (cod \)\\, Leg1 = \\tab\<^sub>1 (cod \)\\\\ else Span.null" lemma Dom_SPN [simp]: assumes "arr \" shows "Dom (SPN \) = \Leg0 = \\tab\<^sub>0 (dom \)\\, Leg1 = \\tab\<^sub>1 (dom \)\\\" using assms SPN_def by simp lemma Cod_SPN [simp]: assumes "arr \" shows "Cod (SPN \) = \Leg0 = \\tab\<^sub>0 (cod \)\\, Leg1 = \\tab\<^sub>1 (cod \)\\\" using assms SPN_def by simp text \Now we have to show this does the right thing for us.\ lemma SPN_in_hom: assumes "arr \" shows "Span.in_hom (SPN \) (SPN\<^sub>0 (dom \)) (SPN\<^sub>0 (cod \))" proof interpret Dom: span_in_category Maps.comp \Dom (SPN \)\ proof interpret r: identity_in_bicategory_of_spans V H \ \ src trg \dom \\ using assms by unfold_locales auto show "Maps.span (Leg0 (Dom (SPN \))) (Leg1 (Dom (SPN \)))" using assms Maps.CLS_in_hom SPN_def by (metis (no_types, lifting) Maps.in_homE bicategory_of_spans.Dom_SPN bicategory_of_spans_axioms r.leg1_is_map r.leg1_simps(3) r.satisfies_T0 span_data.simps(1) span_data.simps(2)) qed interpret Dom': span_in_category Maps.comp \\Leg0 = \\tab\<^sub>0 (dom \)\\, Leg1 = \\tab\<^sub>1 (dom \)\\\\ using assms Dom.span_in_category_axioms SPN_def by simp interpret Cod: span_in_category Maps.comp \Cod (SPN \)\ proof interpret s: identity_in_bicategory_of_spans V H \ \ src trg \cod \\ using assms by unfold_locales auto show "Maps.span (Leg0 (Cod (SPN \))) (Leg1 (Cod (SPN \)))" using assms Maps.CLS_in_hom SPN_def by (metis (no_types, lifting) bicategory_of_spans.Cod_SPN bicategory_of_spans_axioms ide_dom s.base_simps(2) s.base_simps(3) s.determines_span span_in_category.is_span) qed interpret Cod': span_in_category Maps.comp \\Leg0 = \\tab\<^sub>0 (cod \)\\, Leg1 = \\tab\<^sub>1 (cod \)\\\\ using assms Cod.span_in_category_axioms SPN_def by simp show 1: "Span.arr (SPN \)" proof (unfold Span.arr_char) show "arrow_of_spans Maps.comp (SPN \)" proof (unfold_locales) show "Maps.in_hom (Chn (SPN \)) Dom.apex Cod.apex" unfolding SPN_def Maps.in_hom_char using assms Dom'.apex_def Cod'.apex_def Dom'.is_span Cod'.is_span Maps.arr_char by auto show "Cod.leg0 \ Chn (SPN \) = Dom.leg0" unfolding SPN_def using assms spn_props [of \] Maps.comp_CLS [of "tab\<^sub>0 (cod \)" "spn \"] by simp show "Cod.leg1 \ Chn (SPN \) = Dom.leg1" unfolding SPN_def using assms spn_props [of \] Maps.comp_CLS [of "tab\<^sub>1 (cod \)" "spn \"] by simp qed qed show "Span.dom (SPN \) = SPN\<^sub>0 (dom \)" using assms 1 Span.dom_char Dom'.apex_def SPN_def by simp show "Span.cod (SPN \) = SPN\<^sub>0 (cod \)" using assms 1 Span.cod_char Cod'.apex_def SPN_def by simp qed interpretation SPN: "functor" V Span.vcomp SPN proof show "\\. \ arr \ \ SPN \ = Span.null" unfolding SPN_def by simp show 1: "\\. arr \ \ Span.arr (SPN \)" using SPN_in_hom by auto show "\\. arr \ \ Span.dom (SPN \) = SPN (dom \)" proof - fix \ assume \: "arr \" have 1: "Maps.arr (Maps.MkArr (src (tab\<^sub>0 (dom \))) (src \) \tab\<^sub>0 (dom \)\)" proof - have "src (tab\<^sub>0 (dom \)) \ Collect obj" using \ by simp moreover have "src \ \ Collect obj" using \ by simp moreover have "\tab\<^sub>0 (dom \)\ \ Maps.Hom (src (tab\<^sub>0 (local.dom \))) (src \)" proof - have "\tab\<^sub>0 (dom \) : src (tab\<^sub>0 (dom \)) \ src \\" using \ by simp moreover have "is_left_adjoint (tab\<^sub>0 (dom \))" using \ tab\<^sub>0_simps [of "dom \"] by auto ultimately show ?thesis by auto qed ultimately show ?thesis by simp qed have "\spn (dom \)\ = \src (tab\<^sub>0 (dom \))\" using \ spn_ide iso_class_eqI by auto hence "SPN (dom \) = SPN\<^sub>0 (dom \)" unfolding SPN_def using \ 1 Maps.dom_char by simp thus "Span.dom (SPN \) = SPN (dom \)" using \ SPN_in_hom by auto qed show "\\. arr \ \ Span.cod (SPN \) = SPN (cod \)" proof - fix \ assume \: "arr \" have 1: "Maps.arr (Maps.MkArr (src (tab\<^sub>0 (cod \))) (src \) \tab\<^sub>0 (cod \)\)" proof - have "src (tab\<^sub>0 (cod \)) \ Collect obj" using \ by simp moreover have "src \ \ Collect obj" using \ by simp moreover have "\tab\<^sub>0 (cod \)\ \ Maps.Hom (src (tab\<^sub>0 (cod \))) (src \)" proof - have "\tab\<^sub>0 (cod \) : src (tab\<^sub>0 (cod \)) \ src \\" using \ by simp moreover have "is_left_adjoint (tab\<^sub>0 (cod \))" using \ by simp ultimately show ?thesis by auto qed ultimately show ?thesis by simp qed have "\spn (cod \)\ = \src (tab\<^sub>0 (cod \))\" using \ spn_ide iso_class_eqI by auto hence "SPN (cod \) = SPN\<^sub>0 (cod \)" unfolding SPN_def using \ 1 Maps.dom_char by simp thus "Span.cod (SPN \) = SPN (cod \)" using \ SPN_in_hom by auto qed show "\\ \. seq \ \ \ SPN (\ \ \) = SPN \ \ SPN \" proof - fix \ \ assume seq: "seq \ \" have "Dom (SPN (\ \ \)) = Dom (SPN \ \ SPN \)" using seq 1 Span.vcomp_def Span.arr_char by (elim seqE, simp) moreover have "Cod (SPN (\ \ \)) = Cod (SPN \ \ SPN \)" using seq 1 Span.vcomp_def Span.arr_char by (elim seqE, simp) moreover have "Chn (SPN (\ \ \)) = Chn (SPN \ \ SPN \)" proof - have *: "\spn (\ \ \)\ = Maps.Comp \spn \\ \spn \\" proof show "\spn (\ \ \)\ \ Maps.Comp \spn \\ \spn \\" proof fix h assume h: "h \ \spn (\ \ \)\" show "h \ Maps.Comp \spn \\ \spn \\" proof - have 1: "spn \ \ \spn \\" using seq ide_in_iso_class by auto moreover have 2: "spn \ \ \spn \\" using seq ide_in_iso_class by auto moreover have "spn \ \ spn \ \ h" proof - have "spn \ \ spn \ \ spn (\ \ \)" using seq spn_hcomp 1 2 iso_class_def isomorphic_reflexive isomorphic_symmetric by simp thus ?thesis using h isomorphic_transitive iso_class_def by simp qed ultimately show ?thesis unfolding Maps.Comp_def by (metis (mono_tags, lifting) is_iso_classI spn_simps(2) mem_Collect_eq seq seqE) qed qed show "Maps.Comp \spn \\ \spn \\ \ \spn (\ \ \)\" proof fix h assume h: "h \ Maps.Comp \spn \\ \spn \\" show "h \ \spn (\ \ \)\" proof - obtain f g where 1: "g \ \spn \\ \ f \ \spn \\ \ g \ f \ h" using h Maps.Comp_def [of "iso_class (spn \)" "iso_class (spn \)"] iso_class_def iso_class_elems_isomorphic by blast have fg: "g \ spn \ \ f \ spn \ \ g \ f \ h" proof - have "spn \ \ \spn \\ \ spn \ \ \spn \\" using seq ide_in_iso_class by auto thus ?thesis using 1 iso_class_elems_isomorphic isomorphic_symmetric is_iso_classI by (meson spn_simps(2) seq seqE) qed have "g \ f \ \spn (\ \ \)\" using seq fg 1 spn_hcomp iso_class_def isomorphic_symmetric by simp thus ?thesis using fg isomorphic_transitive iso_class_def by blast qed qed qed have "Chn (SPN \ \ SPN \) = Maps.MkArr (src (tab\<^sub>0 (cod \))) (src (tab\<^sub>0 (cod \))) \spn \\ \ Maps.MkArr (src (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (cod \))) \spn \\" using 1 seq SPN_def Span.vcomp_def Span.arr_char apply (elim seqE) apply simp by (metis (no_types, lifting) seq vseq_implies_hpar(1) vseq_implies_hpar(2)) also have "... = Maps.MkArr (src (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (cod \))) (Maps.Comp \spn \\ \spn \\)" proof - have "Maps.seq (Maps.MkArr (src (tab\<^sub>0 (cod \))) (src (tab\<^sub>0 (cod \))) \spn \\) (Maps.MkArr (src (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (cod \))) \spn \\)" proof show "Maps.in_hom (Maps.MkArr (src (tab\<^sub>0 (local.dom \))) (src (tab\<^sub>0 (cod \))) \spn \\) (Maps.MkIde (src (tab\<^sub>0 (dom \)))) (Maps.MkIde (src (tab\<^sub>0 (cod \))))" proof - have "src (tab\<^sub>0 (dom \)) \ Collect obj" using in_hhom_def seq by auto moreover have "src (tab\<^sub>0 (cod \)) \ Collect obj" using seq by auto moreover have "\spn \\ \ Maps.Hom (src (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (cod \)))" using spn_props by (metis (mono_tags, lifting) mem_Collect_eq seq seqE) ultimately show ?thesis using Maps.MkArr_in_hom by simp qed show "Maps.in_hom (Maps.MkArr (src (tab\<^sub>0 (cod \))) (src (tab\<^sub>0 (cod \))) \spn \\) (Maps.MkIde (src (tab\<^sub>0 (cod \)))) (Maps.MkIde (src (tab\<^sub>0 (cod \))))" proof - have "src (tab\<^sub>0 (cod \)) \ Collect obj" using in_hhom_def seq by auto moreover have "src (tab\<^sub>0 (cod \)) \ Collect obj" using seq by auto moreover have "\spn \\ \ Maps.Hom (src (tab\<^sub>0 (cod \))) (src (tab\<^sub>0 (cod \)))" using spn_props by (metis (mono_tags, lifting) mem_Collect_eq seq seqE) ultimately show ?thesis using Maps.MkArr_in_hom by simp qed qed thus ?thesis using Maps.comp_char [of "Maps.MkArr (src (tab\<^sub>0 (cod \))) (src (tab\<^sub>0 (cod \))) \spn \\" "Maps.MkArr (src (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (cod \))) \spn \\"] by simp qed also have "... = Maps.MkArr (src (tab\<^sub>0 (dom \))) (src (tab\<^sub>0 (cod \))) \spn (\ \ \)\" using * by simp also have "... = Chn (SPN (\ \ \))" using seq SPN_def Span.vcomp_def by (elim seqE, simp) finally show ?thesis by simp qed ultimately show "SPN (\ \ \) = SPN \ \ SPN \" by simp qed qed lemma SPN_is_functor: shows "functor V Span.vcomp SPN" .. interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN proof show "\\. arr \ \ Span.isomorphic (SPN (src \)) (Span.src (SPN \))" proof - fix \ assume \: "arr \" let ?src = "Maps.MkIde (src \)" have src: "Maps.ide ?src" using \ by simp interpret src: identity_in_bicategory_of_spans V H \ \ src trg \src \\ using \ by unfold_locales auto interpret src: arrow_of_tabulations_in_maps V H \ \ src trg \src \\ src.tab \tab\<^sub>0 (src \)\ \tab\<^sub>1 (src \)\ \src \\ src.tab \tab\<^sub>0 (src \)\ \tab\<^sub>1 (src \)\ \src \\ using src.is_arrow_of_tabulations_in_maps by simp interpret src: span_in_category Maps.comp \\Leg0 = ?src, Leg1 = ?src\\ using src by (unfold_locales, simp) let ?tab\<^sub>0 = "Maps.MkArr (src (tab\<^sub>0 (src \))) (src \) \tab\<^sub>0 (src \)\" have tab\<^sub>0_src: "\tab\<^sub>0 (src \) : src (tab\<^sub>0 (src \)) \ src \\ \ is_left_adjoint (tab\<^sub>0 (src \)) \ \tab\<^sub>0 (src \)\ = \tab\<^sub>0 (src \)\" using \ by simp have tab\<^sub>0: "Maps.arr ?tab\<^sub>0" using \ Maps.arr_MkArr tab\<^sub>0_src by blast let ?tab\<^sub>1 = "Maps.MkArr (src (tab\<^sub>0 (src \))) (src \) \tab\<^sub>1 (src \)\" have tab\<^sub>1_src: "\tab\<^sub>1 (src \) : src (tab\<^sub>0 (src \)) \ src \\ \ is_left_adjoint (tab\<^sub>1 (src \)) \ \tab\<^sub>1 (src \)\ = \tab\<^sub>1 (src \)\" using \ by simp have tab\<^sub>1: "Maps.arr ?tab\<^sub>1" using \ Maps.arr_MkArr tab\<^sub>1_src by blast interpret tab: span_in_category Maps.comp \\Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\\ using tab\<^sub>0 tab\<^sub>1 Maps.dom_char Maps.cod_char by (unfold_locales, simp) have "src \ \ tab\<^sub>0 (src \) \ tab\<^sub>0 (src \)" using \ iso_lunit isomorphic_def by (metis lunit_in_hom(2) src.ide_u src.u_simps(3) src_src) hence "src \ \ tab\<^sub>0 (src \) \ tab\<^sub>1 (src \)" using \ src.obj_has_symmetric_tab isomorphic_transitive by blast have "\tab\<^sub>0 (src \)\ \ Maps.Hom (src (tab\<^sub>0 (src \))) (src \)" using \ tab\<^sub>0_src by blast have "\src \\ \ Maps.Hom (src \) (src \)" proof - have "\src \ : src \ \ src \\ \ is_left_adjoint (src \) \ \src \\ = \src \\" using \ obj_is_self_adjoint by simp thus ?thesis by auto qed interpret SPN_src: arrow_of_spans Maps.comp \SPN (src \)\ using \ SPN.preserves_reflects_arr Span.arr_char by blast have SPN_src: "SPN (src \) = \Chn = Maps.MkArr (src (tab\<^sub>0 (src \))) (src (tab\<^sub>0 (src \))) \spn (src \)\, Dom = \Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\, Cod = \Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\\" unfolding SPN_def using \ by simp interpret src_SPN: arrow_of_spans Maps.comp \Span.src (SPN \)\ using \ SPN.preserves_reflects_arr Span.arr_char by blast have src_SPN: "Span.src (SPN \) = \Chn = ?src, Dom = \Leg0 = ?src, Leg1 = ?src\, Cod = \Leg0 = ?src, Leg1 = ?src\\" proof - let ?tab\<^sub>0_dom = "Maps.MkArr (src (tab\<^sub>0 (dom \))) (src \) \tab\<^sub>0 (dom \)\" have "Maps.arr ?tab\<^sub>0_dom" proof - have "\tab\<^sub>0 (dom \) : src (tab\<^sub>0 (dom \)) \ src \\ \ is_left_adjoint (tab\<^sub>0 (dom \)) \ \tab\<^sub>0 (dom \)\ = \tab\<^sub>0 (dom \)\" using \ by simp thus ?thesis using \ Maps.arr_MkArr by blast qed thus ?thesis using \ Maps.cod_char Span.src_def by simp qed text \ The idea of the proof is that @{term "iso_class (tab\<^sub>0 (src \))"} is invertible in \Maps(B)\ and determines an invertible arrow of spans from @{term "SPN (src \)"} to @{term "Span.src (SPN \)"}. \ let ?\ = "\Chn = ?tab\<^sub>0, Dom = Dom (SPN (src \)), Cod = Cod (Span.src (SPN \))\" interpret \: arrow_of_spans Maps.comp ?\ apply (unfold_locales, simp_all) proof - show "Maps.in_hom ?tab\<^sub>0 SPN_src.dom.apex src_SPN.cod.apex" using \ tab\<^sub>0 Maps.dom_char Maps.cod_char SPN_src src_SPN tab.apex_def src_SPN.cod.apex_def apply (intro Maps.in_homI) by simp_all show "src_SPN.cod.leg0 \ ?tab\<^sub>0 = SPN_src.dom.leg0" proof - have "Maps.seq src_SPN.cod.leg0 ?tab\<^sub>0" using \ src_SPN tab\<^sub>0 Maps.dom_char Maps.cod_char by (intro Maps.seqI, auto) moreover have "Maps.Comp \src \\ \tab\<^sub>0 (src \)\ = \tab\<^sub>0 (src \)\" proof - have "tab\<^sub>0 (src \) \ Maps.Comp \src \\ \tab\<^sub>0 (src \)\" using \ is_iso_classI ide_in_iso_class [of "src \"] ide_in_iso_class [of "tab\<^sub>0 (src \)"] \src \ \ tab\<^sub>0 (src \) \ tab\<^sub>0 (src \)\ by auto thus ?thesis using Maps.Comp_eq_iso_class_memb \\tab\<^sub>0 (src \)\ \ Maps.Hom (src (tab\<^sub>0 (src \))) (src \)\ \\src \\ \ Maps.Hom (src \) (src \)\ by meson qed ultimately show ?thesis using \ Maps.comp_char [of src_SPN.cod.leg0 ?tab\<^sub>0] src_SPN by simp qed show "src_SPN.cod.leg1 \ ?tab\<^sub>0 = SPN_src.dom.leg1" proof - have "Maps.seq src_SPN.cod.leg1 ?tab\<^sub>0" using \ src_SPN tab\<^sub>0 Maps.dom_char Maps.cod_char by (intro Maps.seqI, auto) moreover have "Maps.Comp \src \\ \tab\<^sub>0 (src \)\ = \tab\<^sub>1 (src \)\" proof - have "tab\<^sub>1 (src \) \ Maps.Comp \src \\ \tab\<^sub>0 (src \)\" using \ is_iso_classI ide_in_iso_class [of "src \"] ide_in_iso_class [of "tab\<^sub>0 (src \)"] \isomorphic (src \ \ tab\<^sub>0 (src \)) (tab\<^sub>1 (src \))\ by auto thus ?thesis using Maps.Comp_eq_iso_class_memb \\tab\<^sub>0 (src \)\ \ Maps.Hom (src (tab\<^sub>0 (src \))) (src \)\ \\src \\ \ Maps.Hom (src \) (src \)\ by meson qed ultimately show ?thesis using \ Maps.comp_char [of src_SPN.cod.leg1 ?tab\<^sub>0] src_SPN by simp qed qed have "Span.in_hom ?\ (SPN (src \)) (Span.src (SPN \))" using \ tab\<^sub>0 spn_ide [of "src \"] iso_class_eqI Span.arr_char Span.dom_char Span.cod_char \.arrow_of_spans_axioms SPN_src src_SPN src.apex_def tab.apex_def Maps.dom_char apply (intro Span.in_homI) by auto (* The preceding cannot be written "by (intro Span.in_homI, auto)", why? *) moreover have "Maps.iso ?tab\<^sub>0" using \ tab\<^sub>0 ide_in_iso_class src.is_map_iff_tab\<^sub>0_is_equivalence obj_is_self_adjoint Maps.iso_char' [of ?tab\<^sub>0] by auto ultimately show "Span.isomorphic (SPN (src \)) (Span.src (SPN \))" using Span.isomorphic_def Span.iso_char by auto qed show "\\. arr \ \ Span.isomorphic (SPN (trg \)) (Span.trg (SPN \))" proof - fix \ assume \: "arr \" let ?trg = "Maps.MkIde (trg \)" have trg: "Maps.ide ?trg" using \ by simp interpret trg: identity_in_bicategory_of_spans V H \ \ src trg \trg \\ using \ by unfold_locales auto interpret trg: span_in_category Maps.comp \\Leg0 = ?trg, Leg1 = ?trg\\ using trg by (unfold_locales, simp) let ?tab\<^sub>0 = "Maps.MkArr (src (tab\<^sub>0 (trg \))) (trg \) \tab\<^sub>0 (trg \)\" have tab\<^sub>0_trg: "\tab\<^sub>0 (trg \) : src (tab\<^sub>0 (trg \)) \ trg \\ \ is_left_adjoint (tab\<^sub>0 (trg \)) \ \tab\<^sub>0 (trg \)\ = \tab\<^sub>0 (trg \)\" using \ by simp have tab\<^sub>0: "Maps.arr ?tab\<^sub>0" using \ Maps.arr_MkArr tab\<^sub>0_trg by blast let ?tab\<^sub>1 = "Maps.MkArr (src (tab\<^sub>0 (trg \))) (trg \) \tab\<^sub>1 (trg \)\" have tab\<^sub>1_trg: "\tab\<^sub>1 (trg \) : src (tab\<^sub>0 (trg \)) \ trg \\ \ is_left_adjoint (tab\<^sub>1 (trg \)) \ \tab\<^sub>1 (trg \)\ = \tab\<^sub>1 (trg \)\" using \ by simp have tab\<^sub>1: "Maps.arr ?tab\<^sub>1" using \ Maps.arr_MkArr tab\<^sub>1_trg by blast interpret tab: span_in_category Maps.comp \\Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\\ using tab\<^sub>0 tab\<^sub>1 Maps.dom_char Maps.cod_char by (unfold_locales, simp) have "trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>0 (trg \)" proof - have "\\[tab\<^sub>1 (trg \)] : trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>1 (trg \)\" using \ by simp moreover have "iso \[tab\<^sub>1 (trg \)]" using \ iso_lunit by simp ultimately have "trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>1 (trg \)" using isomorphic_def by auto also have "tab\<^sub>1 (trg \) \ tab\<^sub>0 (trg \)" using \ trg.obj_has_symmetric_tab isomorphic_symmetric by auto finally show ?thesis by blast qed hence "trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>1 (trg \)" using \ trg.obj_has_symmetric_tab isomorphic_transitive by blast have "\tab\<^sub>1 (trg \)\ \ Maps.Hom (src (tab\<^sub>0 (trg \))) (trg \)" proof - have "\tab\<^sub>1 (trg \) : src (tab\<^sub>0 (trg \)) \ trg \\ \ is_left_adjoint (tab\<^sub>0 (trg \)) \ \tab\<^sub>0 (trg \)\ = \tab\<^sub>0 (trg \)\" using \ by simp thus ?thesis by auto qed have "\trg \\ \ Maps.Hom (trg \) (trg \)" proof - have "\trg \ : trg \ \ trg \\ \ is_left_adjoint (trg \) \ \trg \\ = \trg \\" using \ obj_is_self_adjoint by simp thus ?thesis by auto qed interpret SPN_trg: arrow_of_spans Maps.comp \SPN (trg \)\ using \ SPN.preserves_reflects_arr Span.arr_char by blast have SPN_trg: "SPN (trg \) = \Chn = Maps.MkArr (src (tab\<^sub>1 (trg \))) (src (tab\<^sub>1 (trg \))) \spn (trg \)\, Dom = \Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\, Cod = \Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\\" unfolding SPN_def using \ by simp interpret trg_SPN: arrow_of_spans Maps.comp \Span.trg (SPN \)\ using \ SPN.preserves_reflects_arr Span.arr_char by blast have trg_SPN: "Span.trg (SPN \) = \Chn = ?trg, Dom = \Leg0 = ?trg, Leg1 = ?trg\, Cod = \Leg0 = ?trg, Leg1 = ?trg\\" proof - let ?tab\<^sub>1_dom = "Maps.MkArr (src (tab\<^sub>1 (dom \))) (trg \) \tab\<^sub>1 (dom \)\" have "Maps.arr ?tab\<^sub>1_dom" proof - have "\tab\<^sub>1 (dom \) : src (tab\<^sub>1 (dom \)) \ trg \\ \ is_left_adjoint (tab\<^sub>1 (dom \)) \ \tab\<^sub>1 (dom \)\ = \tab\<^sub>1 (dom \)\" using \ by simp thus ?thesis using \ Maps.arr_MkArr by blast qed thus ?thesis using \ Maps.cod_char Span.trg_def by simp qed let ?\ = "\Chn = ?tab\<^sub>1, Dom = Dom (SPN (trg \)), Cod = Cod (Span.trg (SPN \))\" interpret \: arrow_of_spans Maps.comp ?\ apply (unfold_locales, simp_all) proof - show "Maps.in_hom ?tab\<^sub>1 SPN_trg.dom.apex trg_SPN.cod.apex" using \ tab\<^sub>0 tab\<^sub>1 Maps.dom_char Maps.cod_char SPN_trg trg_SPN tab.apex_def trg_SPN.cod.apex_def apply (intro Maps.in_homI) by simp_all (* The preceding cannot be written "by (intro Maps.in_homI, simp_all)", why? *) show "Maps.comp trg_SPN.cod.leg0 ?tab\<^sub>1 = SPN_trg.dom.leg0" proof - have "Maps.seq trg_SPN.cod.leg0 ?tab\<^sub>1" using \ trg_SPN tab\<^sub>1 Maps.dom_char Maps.cod_char by (intro Maps.seqI, auto) moreover have "Maps.Comp \trg \\ \tab\<^sub>1 (trg \)\ = \tab\<^sub>1 (trg \)\" proof - have "tab\<^sub>1 (trg \) \ Maps.Comp \trg \\ \tab\<^sub>1 (trg \)\" using \ is_iso_classI ide_in_iso_class [of "trg \"] ide_in_iso_class [of "tab\<^sub>1 (trg \)"] \trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>1 (trg \)\ by auto thus ?thesis using Maps.Comp_eq_iso_class_memb \iso_class (tab\<^sub>1 (trg \)) \ Maps.Hom (src (tab\<^sub>0 (trg \))) (trg \)\ \iso_class (trg \) \ Maps.Hom (trg \) (trg \)\ by meson qed moreover have "\tab\<^sub>1 (trg \)\ = \tab\<^sub>0 (trg \)\" using \ iso_class_eqI trg.obj_has_symmetric_tab by auto ultimately show ?thesis using \ Maps.comp_char [of trg_SPN.cod.leg0 ?tab\<^sub>1] trg_SPN by simp qed show "trg_SPN.cod.leg1 \ ?tab\<^sub>1 = SPN_trg.dom.leg1" proof - have "Maps.seq trg_SPN.cod.leg1 ?tab\<^sub>1" using \ trg_SPN tab\<^sub>1 Maps.dom_char Maps.cod_char by (intro Maps.seqI, auto) moreover have "Maps.Comp \trg \\ \tab\<^sub>1 (trg \)\ = \tab\<^sub>1 (trg \)\" proof - have "tab\<^sub>1 (trg \) \ Maps.Comp \trg \\ \tab\<^sub>1 (trg \)\" using \ is_iso_classI ide_in_iso_class [of "trg \"] ide_in_iso_class [of "tab\<^sub>1 (trg \)"] \trg \ \ tab\<^sub>1 (trg \) \ tab\<^sub>1 (trg \)\ by auto thus ?thesis using Maps.Comp_eq_iso_class_memb \\tab\<^sub>1 (trg \)\ \ Maps.Hom (src (tab\<^sub>0 (trg \))) (trg \)\ \\trg \\ \ Maps.Hom (trg \) (trg \)\ by meson qed ultimately show ?thesis using \ Maps.comp_char [of trg_SPN.cod.leg1 ?tab\<^sub>1] trg_SPN by simp qed qed have \: "Span.in_hom ?\ (SPN (trg \)) (Span.trg (SPN \))" using \ tab\<^sub>0 spn_ide [of "trg \"] iso_class_eqI Span.arr_char Span.dom_char Span.cod_char \.arrow_of_spans_axioms SPN_trg trg_SPN trg.apex_def tab.apex_def Maps.dom_char apply (intro Span.in_homI) by auto have "Maps.iso ?tab\<^sub>1" proof - have "Maps.iso ?tab\<^sub>0" using \ tab\<^sub>0 ide_in_iso_class trg.is_map_iff_tab\<^sub>0_is_equivalence obj_is_self_adjoint Maps.iso_char' [of ?tab\<^sub>0] by auto moreover have "?tab\<^sub>0 = ?tab\<^sub>1" proof - have "\tab\<^sub>0 (trg \)\ = \tab\<^sub>1 (trg \)\" using \ iso_class_eqI trg.obj_has_symmetric_tab by auto thus ?thesis by simp qed ultimately show ?thesis by simp qed thus "Span.isomorphic (SPN (trg \)) (Span.trg (SPN \))" using \ Span.isomorphic_def Span.iso_char by auto qed qed lemma SPN_is_weak_arrow_of_homs: shows "weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN" .. end subsubsection "Compositors" text \ To complete the proof that \SPN\ is a pseudofunctor, we need to obtain a natural isomorphism \\\, whose component at \(r, s)\ is an isomorphism \\ (r, s)\ from the horizontal composite \SPN r \ SPN s\ to \SPN (r \ s)\ in \Span(Maps(B))\, and we need to prove that the coherence conditions are satisfied. We have shown that the tabulations of \r\ and \s\ compose to yield a tabulation of \r \ s\. Since tabulations of the same arrow are equivalent, this tabulation must be equivalent to the chosen tabulation of \r \ s\. We therefore obtain an equivalence map from the apex of \SPN r \ SPN s\ to the apex of \SPN (r \ s)\ which commutes with the legs of these spans up to isomorphism. This equivalence map determines an invertible arrow in \Span(Maps(B))\. Moreover, by property \T2\, any two such equivalence maps are connected by a unique 2-cell, which is consequently an isomorphism. This shows that the arrow in \Span(Maps(B))\ is uniquely determined, which fact we can exploit to establish the required coherence conditions. \ locale two_composable_identities_in_bicategory_of_spans = bicategory_of_spans V H \ \ src trg + r: identity_in_bicategory_of_spans V H \ \ src trg r + s: identity_in_bicategory_of_spans V H \ \ src trg s for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and r :: 'a and s :: 'a + assumes composable: "src r = trg s" begin notation isomorphic (infix "\" 50) interpretation r: arrow_in_bicategory_of_spans V H \ \ src trg r r r by unfold_locales auto interpretation r: arrow_of_tabulations_in_maps V H \ \ src trg r r.tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ r r.tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ r using r.is_arrow_of_tabulations_in_maps by simp interpretation s: arrow_in_bicategory_of_spans V H \ \ src trg s s s by unfold_locales auto interpretation s: arrow_of_tabulations_in_maps V H \ \ src trg s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 s\ s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 s\ s using s.is_arrow_of_tabulations_in_maps by simp sublocale identity_in_bicategory_of_spans V H \ \ src trg \r \ s\ apply unfold_locales by (simp add: composable) sublocale horizontal_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg r r.tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 s\ r r.tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 s\ r s using composable by unfold_locales auto abbreviation p\<^sub>0 where "p\<^sub>0 \ \\.p\<^sub>0" abbreviation p\<^sub>1 where "p\<^sub>1 \ \\.p\<^sub>1" text \ We will take as the composition isomorphism from \SPN r \ SPN s\ to \SPN (r \ s)\ the arrow of tabulations, induced by the identity \r \ s\, from the composite of the chosen tabulations of \r\ and \s\ to the chosen tabulation of \r \ s\. As this arrow of tabulations is induced by an identity, it is an equivalence map. \ interpretation cmp: identity_arrow_of_tabulations_in_maps V H \ \ src trg \r \ s\ \\.tab \tab\<^sub>0 s \ \\.p\<^sub>0\ \tab\<^sub>1 r \ \\.p\<^sub>1\ \r \ s\ tab \tab\<^sub>0 (r \ s)\ \tab\<^sub>1 (r \ s)\ \r \ s\ using composable by unfold_locales auto lemma cmp_interpretation: shows "identity_arrow_of_tabulations_in_maps V H \ \ src trg (r \ s) \\.tab (tab\<^sub>0 s \ \\.p\<^sub>0) (tab\<^sub>1 r \ \\.p\<^sub>1) (r \ s) tab (tab\<^sub>0 (r \ s)) (tab\<^sub>1 (r \ s)) (r \ s)" .. definition cmp where "cmp = cmp.chine" lemma cmp_props: shows "\cmp : src \\.tab \ src tab\" and "\cmp : cmp \ cmp\" and "equivalence_map cmp" and "tab\<^sub>0 (r \ s) \ cmp \ tab\<^sub>0 s \ \\.p\<^sub>0" and "tab\<^sub>1 (r \ s) \ cmp \ tab\<^sub>1 r \ \\.p\<^sub>1" using cmp_def cmp.leg0_uniquely_isomorphic(1) cmp.leg1_uniquely_isomorphic(1) isomorphic_symmetric cmp.chine_is_equivalence by auto lemma cmp_in_hom [intro]: shows "\cmp : src \\.tab \ src tab\" and "\cmp : cmp \ cmp\" using cmp_props by auto lemma cmp_simps [simp]: shows "arr cmp" and "ide cmp" and "src cmp = src \\.tab" and "trg cmp = src tab" and "dom cmp = cmp" and "cod cmp = cmp" using cmp_props equivalence_map_is_ide by auto text \ Now we have to use the above properties of the underlying bicategory to exhibit the composition isomorphisms as actual arrows in \Span(Maps(B))\ and to prove the required naturality and coherence conditions. \ interpretation Maps: maps_category V H \ \ src trg .. interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. no_notation Fun.comp (infixl "\" 55) notation Span.vcomp (infixr "\" 55) notation Span.hcomp (infixr "\" 53) notation Maps.comp (infixr "\" 55) interpretation SPN: "functor" V Span.vcomp SPN using SPN_is_functor by simp interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN using SPN_is_weak_arrow_of_homs by simp interpretation SPN_r_SPN_s: arrow_of_spans Maps.comp \SPN r \ SPN s\ using composable Span.ide_char [of "SPN r \ SPN s"] by simp interpretation SPN_r_SPN_s: identity_arrow_of_spans Maps.comp \SPN r \ SPN s\ using composable Span.ide_char [of "SPN r \ SPN s"] by (unfold_locales, simp) interpretation SPN_rs: arrow_of_spans Maps.comp \SPN (r \ s)\ using composable Span.arr_char r.base_simps(2) s.base_simps(2) by blast interpretation SPN_rs: identity_arrow_of_spans Maps.comp \SPN (r \ s)\ using composable Span.ide_char SPN.preserves_ide r.is_ide s.is_ide by (unfold_locales, simp) text \ The following are the legs (as arrows of \Maps\) of the spans \SPN r\ and \SPN s\. \ definition R\<^sub>0 where "R\<^sub>0 = \\tab\<^sub>0 r\\" definition R\<^sub>1 where "R\<^sub>1 = \\tab\<^sub>1 r\\" definition S\<^sub>0 where "S\<^sub>0 = \\tab\<^sub>0 s\\" definition S\<^sub>1 where "S\<^sub>1 = \\tab\<^sub>1 s\\" lemma span_legs_eq: shows "Leg0 (Dom (SPN r)) = R\<^sub>0" and "Leg1 (Dom (SPN r)) = R\<^sub>1" and "Leg0 (Dom (SPN s)) = S\<^sub>0" and "Leg1 (Dom (SPN s)) = S\<^sub>1" using SPN_def R\<^sub>0_def R\<^sub>1_def S\<^sub>0_def S\<^sub>1_def composable by auto lemma R\<^sub>0_in_hom [intro]: shows "Maps.in_hom R\<^sub>0 (Maps.MkIde (src r.s\<^sub>0)) (Maps.MkIde (src r))" by (simp add: Maps.MkArr_in_hom' R\<^sub>0_def) lemma R\<^sub>1_in_hom [intro]: shows "Maps.in_hom R\<^sub>1 (Maps.MkIde (src r.s\<^sub>0)) (Maps.MkIde (trg r))" by (simp add: Maps.MkArr_in_hom' R\<^sub>1_def) lemma S\<^sub>0_in_hom [intro]: shows "Maps.in_hom S\<^sub>0 (Maps.MkIde (src s.s\<^sub>0)) (Maps.MkIde (src s))" by (simp add: Maps.MkArr_in_hom' S\<^sub>0_def) lemma S\<^sub>1_in_hom [intro]: shows "Maps.in_hom S\<^sub>1 (Maps.MkIde (src s.s\<^sub>0)) (Maps.MkIde (trg s))" by (simp add: Maps.MkArr_in_hom' S\<^sub>1_def) lemma RS_simps [simp]: shows "Maps.arr R\<^sub>0" and "Maps.dom R\<^sub>0 = Maps.MkIde (src r.s\<^sub>0)" and "Maps.cod R\<^sub>0 = Maps.MkIde (src r)" and "Maps.Dom R\<^sub>0 = src r.s\<^sub>0" and "Maps.Cod R\<^sub>0 = src r" and "Maps.arr R\<^sub>1" and "Maps.dom R\<^sub>1 = Maps.MkIde (src r.s\<^sub>0)" and "Maps.cod R\<^sub>1 = Maps.MkIde (trg r)" and "Maps.Dom R\<^sub>1 = src r.s\<^sub>0" and "Maps.Cod R\<^sub>1 = trg r" and "Maps.arr S\<^sub>0" and "Maps.dom S\<^sub>0 = Maps.MkIde (src s.s\<^sub>0)" and "Maps.cod S\<^sub>0 = Maps.MkIde (src s)" and "Maps.Dom S\<^sub>0 = src s.s\<^sub>0" and "Maps.Cod S\<^sub>0 = src s" and "Maps.arr S\<^sub>1" and "Maps.dom S\<^sub>1 = Maps.MkIde (src s.s\<^sub>0)" and "Maps.cod S\<^sub>1 = Maps.MkIde (trg s)" and "Maps.Dom S\<^sub>1 = src s.s\<^sub>0" and "Maps.Cod S\<^sub>1 = trg s" using R\<^sub>0_in_hom R\<^sub>1_in_hom S\<^sub>0_in_hom S\<^sub>1_in_hom composable by (auto simp add: R\<^sub>0_def R\<^sub>1_def S\<^sub>0_def S\<^sub>1_def) text \ The apex of the composite span @{term "SPN r \ SPN s"} (defined in terms of pullback) coincides with the apex of the composite tabulation \\\\ (defined using the chosen tabulation of \(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 s)\). We need this to be true in order to define the compositor of a pseudofunctor from the underlying bicategory \B\ to \Span(Maps(B))\. It is only true if we have carefully chosen pullbacks in \Maps(B)\ in order to ensure the relationship with the chosen tabulations. \ lemma SPN_r_SPN_s_apex_eq: shows "SPN_r_SPN_s.apex = Maps.MkIde (src \\.tab)" proof - have "obj (Maps.Cod SPN_r_SPN_s.leg0)" using Maps.arr_char [of "SPN_r_SPN_s.leg0"] by simp moreover have "obj (Maps.Dom SPN_r_SPN_s.leg0)" using Maps.arr_char [of "SPN_r_SPN_s.leg0"] by simp moreover have "SPN_r_SPN_s.leg0 \ Maps.Null" using Maps.arr_char [of "SPN_r_SPN_s.leg0"] by simp moreover have "Maps.Dom SPN_r_SPN_s.leg0 = src \\.tab" proof - interpret REP_S\<^sub>1: map_in_bicategory V H \ \ src trg \Maps.REP S\<^sub>1\ using Maps.REP_in_Map Maps.arr_char Maps.in_HomD S\<^sub>1_def apply unfold_locales by (meson Maps.REP_in_hhom(2) S\<^sub>1_in_hom) interpret REP_R\<^sub>0: map_in_bicategory V H \ \ src trg \Maps.REP R\<^sub>0\ using Maps.REP_in_Map Maps.arr_char Maps.in_HomD R\<^sub>0_def apply unfold_locales by (meson Maps.REP_in_hhom(2) R\<^sub>0_in_hom) have "Maps.Dom SPN_r_SPN_s.leg0 = Maps.Dom (S\<^sub>0 \ Maps.PRJ\<^sub>0 R\<^sub>0 S\<^sub>1)" using composable Span.hcomp_def S\<^sub>0_def R\<^sub>0_def S\<^sub>1_def by simp also have "... = Maps.Dom \\tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)\\" proof - have "is_left_adjoint (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1))" proof - have "ide ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)" proof - have "src (Maps.REP R\<^sub>0)\<^sup>* = trg (Maps.REP S\<^sub>1)" using REP_R\<^sub>0.is_map REP_S\<^sub>1.is_map left_adjoint_is_ide R\<^sub>0_def S\<^sub>1_def by (metis (no_types, lifting) Maps.REP_CLS REP_R\<^sub>0.antipar(2) isomorphic_implies_hpar(4) composable r.leg0_simps(3) r.satisfies_T0 s.leg1_is_map s.leg1_simps(3) s.leg1_simps(4)) thus ?thesis by simp qed thus ?thesis by simp qed moreover have "Maps.Dom (S\<^sub>0 \ \\tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)\\) = src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1))" proof - have "Maps.arr (\\prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\\)" using Maps.CLS_in_hom Maps.prj0_simps(1) Maps.PRJ\<^sub>0_def composable by fastforce moreover have "Maps.Dom S\<^sub>0 = Maps.Cod \\prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\\" proof - have "Maps.Cod \\prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\\ = trg (prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0))" by simp also have "... = src (Maps.REP S\<^sub>1)" proof - have "ide (Maps.REP S\<^sub>1)" by simp moreover have "is_left_adjoint (Maps.REP R\<^sub>0)" by auto moreover have "trg (Maps.REP S\<^sub>1) = trg (Maps.REP R\<^sub>0)" by (simp add: composable) ultimately show ?thesis using S\<^sub>1_def Maps.REP_CLS r.leg0_is_map s.leg1_is_map by simp qed also have "... = src (tab\<^sub>0 s)" using tab\<^sub>0_in_hom(1) by simp also have "... = Maps.Dom S\<^sub>0" using S\<^sub>0_def by simp finally show ?thesis by simp qed ultimately have "Maps.Dom (S\<^sub>0 \ \\tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)\\) = Maps.Dom \\tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)\\" using Maps.CLS_in_hom by simp thus ?thesis by simp qed ultimately show ?thesis using Maps.PRJ\<^sub>0_def composable Maps.Dom.simps(1) RS_simps(1) RS_simps(16) RS_simps(18) RS_simps(3) by presburger qed also have "... = src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1))" by simp finally have "Maps.Dom SPN_r_SPN_s.leg0 = src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1))" by simp also have "... = src (tab\<^sub>0 (r.s\<^sub>0\<^sup>* \ s.s\<^sub>1))" proof - interpret r\<^sub>0's\<^sub>1: identity_in_bicategory_of_spans V H \ \ src trg \r.s\<^sub>0\<^sup>* \ s.s\<^sub>1\ using composable by (unfold_locales, simp) have "(Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1 \ r.s\<^sub>0\<^sup>* \ s.s\<^sub>1" proof - have "(Maps.REP R\<^sub>0)\<^sup>* \ r.s\<^sub>0\<^sup>*" proof - have 1: "adjoint_pair (Maps.REP R\<^sub>0) (Maps.REP R\<^sub>0)\<^sup>*" using REP_R\<^sub>0.is_map left_adjoint_extends_to_adjoint_pair by blast moreover have "adjoint_pair r.s\<^sub>0 (Maps.REP R\<^sub>0)\<^sup>*" proof - have "Maps.REP R\<^sub>0 \ r.s\<^sub>0" unfolding R\<^sub>0_def using Maps.REP_CLS r.leg0_is_map composable by force thus ?thesis using 1 adjoint_pair_preserved_by_iso isomorphic_def REP_R\<^sub>0.triangle_in_hom(4) REP_R\<^sub>0.triangle_right' by auto qed ultimately show ?thesis using r.leg0_is_map left_adjoint_determines_right_up_to_iso left_adjoint_extends_to_adjoint_pair by auto qed moreover have "Maps.REP S\<^sub>1 \ s.s\<^sub>1" unfolding S\<^sub>1_def using Maps.REP_CLS s.leg1_is_map composable by force moreover have "\a. a \ (tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 s \ (Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1 \ a" using calculation composable isomorphic_implies_hpar(3) hcomp_ide_isomorphic hcomp_isomorphic_ide [of "(Maps.REP R\<^sub>0)\<^sup>*" "r.s\<^sub>0\<^sup>*" s.s\<^sub>1] by auto ultimately show ?thesis using isomorphic_transitive by blast qed thus ?thesis using r\<^sub>0's\<^sub>1.isomorphic_implies_same_tab isomorphic_symmetric by metis qed also have "... = src \\.tab" using VV.ide_char VV.arr_char composable Span.hcomp_def \\.tab_simps(2) by auto finally show ?thesis by simp qed ultimately show ?thesis using composable Maps.arr_char Maps.dom_char SPN_r_SPN_s.dom.apex_def apply auto by (metis (no_types, lifting) Maps.not_arr_null SPN_r_SPN_s.chine_eq_apex SPN_r_SPN_s.chine_simps(1)) qed text \ We will be taking the arrow @{term "CLS cmp"} of \Maps\ as the composition isomorphism from @{term "SPN r \ SPN s"} to @{term "SPN (r \ s)"}. The following result shows that it has the right domain and codomain for that purpose. \ lemma iso_class_cmp_in_hom: shows "Maps.in_hom (Maps.MkArr (src \\.tab) (src tab) \cmp\) SPN_r_SPN_s.apex SPN_rs.apex" and "Maps.in_hom \\cmp\\ SPN_r_SPN_s.apex SPN_rs.apex" proof - show "Maps.in_hom (Maps.MkArr (src \\.tab) (src tab) \cmp\) SPN_r_SPN_s.apex SPN_rs.apex" proof - have "obj (src \\.tab)" using obj_src \\.tab_in_hom by blast moreover have "obj (src tab)" using obj_src by simp moreover have "\cmp\ \ Maps.Hom (src \\.tab) (src tab)" by (metis (mono_tags, lifting) cmp.is_map cmp_def cmp_props(1) mem_Collect_eq) moreover have "SPN_r_SPN_s.apex = Maps.MkIde (src \\.tab)" using SPN_r_SPN_s_apex_eq by simp moreover have "SPN_rs.apex = Maps.MkIde (src tab)" using SPN_def composable SPN_rs.cod.apex_def Maps.arr_char Maps.dom_char SPN_rs.dom.leg_simps(1) by fastforce ultimately show ?thesis using Maps.MkArr_in_hom by simp qed thus "Maps.in_hom \\cmp\\ SPN_r_SPN_s.apex SPN_rs.apex" by simp qed interpretation r\<^sub>0's\<^sub>1: two_composable_identities_in_bicategory_of_spans V H \ \ src trg \(Maps.REP R\<^sub>0)\<^sup>*\ \Maps.REP S\<^sub>1\ proof show "ide (Maps.REP S\<^sub>1)" using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide by (meson Maps.REP_in_hhom(2) S\<^sub>1_in_hom) show "ide (Maps.REP R\<^sub>0)\<^sup>*" using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide Maps.REP_in_hhom(2) R\<^sub>0_in_hom by auto show "src (Maps.REP R\<^sub>0)\<^sup>* = trg (Maps.REP S\<^sub>1)" using Maps.REP_in_hhom(2) R\<^sub>0_in_hom composable by auto qed interpretation R\<^sub>0'S\<^sub>1: identity_in_bicategory_of_spans V H \ \ src trg \(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 s\ by (unfold_locales, simp add: composable) lemma prj_tab_agreement: shows "(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 s \ (Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1" and "\\.p\<^sub>0 \ prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)" and "\\.p\<^sub>1 \ prj\<^sub>1 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)" proof - show 1: "(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 s \ (Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1" proof - have "(tab\<^sub>0 r)\<^sup>* \ (Maps.REP R\<^sub>0)\<^sup>*" using Maps.REP_CLS isomorphic_symmetric R\<^sub>0_def composable r.satisfies_T0 isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint by fastforce moreover have "tab\<^sub>1 s \ Maps.REP S\<^sub>1" by (metis Maps.REP_CLS isomorphic_symmetric S\<^sub>1_def s.leg1_is_map s.leg1_simps(3-4)) moreover have "src (Maps.REP R\<^sub>0)\<^sup>* = trg (tab\<^sub>1 s)" using composable r.T0.antipar right_adjoint_simps(2) by fastforce ultimately show ?thesis using hcomp_isomorphic_ide [of "(tab\<^sub>0 r)\<^sup>*" "(Maps.REP R\<^sub>0)\<^sup>*" "tab\<^sub>1 s"] hcomp_ide_isomorphic isomorphic_transitive composable by auto qed show "\\.p\<^sub>0 \ tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)" using 1 R\<^sub>0'S\<^sub>1.isomorphic_implies_same_tab isomorphic_reflexive by auto show "\\.p\<^sub>1 \ tab\<^sub>1 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)" using 1 R\<^sub>0'S\<^sub>1.isomorphic_implies_same_tab isomorphic_reflexive by auto qed lemma chine_hcomp_SPN_SPN: shows "Span.chine_hcomp (SPN r) (SPN s) = Maps.MkIde (src \\.p\<^sub>0)" proof - have "Span.chine_hcomp (SPN r) (SPN s) = Maps.MkIde (src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \ Maps.REP S\<^sub>1)))" using Span.chine_hcomp_ide_ide [of "SPN r" "SPN s"] composable Maps.pbdom_def Maps.PRJ\<^sub>0_def Maps.CLS_in_hom Maps.dom_char R\<^sub>0_def S\<^sub>1_def apply simp using Maps.prj0_simps(1) RS_simps(1) RS_simps(16) RS_simps(18) RS_simps(3) by presburger also have "... = Maps.MkIde (src \\.p\<^sub>0)" using prj_tab_agreement isomorphic_implies_hpar(3) by force finally show ?thesis by simp qed end text \ The development above focused on two specific composable 1-cells in bicategory \B\. Here we reformulate those results as statements about the entire bicategory. \ context bicategory_of_spans begin interpretation Maps: maps_category V H \ \ src trg .. interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. no_notation Fun.comp (infixl "\" 55) notation Span.vcomp (infixr "\" 55) notation Span.hcomp (infixr "\" 53) notation Maps.comp (infixr "\" 55) notation isomorphic (infix "\" 50) interpretation SPN: "functor" V Span.vcomp SPN using SPN_is_functor by simp interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN using SPN_is_weak_arrow_of_homs by simp interpretation HoSPN_SPN: composite_functor VV.comp Span.VV.comp Span.vcomp SPN.FF \\\\. fst \\ \ snd \\\ .. interpretation SPNoH: composite_functor VV.comp V Span.vcomp \\\\. fst \\ \ snd \\\ SPN .. text \ Given arbitrary composable 1-cells \r\ and \s\, obtain an arrow of spans in \Maps\ having the isomorphism class of \rs.cmp\ as its chine. \ definition CMP where "CMP r s \ \Chn = \\two_composable_identities_in_bicategory_of_spans.cmp V H \ \ src trg r s\\, Dom = Dom (SPN r \ SPN s), Cod = Dom (SPN (r \ s))\" lemma compositor_in_hom [intro]: assumes "ide r" and "ide s" and "src r = trg s" shows "Span.in_hhom (CMP r s) (SPN.map\<^sub>0 (src s)) (SPN.map\<^sub>0 (trg r))" and "Span.in_hom (CMP r s) (HoSPN_SPN.map (r, s)) (SPNoH.map (r, s))" proof - have rs: "VV.ide (r, s)" using assms VV.ide_char VV.arr_char by simp interpret rs: two_composable_identities_in_bicategory_of_spans V H \ \ src trg r s using rs VV.ide_char VV.arr_char by unfold_locales auto interpret cmp: arrow_of_tabulations_in_maps V H \ \ src trg \r \ s\ rs.\\.tab \tab\<^sub>0 s \ rs.\\.p\<^sub>0\ \tab\<^sub>1 r \ rs.\\.p\<^sub>1\ \r \ s\ rs.tab \tab\<^sub>0 (r \ s)\ \tab\<^sub>1 (r \ s)\ \r \ s\ by unfold_locales auto have "rs.cmp = cmp.chine" using rs.cmp_def by simp interpret SPN_r_SPN_s: arrow_of_spans Maps.comp \SPN r \ SPN s\ using rs.composable Span.ide_char [of "SPN r \ SPN s"] by simp interpret SPN_r_SPN_s: identity_arrow_of_spans Maps.comp \SPN r \ SPN s\ using rs.composable Span.ide_char [of "SPN r \ SPN s"] by (unfold_locales, simp) interpret SPN_rs: arrow_of_spans Maps.comp \SPN (r \ s)\ using Span.arr_char rs.is_ide SPN.preserves_arr by blast interpret SPN_rs: identity_arrow_of_spans Maps.comp \SPN (r \ s)\ using Span.ide_char rs.is_ide SPN.preserves_ide by (unfold_locales, simp) interpret Dom: span_in_category Maps.comp \Dom (CMP r s)\ by (unfold_locales, simp add: CMP_def) interpret Cod: span_in_category Maps.comp \Cod (CMP r s)\ proof - (* TODO: I don't understand what makes this so difficult. *) have "\tab\<^sub>0 (r \ s) : src (tab\<^sub>0 (r \ s)) \ src s\ \ is_left_adjoint (tab\<^sub>0 (r \ s)) \ \tab\<^sub>0 (r \ s)\ = \tab\<^sub>0 (r \ s)\" by simp hence "\f. \f : src (tab\<^sub>0 (r \ s)) \ src s\ \ is_left_adjoint f \ \tab\<^sub>0 (r \ s)\ = \f\" by blast moreover have "\f. \f : src (tab\<^sub>0 (r \ s)) \ trg r\ \ is_left_adjoint f \ \tab\<^sub>1 (r \ s)\ = \f\" by (metis rs.base_simps(2) rs.leg1_in_hom(1) rs.leg1_is_map trg_hcomp) ultimately show "span_in_category Maps.comp (Cod (CMP r s))" using assms Maps.arr_char Maps.dom_char CMP_def by unfold_locales auto qed interpret r\<^sub>0's\<^sub>1: two_composable_identities_in_bicategory_of_spans V H \ \ src trg \(Maps.REP rs.R\<^sub>0)\<^sup>*\ \Maps.REP rs.S\<^sub>1\ proof show "ide (Maps.REP rs.S\<^sub>1)" using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide by (meson Maps.REP_in_hhom(2) rs.S\<^sub>1_in_hom) show "ide (Maps.REP rs.R\<^sub>0)\<^sup>*" using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide Maps.REP_in_hhom(2) rs.R\<^sub>0_in_hom by auto show "src (Maps.REP rs.R\<^sub>0)\<^sup>* = trg (Maps.REP rs.S\<^sub>1)" using Maps.REP_in_hhom(2) rs.R\<^sub>0_in_hom rs.composable by auto qed interpret R\<^sub>0'S\<^sub>1: identity_in_bicategory_of_spans V H \ \ src trg \(tab\<^sub>0 r)\<^sup>* \ tab\<^sub>1 s\ by (unfold_locales, simp add: rs.composable) text \ Here we obtain explicit formulas for the legs and apex of \SPN_r_SPN_s\ and \SPN_rs\. \ have SPN_r_SPN_s_leg0_eq: "SPN_r_SPN_s.leg0 = Maps.comp rs.S\<^sub>0 (Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1)" using rs.composable Span.hcomp_def rs.S\<^sub>0_def rs.R\<^sub>0_def rs.S\<^sub>1_def by simp have SPN_r_SPN_s_leg1_eq: "SPN_r_SPN_s.leg1 = Maps.comp rs.R\<^sub>1 (Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1)" using rs.composable Span.hcomp_def rs.R\<^sub>1_def rs.R\<^sub>0_def rs.S\<^sub>1_def by simp have "SPN_r_SPN_s.apex = Maps.MkIde (src rs.\\.tab)" using rs.SPN_r_SPN_s_apex_eq by auto have SPN_rs_leg0_eq: "SPN_rs.leg0 = \\tab\<^sub>0 (r \ s)\\" unfolding SPN_def using rs by simp have SPN_rs_leg1_eq: "SPN_rs.leg1 = \\tab\<^sub>1 (r \ s)\\" unfolding SPN_def using rs by simp have "SPN_rs.apex = Maps.MkIde (src (tab_of_ide (r \ s)))" using SPN_rs.dom.apex_def Maps.dom_char SPN_rs_leg0_eq SPN_rs.dom.leg_simps(1) by simp text \ The composition isomorphism @{term "CMP r s"} is an arrow of spans in \Maps(B)\. \ interpret arrow_of_spans Maps.comp \CMP r s\ proof show 1: "Maps.in_hom (Chn (CMP r s)) Dom.apex Cod.apex" using rs.iso_class_cmp_in_hom rs.composable CMP_def by simp show "Cod.leg0 \ Chn (CMP r s) = Dom.leg0" proof (intro Maps.arr_eqI) show 2: "Maps.seq Cod.leg0 (Chn (CMP r s))" using 1 Maps.dom_char Maps.cod_char by blast show 3: "Maps.arr Dom.leg0" by simp show "Maps.Dom (Cod.leg0 \ Chn (CMP r s)) = Maps.Dom Dom.leg0" using 1 2 Maps.dom_char Maps.cod_char Maps.comp_char Dom.leg_in_hom Maps.in_hom_char Maps.seq_char by auto show "Maps.Cod (Cod.leg0 \ Chn (CMP r s)) = Maps.Cod Dom.leg0" using 2 3 Maps.comp_char [of Cod.leg0 "Chn (CMP r s)"] Dom.leg_simps Dom.apex_def Maps.dom_char SPN_r_SPN_s_leg0_eq Maps.comp_char [of rs.S\<^sub>0 "Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1"] CMP_def by simp show "Maps.Map (Cod.leg0 \ Chn (CMP r s)) = Maps.Map Dom.leg0" proof - have "Maps.Map (Cod.leg0 \ Chn (CMP r s)) = Maps.Comp \tab\<^sub>0 (r \ s)\ \rs.cmp\" using 1 2 Maps.dom_char Maps.cod_char Maps.comp_char [of Cod.leg0 "Chn (CMP r s)"] CMP_def by simp also have "... = Maps.Comp \tab\<^sub>0 s\ \rs.\\.p\<^sub>0\" proof - have "Maps.Comp \tab\<^sub>0 (r \ s)\ \rs.cmp\ = \tab\<^sub>0 (r \ s) \ rs.cmp\" using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp cmp.is_map rs.cmp_def by auto also have "... = Maps.Comp \tab\<^sub>0 s\ \rs.\\.p\<^sub>0\" using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp iso_class_eqI rs.cmp_props(4) by auto finally show ?thesis by simp qed also have "... = Maps.Map Dom.leg0" proof - have "Maps.seq rs.S\<^sub>0 (Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1)" by (intro Maps.seqI, simp_all add: rs.composable) moreover have "\prj\<^sub>0 (Maps.REP rs.S\<^sub>1) (Maps.REP rs.R\<^sub>0)\ = \rs.\\.p\<^sub>0\" using "rs.prj_tab_agreement" iso_class_eqI by auto moreover have "Maps.Dom (Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1) = src rs.\\.p\<^sub>0" using rs.prj_tab_agreement Maps.PRJ\<^sub>0_def rs.composable isomorphic_implies_hpar(3) by auto ultimately show ?thesis using SPN_r_SPN_s_leg0_eq Maps.comp_char [of rs.S\<^sub>0 "Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1"] rs.S\<^sub>0_def Maps.PRJ\<^sub>0_def rs.composable CMP_def by simp qed finally show ?thesis by simp qed qed show "Cod.leg1 \ Chn (CMP r s) = Dom.leg1" proof (intro Maps.arr_eqI) show 2: "Maps.seq Cod.leg1 (Chn (CMP r s))" using 1 Maps.dom_char Maps.cod_char by blast show 3: "Maps.arr Dom.leg1" by simp show "Maps.Dom (Cod.leg1 \ Chn (CMP r s)) = Maps.Dom Dom.leg1" using 1 2 Maps.dom_char Maps.cod_char Maps.comp_char Dom.leg_in_hom Maps.in_hom_char Maps.seq_char by auto show "Maps.Cod (Cod.leg1 \ Chn (CMP r s)) = Maps.Cod Dom.leg1" using 2 3 Maps.comp_char [of Cod.leg1 "Chn (CMP r s)"] Dom.apex_def Maps.dom_char SPN_r_SPN_s_leg1_eq Maps.comp_char [of rs.R\<^sub>1 "Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1"] CMP_def by simp show "Maps.Map (Cod.leg1 \ Chn (CMP r s)) = Maps.Map Dom.leg1" proof - have "Maps.Map (Cod.leg1 \ Chn (CMP r s)) = Maps.Comp \tab\<^sub>1 (r \ s)\ \rs.cmp\" using 1 2 Maps.dom_char Maps.cod_char Maps.comp_char [of Cod.leg1 "Chn (CMP r s)"] CMP_def by simp also have "... = Maps.Comp \tab\<^sub>1 r\ \rs.\\.p\<^sub>1\" proof - have "Maps.Comp \tab\<^sub>1 (r \ s)\ \rs.cmp\ = \tab\<^sub>1 (r \ s) \ rs.cmp\" using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp cmp.is_map rs.cmp_def by auto also have "... = Maps.Comp \tab\<^sub>1 r\ \rs.\\.p\<^sub>1\" using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp [of "tab\<^sub>1 r" rs.\\.p\<^sub>1] iso_class_eqI rs.cmp_props(5) by auto finally show ?thesis by simp qed also have "... = Maps.Map Dom.leg1" proof - have "Maps.seq rs.R\<^sub>1 (Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1)" by (intro Maps.seqI, simp_all add: rs.composable) moreover have "\prj\<^sub>1 (Maps.REP rs.S\<^sub>1) (Maps.REP rs.R\<^sub>0)\ = \rs.\\.p\<^sub>1\" using rs.prj_tab_agreement iso_class_eqI by auto moreover have "Maps.Dom (Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1) = src rs.\\.p\<^sub>1" using rs.prj_tab_agreement Maps.PRJ\<^sub>1_def rs.composable isomorphic_implies_hpar(3) by auto ultimately show ?thesis using SPN_r_SPN_s_leg1_eq Maps.comp_char [of rs.R\<^sub>1 "Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1"] rs.R\<^sub>1_def Maps.PRJ\<^sub>1_def rs.composable CMP_def by simp qed finally show ?thesis by simp (* * Very simple, right? Yeah, once you sort through the notational morass and * figure out what equals what. *) qed qed qed show "Span.in_hom (CMP r s) (HoSPN_SPN.map (r, s)) (SPNoH.map (r, s))" using Span.arr_char arrow_of_spans_axioms Span.dom_char Span.cod_char CMP_def SPN.FF_def VV.arr_char rs.composable by auto thus "Span.in_hhom (CMP r s) (SPN.map\<^sub>0 (src s)) (SPN.map\<^sub>0 (trg r))" using assms VV.ide_char VV.arr_char VV.in_hom_char SPN.FF_def apply (intro Span.in_hhomI) apply auto using Span.src_dom [of "CMP r s"] Span.trg_dom [of "CMP r s"] apply (elim Span.in_homE) apply auto using Span.src_dom [of "CMP r s"] Span.trg_dom [of "CMP r s"] apply (elim Span.in_homE) by auto qed lemma compositor_simps [simp]: assumes "ide r" and "ide s" and "src r = trg s" shows "Span.arr (CMP r s)" and "Span.src (CMP r s) = SPN.map\<^sub>0 (src s)" and "Span.trg (CMP r s) = SPN.map\<^sub>0 (trg r)" and "Span.dom (CMP r s) = HoSPN_SPN.map (r, s)" and "Span.cod (CMP r s) = SPNoH.map (r, s)" using assms compositor_in_hom [of r s] by auto lemma compositor_is_iso: assumes "ide r" and "ide s" and "src r = trg s" shows "Span.iso (CMP r s)" proof - interpret rs: two_composable_identities_in_bicategory_of_spans V H \ \ src trg r s using assms by unfold_locales auto have "Span.arr (CMP r s)" using assms compositor_in_hom by blast moreover have "Maps.iso \\rs.cmp\\" using assms Maps.iso_char' by (metis (mono_tags, lifting) Maps.CLS_in_hom Maps.Map.simps(1) Maps.in_homE equivalence_is_left_adjoint ide_in_iso_class rs.cmp_props(3) rs.cmp_simps(2)) ultimately show ?thesis unfolding CMP_def using assms Span.iso_char by simp qed interpretation \: transformation_by_components VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map \\rs. CMP (fst rs) (snd rs)\ proof fix rs assume rs: "VV.ide rs" let ?r = "fst rs" let ?s = "snd rs" show "Span.in_hom (CMP ?r ?s) (HoSPN_SPN.map rs) (SPNoH.map rs)" using rs compositor_in_hom [of ?r ?s] VV.ide_char VV.arr_char by simp next fix \\ assume \\: "VV.arr \\" let ?\ = "fst \\" let ?\ = "snd \\" show "CMP (fst (VV.cod \\)) (snd (VV.cod \\)) \ HoSPN_SPN.map \\ = SPNoH.map \\ \ CMP (fst (VV.dom \\)) (snd (VV.dom \\))" proof - let ?LHS = "CMP (fst (VV.cod \\)) (snd (VV.cod \\)) \ HoSPN_SPN.map \\" let ?RHS = "SPNoH.map \\ \ CMP (fst (VV.dom \\)) (snd (VV.dom \\))" have LHS: "Span.in_hom ?LHS (HoSPN_SPN.map (VV.dom \\)) (SPNoH.map (VV.cod \\))" proof show "Span.in_hom (HoSPN_SPN.map \\) (HoSPN_SPN.map (VV.dom \\)) (HoSPN_SPN.map (VV.cod \\))" using \\ by blast show "Span.in_hom (CMP (fst (VV.cod \\)) (snd (VV.cod \\))) (HoSPN_SPN.map (VV.cod \\)) (SPNoH.map (VV.cod \\))" using \\ VV.cod_simp by (auto simp add: VV.arr_char) qed have RHS: "Span.in_hom ?RHS (HoSPN_SPN.map (VV.dom \\)) (SPNoH.map (VV.cod \\))" using \\ VV.dom_simp VV.cod_simp by (auto simp add: VV.arr_char) show "?LHS = ?RHS" proof (intro Span.arr_eqI) show "Span.par ?LHS ?RHS" apply (intro conjI) using LHS RHS apply auto[2] proof - show "Span.dom ?LHS = Span.dom ?RHS" proof - have "Span.dom ?LHS = HoSPN_SPN.map (VV.dom \\)" using LHS by auto also have "... = Span.dom ?RHS" using RHS by auto finally show ?thesis by simp qed show "Span.cod ?LHS = Span.cod ?RHS" proof - have "Span.cod ?LHS = SPNoH.map (VV.cod \\)" using LHS by auto also have "... = Span.cod ?RHS" using RHS by auto finally show ?thesis by simp qed qed show "Chn ?LHS = Chn ?RHS" proof - interpret dom_\_\: two_composable_identities_in_bicategory_of_spans V H \ \ src trg \dom ?\\ \dom ?\\ using \\ VV.ide_char VV.arr_char by unfold_locales auto interpret cod_\_\: two_composable_identities_in_bicategory_of_spans V H \ \ src trg \cod ?\\ \cod ?\\ using \\ VV.ide_char VV.arr_char by unfold_locales auto interpret \_\: horizontal_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg \dom ?\\ \tab_of_ide (dom ?\)\ \tab\<^sub>0 (dom ?\)\ \tab\<^sub>1 (dom ?\)\ \dom ?\\ \tab_of_ide (dom ?\)\ \tab\<^sub>0 (dom ?\)\ \tab\<^sub>1 (dom ?\)\ \cod ?\\ \tab_of_ide (cod ?\)\ \tab\<^sub>0 (cod ?\)\ \tab\<^sub>1 (cod ?\)\ \cod ?\\ \tab_of_ide (cod ?\)\ \tab\<^sub>0 (cod ?\)\ \tab\<^sub>1 (cod ?\)\ ?\ ?\ using \\ VV.arr_char by unfold_locales auto let ?\\ = "?\ \ ?\" interpret dom_\\: identity_in_bicategory_of_spans V H \ \ src trg \dom ?\\\ using \\ by (unfold_locales, simp) interpret cod_\\: identity_in_bicategory_of_spans V H \ \ src trg \cod ?\\\ using \\ by (unfold_locales, simp) interpret \\: arrow_of_tabulations_in_maps V H \ \ src trg \dom ?\\\ \tab_of_ide (dom ?\\)\ \tab\<^sub>0 (dom ?\\)\ \tab\<^sub>1 (dom ?\\)\ \cod ?\\\ \tab_of_ide (cod ?\\)\ \tab\<^sub>0 (cod ?\\)\ \tab\<^sub>1 (cod ?\\)\ ?\\ using \\ by unfold_locales auto have Chn_LHS_eq: "Chn ?LHS = \\cod_\_\.cmp\\ \ Span.chine_hcomp (SPN (fst \\)) (SPN (snd \\))" proof - have "Chn ?LHS = Chn (CMP (fst (VV.cod \\)) (snd (VV.cod \\))) \ Chn (HoSPN_SPN.map \\)" using \\ LHS Span.Chn_vcomp by blast also have "... = \\cod_\_\.cmp\\ \ Chn (HoSPN_SPN.map \\)" using \\ VV.arr_char VV.cod_char CMP_def by simp also have "... = \\cod_\_\.cmp\\ \ Span.chine_hcomp (SPN (fst \\)) (SPN (snd \\))" using \\ VV.arr_char SPN.FF_def Span.hcomp_def by simp finally show ?thesis by blast qed have Chn_RHS_eq: "Chn ?RHS = Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\))) \\\.chine\ \ Maps.MkArr (src dom_\_\.\\.p\<^sub>0) (src (tab_of_ide (dom ?\ \ dom ?\))) \dom_\_\.cmp\" proof - have "Chn ?RHS = Chn (SPN (?\ \ ?\)) \ Maps.MkArr (src dom_\_\.\\.p\<^sub>0) (src (tab_of_ide (dom ?\ \ dom ?\))) \dom_\_\.cmp\" using \\ RHS Span.vcomp_def VV.arr_char CMP_def Span.arr_char Span.not_arr_Null VV.dom_simp by auto moreover have "Chn (SPN (?\ \ ?\)) = Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\))) \\\.chine\" proof - have "Chn (SPN (?\ \ ?\)) = Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\))) \spn ?\\\" using \\ SPN_def by simp also have "... = Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\))) \\\.chine\" using spn_def by simp finally show ?thesis by simp qed ultimately show ?thesis by simp qed let ?Chn_LHS = "Maps.MkArr (src cod_\_\.\\.p\<^sub>0) (src (tab_of_ide (cod ?\ \ cod ?\))) \cod_\_\.cmp\ \ Span.chine_hcomp (SPN ?\) (SPN ?\)" let ?Chn_RHS = "Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\))) \\\.chine\ \ Maps.MkArr (src dom_\_\.\\.p\<^sub>0) (src (tab_of_ide (dom ?\ \ dom ?\))) \dom_\_\.cmp\" have "?Chn_LHS = ?Chn_RHS" proof (intro Maps.arr_eqI) interpret LHS: arrow_of_spans Maps.comp ?LHS using LHS Span.arr_char by auto interpret RHS: arrow_of_spans Maps.comp ?RHS using RHS Span.arr_char by auto show 1: "Maps.arr ?Chn_LHS" using LHS.chine_in_hom Chn_LHS_eq by auto show 2: "Maps.arr ?Chn_RHS" using RHS.chine_in_hom Chn_RHS_eq by auto text \ Here is where we use \dom_\_\.chine_hcomp_SPN_SPN\, which depends on our having chosen the ``right'' pullbacks for \Maps(B)\. The map \Chn_LHS\ has as its domain the apex of the horizontal composite of the components of @{term "VV.dom \\"}, whereas \Chn_RHS\ has as its domain the apex of the chosen tabulation of \r\<^sub>0\<^sup>* \ s\<^sub>1\. We need these to be equal in order for \Chn_LHS\ and \Chn_RHS\ to be equal. \ show "Maps.Dom ?Chn_LHS = Maps.Dom ?Chn_RHS" proof - have "Maps.Dom ?Chn_LHS = Maps.Dom (Maps.dom ?Chn_LHS)" using \\ 1 Maps.Dom_dom by presburger also have "... = Maps.Dom (Span.chine_hcomp (SPN (dom ?\)) (SPN (dom ?\)))" proof - have "... = Maps.Dom (Maps.dom (Span.chine_hcomp (SPN ?\) (SPN ?\)))" using 1 Maps.seq_char Maps.Dom_comp by auto also have "... = Maps.Dom (Maps.pbdom (Leg0 (Dom (SPN ?\))) (Leg1 (Dom (SPN ?\))))" using \\ VV.arr_char Span.chine_hcomp_in_hom [of "SPN ?\" "SPN ?\"] by auto also have "... = Maps.Dom (Maps.dom (Maps.pbdom (Leg0 (Dom (SPN ?\))) (Leg1 (Dom (SPN ?\)))))" proof - have "Maps.cospan (Leg0 (Dom (SPN (fst \\)))) (Leg1 (Dom (SPN (snd \\))))" using \\ VV.arr_char SPN_in_hom Span.arr_char Span.dom_char SPN_def Maps.CLS_in_hom Maps.arr_char Maps.cod_char dom_\_\.composable dom_\_\.RS_simps(16) dom_\_\.S\<^sub>1_def dom_\_\.RS_simps(1) dom_\_\.R\<^sub>0_def Maps.pbdom_in_hom by simp thus ?thesis using \\ VV.arr_char Maps.pbdom_in_hom by simp qed also have "... = Maps.Dom (Maps.dom (Maps.pbdom (Leg0 (Dom (SPN (dom ?\)))) (Leg1 (Dom (SPN (dom ?\))))))" using \\ SPN_def VV.arr_char by simp also have "... = Maps.Dom (Maps.dom (Span.chine_hcomp (SPN (dom ?\)) (SPN (dom ?\))))" using \\ VV.arr_char ide_dom by (simp add: Span.chine_hcomp_ide_ide) also have "... = Maps.Dom (Span.chine_hcomp (SPN (dom ?\)) (SPN (dom ?\)))" using Maps.Dom_dom Maps.in_homE SPN.preserves_reflects_arr SPN.preserves_src SPN.preserves_trg Span.chine_hcomp_in_hom dom_\_\.composable dom_\_\.r.base_simps(2) dom_\_\.s.base_simps(2) by (metis (no_types, lifting)) finally show ?thesis by simp qed also have "... = src dom_\_\.\\.p\<^sub>0" using "dom_\_\.chine_hcomp_SPN_SPN" by simp also have "... = Maps.Dom ?Chn_RHS" using 2 Maps.seq_char Maps.Dom_comp by auto finally show ?thesis by simp qed show "Maps.Cod ?Chn_LHS = Maps.Cod ?Chn_RHS" proof - have "Maps.Cod ?Chn_LHS = src (tab_of_ide (cod ?\ \ cod ?\))" using \\ 1 VV.arr_char Maps.seq_char by auto also have "... = src (tab\<^sub>0 (cod ?\ \ cod ?\))" using \\ VV.arr_char cod_\\.tab_simps(2) by auto also have "... = Maps.Cod ?Chn_RHS" by (metis (no_types, lifting) "2" Maps.Cod.simps(1) Maps.Cod_comp Maps.seq_char) finally show ?thesis by simp qed show "Maps.Map ?Chn_LHS = Maps.Map ?Chn_RHS" proof - have RHS: "Maps.Map ?Chn_RHS = iso_class (\\.chine \ dom_\_\.cmp)" proof - have "Maps.Map ?Chn_RHS = Maps.Comp \\\.chine\ \dom_\_\.cmp\" using \\ 2 VV.arr_char Maps.Map_comp Maps.comp_char [of "Maps.MkArr (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\))) \\\.chine\" "Maps.MkArr (src dom_\_\.\\.p\<^sub>0) (src (tab_of_ide (dom ?\ \ dom ?\))) \dom_\_\.cmp\"] by simp also have "... = \\\.chine \ dom_\_\.cmp\" proof - have "\dom_\_\.cmp\ \ Maps.Hom (src dom_\_\.\\.p\<^sub>0) (src (tab\<^sub>0 (dom ?\ \ dom ?\)))" proof - have "\dom_\_\.cmp\ \ Maps.Hom (src dom_\_\.\\.tab) (src (tab_of_ide (dom ?\ \ dom ?\)))" using \\ VV.arr_char dom_\_\.cmp_props(1-3) by (metis (mono_tags, lifting) equivalence_is_left_adjoint mem_Collect_eq) thus ?thesis using \\ VV.arr_char dom_\\.tab_simps(2) by simp qed moreover have "\\\.chine\ \ Maps.Hom (src (tab\<^sub>0 (dom ?\ \ dom ?\))) (src (tab\<^sub>0 (cod ?\ \ cod ?\)))" using \\ VV.arr_char \\.chine_in_hom \\.is_map by auto moreover have "\\.chine \ dom_\_\.cmp \ Maps.Comp \\\.chine\ \dom_\_\.cmp\" proof show "is_iso_class \dom_\_\.cmp\" using is_iso_classI by simp show "is_iso_class \\\.chine\" using is_iso_classI by simp show "dom_\_\.cmp \ \dom_\_\.cmp\" using ide_in_iso_class [of dom_\_\.cmp] by simp show "\\.chine \ \\\.chine\" using ide_in_iso_class by simp show "\\.chine \ dom_\_\.cmp \ \\.chine \ dom_\_\.cmp" using \\ VV.arr_char \\.chine_simps dom_\_\.cmp_simps dom_\\.tab_simps(2) isomorphic_reflexive by auto qed ultimately show ?thesis using \\ dom_\_\.cmp_props \\.chine_in_hom \\.chine_is_induced_map Maps.Comp_eq_iso_class_memb by blast qed finally show ?thesis by simp qed have LHS: "Maps.Map ?Chn_LHS = \cod_\_\.cmp \ \_\.chine\" proof - have "Maps.Map ?Chn_LHS = Maps.Comp \cod_\_\.cmp\ (Maps.Map (Maps.tuple (Maps.CLS (spn ?\ \ dom_\_\.\\.p\<^sub>1)) (Maps.CLS (tab\<^sub>0 (cod ?\))) (Maps.CLS (tab\<^sub>1 (cod ?\))) (Maps.CLS (spn ?\ \ dom_\_\.\\.p\<^sub>0))))" proof - have "Maps.Map ?Chn_LHS = Maps.Comp \cod_\_\.cmp\ (Maps.Map (Span.chine_hcomp (SPN ?\) (SPN ?\)))" using \\ 1 VV.arr_char Maps.Map_comp cod_\\.tab_simps(2) Maps.comp_char [of "Maps.MkArr (src cod_\_\.\\.p\<^sub>0) (src (tab_of_ide (cod ?\ \ cod ?\))) \cod_\_\.cmp\" "Span.chine_hcomp (SPN ?\) (SPN ?\)"] by simp moreover have "Span.chine_hcomp (SPN ?\) (SPN ?\) = Maps.tuple (Maps.CLS (spn ?\ \ dom_\_\.\\.p\<^sub>1)) (Maps.CLS (tab\<^sub>0 (cod ?\))) (Maps.CLS (tab\<^sub>1 (cod ?\))) (Maps.CLS (spn ?\ \ dom_\_\.\\.p\<^sub>0))" proof - have "Maps.PRJ\<^sub>0 (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>0 (dom ?\)\) (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\) = \\dom_\_\.\\.p\<^sub>0\\ \ Maps.PRJ\<^sub>1 (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>0 (dom ?\)\) (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\) = \\dom_\_\.\\.p\<^sub>1\\" proof - interpret X: identity_in_bicategory_of_spans V H \ \ src trg \(tab\<^sub>0 (dom ?\))\<^sup>* \ tab\<^sub>1 (dom ?\)\ using \\ VV.arr_char by (unfold_locales, simp) have "Maps.PRJ\<^sub>0 (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>0 (dom ?\)\) (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\) = \\tab\<^sub>0 ((Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg (snd \\)) \tab\<^sub>0 (dom ?\)\))\<^sup>* \ Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\))\\" unfolding Maps.PRJ\<^sub>0_def using \\ VV.arr_char dom_\_\.RS_simps(1) dom_\_\.RS_simps(16) dom_\_\.RS_simps(18) dom_\_\.RS_simps(3) dom_\_\.R\<^sub>0_def dom_\_\.S\<^sub>1_def by auto moreover have "Maps.PRJ\<^sub>1 (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>0 (dom ?\)\) (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\) = \\tab\<^sub>1 ((Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg (snd \\)) \tab\<^sub>0 (dom ?\)\))\<^sup>* \ Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\))\\" unfolding Maps.PRJ\<^sub>1_def using \\ VV.arr_char dom_\_\.RS_simps(1) dom_\_\.RS_simps(16) dom_\_\.RS_simps(18) dom_\_\.RS_simps(3) dom_\_\.R\<^sub>0_def dom_\_\.S\<^sub>1_def by auto moreover have "(Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg (snd \\)) \tab\<^sub>0 (dom ?\)\))\<^sup>* \ Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\))) (trg ?\) \tab\<^sub>1 (dom ?\)\) \ (tab\<^sub>0 (dom ?\))\<^sup>* \ tab\<^sub>1 (dom ?\)" using VV.arr_char \\ dom_\_\.S\<^sub>1_def dom_\_\.s.leg1_simps(3) dom_\_\.s.leg1_simps(4) trg_dom dom_\_\.R\<^sub>0_def dom_\_\.prj_tab_agreement(1) isomorphic_symmetric by simp ultimately show ?thesis using X.isomorphic_implies_same_tab isomorphic_symmetric by metis qed thus ?thesis unfolding Span.chine_hcomp_def using \\ VV.arr_char SPN_def isomorphic_reflexive Maps.comp_CLS [of "spn ?\" dom_\_\.\\.p\<^sub>1 "spn ?\ \ dom_\_\.\\.p\<^sub>1"] Maps.comp_CLS [of "spn ?\" dom_\_\.\\.p\<^sub>0 "spn ?\ \ dom_\_\.\\.p\<^sub>0"] by simp qed moreover have "Maps.Dom (Span.chine_hcomp (SPN ?\) (SPN ?\)) = src dom_\_\.\\.p\<^sub>0" by (metis (no_types, lifting) "1" "2" Maps.Dom.simps(1) Maps.comp_char \Maps.Dom ?Chn_LHS = Maps.Dom ?Chn_RHS\) ultimately show ?thesis by simp qed also have "... = Maps.Comp \cod_\_\.cmp\ \\_\.chine\" proof - let ?tuple = "Maps.tuple \\spn (fst \\) \ dom_\_\.\\.p\<^sub>1\\ \\tab\<^sub>0 (cod ?\)\\ \\tab\<^sub>1 (cod ?\)\\ \\spn (snd \\) \ dom_\_\.\\.p\<^sub>0\\" have "iso_class \_\.chine = Maps.Map ?tuple" using \_\.CLS_chine spn_def Maps.Map.simps(1) by (metis (no_types, lifting)) thus ?thesis by simp qed also have "... = \cod_\_\.cmp \ \_\.chine\" proof - have "\\_\.chine\ \ Maps.Hom (src dom_\_\.\\.p\<^sub>0) (src cod_\_\.\\.p\<^sub>0)" proof - have "\\_\.chine : src dom_\_\.\\.p\<^sub>0 \ src cod_\_\.\\.p\<^sub>0\" using \\ VV.arr_char by simp thus ?thesis using \_\.is_map ide_in_iso_class left_adjoint_is_ide by blast qed moreover have "\cod_\_\.cmp\ \ Maps.Hom (src cod_\_\.\\.p\<^sub>0) (src (tab\<^sub>0 (cod ?\ \ cod ?\)))" proof - have "\cod_\_\.cmp : src cod_\_\.\\.p\<^sub>0 \ src (tab\<^sub>0 (cod ?\ \ cod ?\))\" using \\ VV.arr_char cod_\_\.cmp_in_hom cod_\\.tab_simps(2) by simp thus ?thesis using cod_\_\.cmp_props equivalence_is_left_adjoint left_adjoint_is_ide ide_in_iso_class [of cod_\_\.cmp] by (metis (mono_tags, lifting) mem_Collect_eq) qed moreover have "cod_\_\.cmp \ \_\.chine \ Maps.Comp \cod_\_\.cmp\ \\_\.chine\" proof show "is_iso_class \\_\.chine\" using \_\.w_simps(1) is_iso_classI by blast show "is_iso_class \cod_\_\.cmp\" using cod_\_\.cmp_simps(2) is_iso_classI by blast show "\_\.chine \ \\_\.chine\" using ide_in_iso_class by simp show "cod_\_\.cmp \ \cod_\_\.cmp\" using ide_in_iso_class by simp show "cod_\_\.cmp \ \_\.chine \ cod_\_\.cmp \ \_\.chine" by (simp add: isomorphic_reflexive) qed ultimately show ?thesis using \\ cod_\_\.cmp_props \_\.chine_in_hom \_\.chine_is_induced_map Maps.Comp_eq_iso_class_memb by simp qed finally show ?thesis by simp qed have EQ: "\\\.chine \ dom_\_\.cmp\ = \cod_\_\.cmp \ \_\.chine\" proof (intro iso_class_eqI) show "\\.chine \ dom_\_\.cmp \ cod_\_\.cmp \ \_\.chine" proof - interpret dom_cmp: identity_arrow_of_tabulations_in_maps V H \ \ src trg \dom ?\\\ dom_\_\.\\.tab \tab\<^sub>0 (dom ?\) \ dom_\_\.\\.p\<^sub>0\ \tab\<^sub>1 (dom ?\) \ dom_\_\.\\.p\<^sub>1\ \dom ?\\\ \tab_of_ide (dom ?\ \ dom ?\)\ \tab\<^sub>0 (dom ?\ \ dom ?\)\ \tab\<^sub>1 (dom ?\ \ dom ?\)\ \dom ?\\\ using \\ VV.arr_char dom_\_\.cmp_interpretation by simp interpret cod_cmp: identity_arrow_of_tabulations_in_maps V H \ \ src trg \cod ?\\\ cod_\_\.\\.tab \tab\<^sub>0 (cod ?\) \ cod_\_\.\\.p\<^sub>0\ \tab\<^sub>1 (cod ?\) \ cod_\_\.\\.p\<^sub>1\ \cod ?\\\ \tab_of_ide (cod ?\ \ cod ?\)\ \tab\<^sub>0 (cod ?\ \ cod ?\)\ \tab\<^sub>1 (cod ?\ \ cod ?\)\ \cod ?\\\ using \\ VV.arr_char cod_\_\.cmp_interpretation by simp interpret L: vertical_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg \dom ?\\\ \dom_\_\.\\.tab\ \tab\<^sub>0 (dom ?\) \ dom_\_\.\\.p\<^sub>0\ \tab\<^sub>1 (dom ?\) \ dom_\_\.\\.p\<^sub>1\ \dom ?\\\ \tab_of_ide (dom ?\\)\ \tab\<^sub>0 (dom ?\\)\ \tab\<^sub>1 (dom ?\\)\ \cod ?\\\ cod_\\.tab \tab\<^sub>0 (cod ?\\)\ \tab\<^sub>1 (cod ?\\)\ \dom ?\\\ \?\ \ ?\\ using \\ VV.arr_char dom_\_\.cmp_in_hom by unfold_locales auto interpret R: vertical_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg \dom ?\\\ \dom_\_\.\\.tab\ \tab\<^sub>0 (dom ?\) \ dom_\_\.\\.p\<^sub>0\ \tab\<^sub>1 (dom ?\) \ dom_\_\.\\.p\<^sub>1\ \cod ?\\\ \cod_\_\.\\.tab\ \tab\<^sub>0 (cod ?\) \ cod_\_\.\\.p\<^sub>0\ \tab\<^sub>1 (cod ?\) \ cod_\_\.\\.p\<^sub>1\ \cod ?\\\ cod_\\.tab \tab\<^sub>0 (cod ?\\)\ \tab\<^sub>1 (cod ?\\)\ \?\ \ ?\\ \cod ?\\\ using \\ VV.arr_char cod_\_\.cmp_in_hom by unfold_locales auto have "\\.chine \ dom_\_\.cmp \ L.chine" using \\ VV.arr_char L.chine_char dom_\_\.cmp_def isomorphic_symmetric by simp also have "... = R.chine" using L.is_ide \\ comp_arr_dom comp_cod_arr isomorphic_reflexive by force also have "... \ cod_\_\.cmp \ \_\.chine" using \\ VV.arr_char R.chine_char cod_\_\.cmp_def by simp finally show ?thesis by simp qed qed show ?thesis using LHS RHS EQ by simp qed qed thus ?thesis using Chn_LHS_eq Chn_RHS_eq by simp qed qed qed qed interpretation \: natural_isomorphism VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map \.map using VV.ide_char VV.arr_char \.map_simp_ide compositor_is_iso by (unfold_locales, simp) lemma compositor_is_natural_transformation: shows "transformation_by_components VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map (\rs. CMP (fst rs) (snd rs))" .. lemma compositor_is_natural_isomorphism: shows "natural_isomorphism VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map \.map" .. end subsubsection "Associativity Coherence" locale three_composable_identities_in_bicategory_of_spans = bicategory_of_spans V H \ \ src trg + f: identity_in_bicategory_of_spans V H \ \ src trg f + g: identity_in_bicategory_of_spans V H \ \ src trg g + h: identity_in_bicategory_of_spans V H \ \ src trg h for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and f :: 'a and g :: 'a and h :: 'a + assumes fg: "src f = trg g" and gh: "src g = trg h" begin interpretation f: arrow_of_tabulations_in_maps V H \ \ src trg f f.tab \tab\<^sub>0 f\ \tab\<^sub>1 f\ f f.tab \tab\<^sub>0 f\ \tab\<^sub>1 f\ f using f.is_arrow_of_tabulations_in_maps by simp interpretation h: arrow_of_tabulations_in_maps V H \ \ src trg h h.tab \tab\<^sub>0 h\ \tab\<^sub>1 h\ h h.tab \tab\<^sub>0 h\ \tab\<^sub>1 h\ h using h.is_arrow_of_tabulations_in_maps by simp interpretation E: self_evaluation_map V H \ \ src trg .. notation E.eval ("\_\") interpretation Maps: maps_category V H \ \ src trg .. interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. no_notation Fun.comp (infixl "\" 55) notation Span.vcomp (infixr "\" 55) notation Span.hcomp (infixr "\" 53) notation Maps.comp (infixr "\" 55) notation isomorphic (infix "\" 50) interpretation SPN: "functor" V Span.vcomp SPN using SPN_is_functor by simp interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN using SPN_is_weak_arrow_of_homs by simp interpretation SPN_SPN: "functor" VV.comp Span.VV.comp SPN.FF using SPN.functor_FF by auto interpretation HoSPN_SPN: composite_functor VV.comp Span.VV.comp Span.vcomp SPN.FF \\\\. fst \\ \ snd \\\ .. interpretation SPNoH: composite_functor VV.comp V Span.vcomp \\\\. fst \\ \ snd \\\ SPN .. text \ Here come a lot of interpretations for ``composite things''. We need these in order to have relatively short, systematic names for entities that will appear in the lemmas to follow. The names of the interpretations use a prefix notation, where \H\ refers to horizontal composition of 1-cells and \T\ refers to composite of tabulations. So, for example, \THfgh\ refers to the composite of the tabulation associated with the horizontal composition \f \ g\ with the tabulation associated with \h\. \ interpretation HHfgh: identity_in_bicategory_of_spans V H \ \ src trg \(f \ g) \ h\ using fg gh by unfold_locales auto interpretation HfHgh: identity_in_bicategory_of_spans V H \ \ src trg \f \ g \ h\ using fg gh by unfold_locales auto interpretation Tfg: two_composable_identities_in_bicategory_of_spans V H \ \ src trg f g using fg gh by unfold_locales auto interpretation Tgh: two_composable_identities_in_bicategory_of_spans V H \ \ src trg g h using fg gh by unfold_locales auto interpretation THfgh: two_composable_identities_in_bicategory_of_spans V H \ \ src trg \f \ g\ h using fg gh by unfold_locales auto interpretation THfgh: tabulation V H \ \ src trg \(f \ g) \ h\ THfgh.\\.tab \tab\<^sub>0 h \ THfgh.\\.p\<^sub>0\ \tab\<^sub>1 (f \ g) \ THfgh.\\.p\<^sub>1\ using THfgh.\\.composite_is_tabulation by simp interpretation TfHgh: two_composable_identities_in_bicategory_of_spans V H \ \ src trg f \g \ h\ using fg gh by unfold_locales auto interpretation TfHgh: tabulation V H \ \ src trg \f \ g \ h\ TfHgh.\\.tab \tab\<^sub>0 (g \ h) \ TfHgh.\\.p\<^sub>0\ \tab\<^sub>1 f \ TfHgh.\\.p\<^sub>1\ using TfHgh.\\.composite_is_tabulation by simp interpretation Tfg_Hfg: arrow_of_tabulations_in_maps V H \ \ src trg \f \ g\ Tfg.\\.tab \tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\ \tab\<^sub>1 f \ Tfg.\\.p\<^sub>1\ \f \ g\ \tab_of_ide (f \ g)\ \tab\<^sub>0 (f \ g)\ \tab\<^sub>1 (f \ g)\ \f \ g\ by unfold_locales auto interpretation Tgh_Hgh: arrow_of_tabulations_in_maps V H \ \ src trg \g \ h\ Tgh.\\.tab \tab\<^sub>0 h \ Tgh.\\.p\<^sub>0\ \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ \g \ h\ \tab_of_ide (g \ h)\ \tab\<^sub>0 (g \ h)\ \tab\<^sub>1 (g \ h)\ \g \ h\ by unfold_locales auto interpretation THfgh_HHfgh: arrow_of_tabulations_in_maps V H \ \ src trg \(f \ g) \ h\ THfgh.\\.tab \tab\<^sub>0 h \ THfgh.\\.p\<^sub>0\ \tab\<^sub>1 (f \ g) \ THfgh.\\.p\<^sub>1\ \(f \ g) \ h\ \tab_of_ide ((f \ g) \ h)\ \tab\<^sub>0 ((f \ g) \ h)\ \tab\<^sub>1 ((f \ g) \ h)\ \(f \ g) \ h\ using fg gh by unfold_locales auto interpretation TfHgh_HfHgh: arrow_of_tabulations_in_maps V H \ \ src trg \f \ g \ h\ TfHgh.\\.tab \tab\<^sub>0 (g \ h) \ TfHgh.\\.p\<^sub>0\ \tab\<^sub>1 f \ TfHgh.\\.p\<^sub>1\ \f \ g \ h\ \tab_of_ide (f \ g \ h)\ \tab\<^sub>0 (f \ g \ h)\ \tab\<^sub>1 (f \ g \ h)\ \f \ g \ h\ using fg gh by unfold_locales auto interpretation TTfgh: composite_tabulation_in_maps V H \ \ src trg \f \ g\ Tfg.\\.tab \tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\ \tab\<^sub>1 f \ Tfg.\\.p\<^sub>1\ h \tab_of_ide h\ \tab\<^sub>0 h\ \tab\<^sub>1 h\ using fg gh by unfold_locales auto interpretation TTfgh_THfgh: horizontal_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg \f \ g\ Tfg.\\.tab \tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\ \tab\<^sub>1 f \ Tfg.\\.p\<^sub>1\ h \tab_of_ide h\ \tab\<^sub>0 h\ \tab\<^sub>1 h\ \f \ g\ \tab_of_ide (f \ g)\ \tab\<^sub>0 (f \ g)\ \tab\<^sub>1 (f \ g)\ h \tab_of_ide h\ \tab\<^sub>0 h\ \tab\<^sub>1 h\ \f \ g\ h .. interpretation TfTgh: composite_tabulation_in_maps V H \ \ src trg f \tab_of_ide f\ \tab\<^sub>0 f\ \tab\<^sub>1 f\ \g \ h\ Tgh.\\.tab \tab\<^sub>0 h \ Tgh.\\.p\<^sub>0\ \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ using fg gh by unfold_locales auto interpretation TfTgh_TfHgh: horizontal_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg f \tab_of_ide f\ \tab\<^sub>0 f\ \tab\<^sub>1 f\ \g \ h\ Tgh.\\.tab \tab\<^sub>0 h \ Tgh.\\.p\<^sub>0\ \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ f \tab_of_ide f\ \tab\<^sub>0 f\ \tab\<^sub>1 f\ \g \ h\ \tab_of_ide (g \ h)\ \tab\<^sub>0 (g \ h)\ \tab\<^sub>1 (g \ h)\ f \g \ h\ .. interpretation TfTgh_TfTgh: horizontal_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg f \tab_of_ide f\ \tab\<^sub>0 f\ \tab\<^sub>1 f\ \g \ h\ Tgh.\\.tab \tab\<^sub>0 h \ Tgh.\\.p\<^sub>0\ \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ f \tab_of_ide f\ \tab\<^sub>0 f\ \tab\<^sub>1 f\ \g \ h\ Tgh.\\.tab \tab\<^sub>0 h \ Tgh.\\.p\<^sub>0\ \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ f \g \ h\ .. text \ The following interpretation defines the associativity between the peaks of the two composite tabulations \TTfgh\ (associated to the left) and \TfTgh\ (associated to the right). \ (* TODO: Try to get rid of the .\\ in, e.g., Tfg.\\.p\<^sub>1. *) interpretation TTfgh_TfTgh: arrow_of_tabulations_in_maps V H \ \ src trg \(f \ g) \ h\ TTfgh.tab \tab\<^sub>0 h \ TTfgh.p\<^sub>0\ \(tab\<^sub>1 f \ Tfg.\\.p\<^sub>1) \ TTfgh.p\<^sub>1\ \f \ g \ h\ TfTgh.tab \(tab\<^sub>0 h \ Tgh.\\.p\<^sub>0) \ TfTgh.p\<^sub>0\ \tab\<^sub>1 f \ TfTgh.p\<^sub>1\ \\[f, g, h]\ using fg gh by unfold_locales auto text \ This interpretation defines the map, from the apex of the tabulation associated with the horizontal composite \(f \ g) \ h\ to the apex of the tabulation associated with the horizontal composite \f \ g \ h\, induced by the associativity isomorphism \\[f, g, h]\ from \(f \ g) \ h\ to \f \ g \ h\. \ interpretation HHfgh_HfHgh: arrow_of_tabulations_in_maps V H \ \ src trg \dom (\ (f, g, h))\ \tab_of_ide (dom (\ (f, g, h)))\ \tab\<^sub>0 (dom (\ (f, g, h)))\ \tab\<^sub>1 (dom (\ (f, g, h)))\ \cod (\ (f, g, h))\ \tab_of_ide (cod (\ (f, g, h)))\ \tab\<^sub>0 (cod (\ (f, g, h)))\ \tab\<^sub>1 (cod (\ (f, g, h)))\ \\ (f, g, h)\ proof - have "arrow_of_tabulations_in_maps V H \ \ src trg ((f \ g) \ h) (tab_of_ide ((f \ g) \ h)) (tab\<^sub>0 ((f \ g) \ h)) (tab\<^sub>1 ((f \ g) \ h)) (f \ g \ h) (tab_of_ide (f \ g \ h)) (tab\<^sub>0 (f \ g \ h)) (tab\<^sub>1 (f \ g \ h)) \[f, g, h]" using fg gh by unfold_locales auto thus "arrow_of_tabulations_in_maps V H \ \ src trg (dom (\ (f, g, h))) (tab_of_ide (dom (\ (f, g, h)))) (tab\<^sub>0 (dom (\ (f, g, h)))) (tab\<^sub>1 (dom (\ (f, g, h)))) (cod (\ (f, g, h))) (tab_of_ide (cod (\ (f, g, h)))) (tab\<^sub>0 (cod (\ (f, g, h)))) (tab\<^sub>1 (cod (\ (f, g, h)))) (\ (f, g, h))" using fg gh \_def by auto qed interpretation SPN_f: arrow_of_spans Maps.comp \SPN f\ using SPN_in_hom Span.arr_char [of "SPN f"] by simp interpretation SPN_g: arrow_of_spans Maps.comp \SPN g\ using SPN_in_hom Span.arr_char [of "SPN g"] by simp interpretation SPN_h: arrow_of_spans Maps.comp \SPN h\ using SPN_in_hom Span.arr_char [of "SPN h"] by simp interpretation SPN_fgh: three_composable_identity_arrows_of_spans Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 \SPN f\ \SPN g\ \SPN h\ using fg gh Span.arr_char SPN_in_hom SPN.preserves_ide Span.ide_char apply unfold_locales by auto text \ The following relates the projections associated with the composite span \SPN_fgh\ with tabulations in the underlying bicategory. \ lemma prj_char: shows "SPN_fgh.Prj\<^sub>1\<^sub>1 = \\Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1\\" and "SPN_fgh.Prj\<^sub>0\<^sub>1 = \\Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1\\" and "SPN_fgh.Prj\<^sub>0 = \\TTfgh.p\<^sub>0\\" and "SPN_fgh.Prj\<^sub>1 = \\TfTgh.p\<^sub>1\\" and "SPN_fgh.Prj\<^sub>1\<^sub>0 = \\Tgh.\\.p\<^sub>1 \ TfTgh.p\<^sub>0\\" and "SPN_fgh.Prj\<^sub>0\<^sub>0 = \\Tgh.\\.p\<^sub>0 \ TfTgh.p\<^sub>0\\" proof - show "SPN_fgh.Prj\<^sub>1\<^sub>1 = \\Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1\\" proof - have "ide (Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1)" by (metis TTfgh.composable TTfgh.leg1_simps(2) Tfg.\\.T0.antipar(2) Tfg.\\.T0.ide_right Tfg_Hfg.u_simps(3) f.T0.antipar(2) f.T0.ide_right f.u_simps(3) fg g.ide_leg1 g.leg1_simps(4) h.ide_leg1 h.leg1_simps(4) ide_hcomp hseqE hcomp_simps(1) tab\<^sub>1_simps(1)) thus ?thesis using fg gh Tfg.\\.prj_char TTfgh.prj_char isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\\.p\<^sub>0 "tab\<^sub>0 g \ Tfg.\\.p\<^sub>0"] Maps.comp_CLS [of Tfg.\\.p\<^sub>1 TTfgh.p\<^sub>1 "Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1"] by (simp add: TTfgh.composable Tfg.\\.T0.antipar(2)) qed show "SPN_fgh.Prj\<^sub>0\<^sub>1 = \\Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1\\" proof - have "ide (Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1)" by (metis TTfgh.leg1_simps(2) bicategory_of_spans.tab\<^sub>0_simps(1) bicategory_of_spans.tab\<^sub>1_simps(1) bicategory_of_spans_axioms Tfg.\\.T0.antipar(2) Tfg.\\.T0.ide_right Tfg.composable f.T0.antipar(2) f.T0.ide_right f.u_simps(3) g.ide_leg1 g.leg1_simps(4) Tfg.u_simps(3) THfgh.composable h.ide_leg1 h.leg1_simps(4) ide_hcomp hseqE hcomp_simps(1) tab\<^sub>1_simps(3)) thus ?thesis using fg gh Tfg.\\.prj_char TTfgh.prj_char isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\\.p\<^sub>0 "tab\<^sub>0 g \ Tfg.\\.p\<^sub>0"] Maps.comp_CLS [of Tfg.\\.p\<^sub>0 TTfgh.p\<^sub>1 "Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1"] by (simp add: Tfg.\\.T0.antipar(2) THfgh.composable) qed show "SPN_fgh.Prj\<^sub>0 = \\TTfgh.p\<^sub>0\\" using isomorphic_reflexive TTfgh.prj_char Tfg.\\.prj_char Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\\.p\<^sub>0 "tab\<^sub>0 g \ Tfg.\\.p\<^sub>0"] by (simp add: Tfg.composable) show "SPN_fgh.Prj\<^sub>1 = \\TfTgh.p\<^sub>1\\" using Tgh.\\.prj_char isomorphic_reflexive Tgh.composable Maps.comp_CLS [of "tab\<^sub>1 g" Tgh.\\.p\<^sub>1 "tab\<^sub>1 g \ Tgh.\\.p\<^sub>1"] TfTgh.prj_char by simp show "SPN_fgh.Prj\<^sub>1\<^sub>0 = \\Tgh.\\.p\<^sub>1 \ TfTgh.p\<^sub>0\\" using fg gh Tgh.\\.prj_char TfTgh.prj_char isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>1 g" "prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)" "tab\<^sub>1 g \ Tgh.\\.p\<^sub>1"] Maps.comp_CLS [of Tgh.\\.p\<^sub>1 TfTgh.p\<^sub>0 "Tgh.\\.p\<^sub>1 \ TfTgh.p\<^sub>0"] by simp show "SPN_fgh.Prj\<^sub>0\<^sub>0 = \\Tgh.\\.p\<^sub>0 \ TfTgh.p\<^sub>0\\" using fg gh Tgh.\\.prj_char TfTgh.prj_char isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>1 g" "Tgh.\\.p\<^sub>1" "tab\<^sub>1 g \ Tgh.\\.p\<^sub>1"] Maps.comp_CLS [of Tgh.\\.p\<^sub>0 TfTgh.p\<^sub>0 "Tgh.\\.p\<^sub>0 \ TfTgh.p\<^sub>0"] by simp qed interpretation \: transformation_by_components VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map \\rs. CMP (fst rs) (snd rs)\ using compositor_is_natural_transformation by simp interpretation \: natural_isomorphism VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map \.map using compositor_is_natural_isomorphism by simp (* * TODO: Figure out how this subcategory gets introduced. * The simps in the locale are used in the subsequent proofs. *) interpretation VVV': subcategory VxVxV.comp \\\\\. arr (fst \\\) \ arr (fst (snd \\\)) \ arr (snd (snd \\\)) \ src (fst (snd \\\)) = trg (snd (snd \\\)) \ src (fst \\\) = trg (fst (snd \\\))\ using fg gh VVV.arr_char VV.arr_char VVV.subcategory_axioms by simp text \ We define abbreviations for the left and right-hand sides of the equation for associativity coherence. \ (* * TODO: \ doesn't really belong in this locale. Replace it with CMP and rearrange * material so that this locale comes first and the definition of \ comes later * in bicategory_of_spans. *) abbreviation LHS where "LHS \ SPN \[f, g, h] \ \.map (f \ g, h) \ (\.map (f, g) \ SPN h)" abbreviation RHS where "RHS \ \.map (f, g \ h) \ (SPN f \ \.map (g, h)) \ Span.assoc (SPN f) (SPN g) (SPN h)" lemma arr_LHS: shows "Span.arr LHS" using fg gh VVV.arr_char VVV.ide_char VV.arr_char VV.ide_char Span.hseqI' HoHV_def compositor_in_hom \_def apply (intro Span.seqI) apply simp_all using SPN.FF_def apply simp proof - have "SPN ((f \ g) \ h) = Span.cod (CMP (f \ g) h)" using fg gh compositor_in_hom by simp also have "... = Span.cod (CMP (f \ g) h \ (CMP f g \ SPN h))" proof - have "Span.seq (CMP (f \ g) h) (CMP f g \ SPN h)" proof (intro Span.seqI Span.hseqI) show 1: "Span.in_hhom (SPN h) (SPN.map\<^sub>0 (src h)) (SPN.map\<^sub>0 (trg h))" using SPN.preserves_src SPN.preserves_trg by simp show 2: "Span.in_hhom (CMP f g) (SPN.map\<^sub>0 (trg h)) (SPN.map\<^sub>0 (trg f))" using compositor_in_hom SPN_fgh.\\.composable fg by auto show 3: "Span.arr (CMP (f \ g) h)" using TTfgh.composable Tfg.\\.ide_base compositor_simps(1) h.is_ide by auto show "Span.dom (CMP (f \ g) h) = Span.cod (CMP f g \ SPN h)" using 1 2 3 fg gh compositor_in_hom SPN_fgh.\\.composable SPN_in_hom SPN.FF_def by auto qed thus ?thesis by simp qed finally show "SPN ((f \ g) \ h) = Span.cod (CMP (f \ g) h \ (CMP f g \ SPN h))" by blast qed lemma arr_RHS: shows "Span.arr RHS" using fg gh VV.ide_char VV.arr_char \.map_simp_ide SPN.FF_def Span.hseqI' by (intro Span.seqI, simp_all) lemma par_LHS_RHS: shows "Span.par LHS RHS" proof (intro conjI) show "Span.arr LHS" using arr_LHS by simp show "Span.arr RHS" using arr_RHS by simp show "Span.dom LHS = Span.dom RHS" proof - have "Span.dom LHS = Span.dom (\.map (f, g) \ SPN h)" using arr_LHS by auto also have "... = Span.dom (\.map (f, g)) \ Span.dom (SPN h)" using arr_LHS Span.dom_hcomp [of "SPN h" "\.map (f, g)"] by blast also have "... = (SPN f \ SPN g) \ SPN h" using fg \.map_simp_ide VV.ide_char VV.arr_char SPN.FF_def by simp also have "... = Span.dom (Span.assoc (SPN f) (SPN g) (SPN h))" using fg gh VVV.arr_char VVV.ide_char VV.arr_char VV.ide_char by simp also have "... = Span.dom RHS" using \Span.arr RHS\ by auto finally show ?thesis by blast qed show "Span.cod LHS = Span.cod RHS" proof - have "Span.cod LHS = Span.cod (SPN \[f, g, h])" using arr_LHS by simp also have "... = SPN (f \ g \ h)" unfolding \_def using fg gh VVV.ide_char VVV.arr_char VV.ide_char VV.arr_char HoVH_def by simp also have "... = Span.cod RHS" using arr_RHS fg gh \.map_simp_ide VV.ide_char VV.arr_char SPN.FF_def compositor_in_hom by simp finally show ?thesis by blast qed qed lemma Chn_LHS_eq: shows "Chn LHS = \\HHfgh_HfHgh.chine\\ \ \\THfgh_HHfgh.chine\\ \ \\TTfgh_THfgh.chine\\" proof - have "Chn LHS = \\HHfgh_HfHgh.chine\\ \ \\THfgh_HHfgh.chine\\ \ Span.chine_hcomp (CMP f g) (SPN h)" proof - have "Chn LHS = Chn (SPN \[f, g, h]) \ Chn (CMP (f \ g) h) \ Chn (CMP f g \ SPN h)" using fg gh arr_LHS \.map_simp_ide VV.ide_char VV.arr_char Span.Chn_vcomp by auto moreover have "Chn (SPN \[f, g, h]) = Maps.CLS HHfgh_HfHgh.chine" using fg gh SPN_def VVV.arr_char VV.arr_char spn_def \_def by simp moreover have "Chn (CMP (f \ g) h) = Maps.CLS THfgh_HHfgh.chine" using fg gh CMP_def THfgh.cmp_def by simp moreover have "Chn (CMP f g \ SPN h) = Span.chine_hcomp (CMP f g) (SPN h)" using fg gh Span.hcomp_def by simp ultimately show ?thesis by simp qed moreover have "Span.chine_hcomp (CMP f g) (SPN h) = \\TTfgh_THfgh.chine\\" proof - have "Span.chine_hcomp (CMP f g) (SPN h) = Maps.tuple (\\Tfg.cmp\\ \ Maps.PRJ\<^sub>1 \\tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\\ \\tab\<^sub>1 h\\) \\tab\<^sub>0 (f \ g)\\ \\tab\<^sub>1 h\\ (\\spn h\\ \ Maps.PRJ\<^sub>0 \\tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\\ \\tab\<^sub>1 h\\)" proof - have "\\tab\<^sub>0 g\\ \ \\Tfg.\\.p\<^sub>0\\ = \\tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\\" using fg isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>0 g" "Tfg.\\.p\<^sub>0" "tab\<^sub>0 g \ Tfg.\\.p\<^sub>0"] by simp moreover have "span_in_category.apex Maps.comp \Leg0 = \\tab\<^sub>0 h\\, Leg1 = \\tab\<^sub>1 h\\\ = \\spn h\\" proof - interpret h: span_in_category Maps.comp \\Leg0 = \\tab\<^sub>0 h\\, Leg1 = \\tab\<^sub>1 h\\\\ using h.determines_span by simp interpret dom_h: identity_arrow_of_tabulations_in_maps V H \ \ src trg \dom h\ \tab_of_ide (dom h)\ \tab\<^sub>0 (dom h)\ \tab\<^sub>1 (dom h)\ \cod h\ \tab_of_ide (cod h)\ \tab\<^sub>0 (cod h)\ \tab\<^sub>1 (cod h)\ h by (simp add: h.is_arrow_of_tabulations_in_maps identity_arrow_of_tabulations_in_maps.intro identity_arrow_of_tabulations_in_maps_axioms.intro) have "Maps.arr h.leg0" using h.leg_simps(1) by simp hence "Maps.dom h.leg0 = \\dom_h.chine\\" using Maps.dom_char Maps.CLS_in_hom apply simp proof - have "h.is_induced_map (src (tab\<^sub>0 h))" using h.is_induced_map_iff dom_h.\_eq_\ h.apex_is_induced_by_cell by force hence "src (tab\<^sub>0 h) \ h.chine" using h.chine_is_induced_map h.induced_map_unique by simp thus "\src (tab\<^sub>0 h)\ = \h.chine\" using iso_class_eqI by simp qed thus ?thesis using h.apex_def spn_def by simp qed ultimately show ?thesis unfolding Span.chine_hcomp_def using fg gh CMP_def Tfg.\\.prj_char Span.hcomp_def by simp qed also have "... = \\TTfgh_THfgh.chine\\" proof - have "\\TTfgh_THfgh.chine\\ = Maps.tuple \\Tfg_Hfg.chine \ TTfgh.p\<^sub>1\\ \\tab\<^sub>0 (f \ g)\\ \\tab\<^sub>1 h\\ \\h.chine \ TTfgh.p\<^sub>0\\" using TTfgh_THfgh.CLS_chine by simp also have "... = Maps.tuple (\\Tfg_Hfg.chine\\ \ \\TTfgh.p\<^sub>1\\) \\tab\<^sub>0 (f \ g)\\ \\tab\<^sub>1 h\\ (\\h.chine\\ \ \\TTfgh.p\<^sub>0\\)" proof - have "\\Tfg_Hfg.chine \ TTfgh.p\<^sub>1\\ = \\Tfg_Hfg.chine\\ \ \\TTfgh.p\<^sub>1\\" proof - have "is_left_adjoint TTfgh.p\<^sub>1" using Tfg.\\.T0.antipar(2) THfgh.composable by simp moreover have "Tfg_Hfg.chine \ TTfgh.p\<^sub>1 \ Tfg_Hfg.chine \ TTfgh.p\<^sub>1" using TTfgh_THfgh.prj_chine(2) isomorphic_reflexive isomorphic_implies_hpar(2) by blast ultimately show ?thesis using Tfg_Hfg.is_map Maps.comp_CLS [of Tfg_Hfg.chine TTfgh.p\<^sub>1 "Tfg_Hfg.chine \ TTfgh.p\<^sub>1"] by simp qed moreover have "\\h.chine \ TTfgh.p\<^sub>0\\ = \\h.chine\\ \ \\TTfgh.p\<^sub>0\\" proof - have "is_left_adjoint TTfgh.p\<^sub>0" by (simp add: Tfg.\\.T0.antipar(2) THfgh.composable) moreover have "h.chine \ TTfgh.p\<^sub>0 \ h.chine \ TTfgh.p\<^sub>0" using TTfgh_THfgh.prj_chine(1) isomorphic_reflexive isomorphic_implies_hpar(2) by blast ultimately show ?thesis using h.is_map Maps.comp_CLS [of h.chine TTfgh.p\<^sub>0 "h.chine \ TTfgh.p\<^sub>0"] by simp qed ultimately show ?thesis by argo qed also have "... = Maps.tuple (\\Tfg.cmp\\ \ Maps.PRJ\<^sub>1 \\tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\\ \\tab\<^sub>1 h\\) \\tab\<^sub>0 (f \ g)\\ \\tab\<^sub>1 h\\ (\\spn h\\ \ Maps.PRJ\<^sub>0 \\tab\<^sub>0 g \ Tfg.\\.p\<^sub>0\\ \\tab\<^sub>1 h\\)" using Tfg.cmp_def spn_def TTfgh.prj_char by simp finally show ?thesis by simp qed finally show ?thesis by blast qed ultimately show ?thesis by simp qed abbreviation tuple_BC where "tuple_BC \ Maps.tuple SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.\.leg0 SPN_fgh.\.leg1 SPN_fgh.Prj\<^sub>0" abbreviation tuple_ABC where "tuple_ABC \ Maps.tuple SPN_fgh.Prj\<^sub>1\<^sub>1 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) tuple_BC" abbreviation tuple_BC' where "tuple_BC' \ Maps.tuple \\Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1\\ \\tab\<^sub>0 g\\ \\tab\<^sub>1 h\\ \\TTfgh.p\<^sub>0\\" abbreviation tuple_ABC' where "tuple_ABC' \ Maps.tuple \\Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1\\ \\tab\<^sub>0 f\\ \\tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\\ tuple_BC'" lemma csq: shows "Maps.commutative_square SPN_fgh.\.leg0 SPN_fgh.\.leg1 SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0" and "Maps.commutative_square SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC" proof - show 1: "Maps.commutative_square SPN_fgh.\.leg0 SPN_fgh.\.leg1 SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0" proof show "Maps.cospan SPN_fgh.\.leg0 SPN_fgh.\.leg1" using SPN_fgh.\\.legs_form_cospan(1) by simp show "Maps.span SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0" using SPN_fgh.prj_simps(2-3,5-6) by presburger show "Maps.dom SPN_fgh.\.leg0 = Maps.cod SPN_fgh.Prj\<^sub>0\<^sub>1" using SPN_fgh.prj_simps(8) SPN_g.dom.is_span SPN_g.dom.leg_simps(2) by auto show "SPN_fgh.\.leg0 \ SPN_fgh.Prj\<^sub>0\<^sub>1 = SPN_fgh.\.leg1 \ SPN_fgh.Prj\<^sub>0" by (metis (no_types, lifting) Maps.cod_comp Maps.comp_assoc Maps.pullback_commutes' SPN_fgh.\\.dom.leg_simps(1) SPN_fgh.\\.leg0_composite SPN_fgh.cospan_\\) qed show "Maps.commutative_square SPN_fgh.\.leg0 (Maps.comp SPN_fgh.\.leg1 SPN_fgh.\\.prj\<^sub>1) SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC" proof show "Maps.cospan SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" using fg gh SPN_fgh.prj_simps(10) by blast show "Maps.span SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC" using fg gh 1 Maps.tuple_simps(1) Maps.tuple_simps(2) SPN_fgh.prj_simps(1) SPN_fgh.prj_simps(4) SPN_fgh.prj_simps(5) by presburger show "Maps.dom SPN_fgh.\.leg0 = Maps.cod SPN_fgh.Prj\<^sub>1\<^sub>1" using fg gh SPN_f.dom.leg_simps(2) SPN_fgh.prj_simps(7) by auto show "SPN_fgh.\.leg0 \ SPN_fgh.Prj\<^sub>1\<^sub>1 = (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ tuple_BC" using 1 fg gh Maps.comp_assoc Maps.prj_tuple by (metis (no_types, lifting) Maps.pullback_commutes' SPN_fgh.cospan_\\) qed qed lemma tuple_ABC_eq_ABC': shows "tuple_BC = tuple_BC'" and "tuple_ABC = tuple_ABC'" proof - have "SPN_fgh.Prj\<^sub>1\<^sub>1 = \\Tfg.\\.p\<^sub>1 \ TTfgh.p\<^sub>1\\" using prj_char by simp moreover have "SPN_fgh.\.leg0 = \\tab\<^sub>0 f\\" by simp moreover have "SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1 = \\tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\\" using Tgh.\\.prj_char isomorphic_reflexive Tgh.composable Maps.comp_CLS [of "tab\<^sub>1 g" Tgh.\\.p\<^sub>1 "tab\<^sub>1 g \ Tgh.\\.p\<^sub>1"] by (simp add: g.T0.antipar(2)) moreover show "tuple_BC = tuple_BC'" using prj_char Tfg.\\.prj_char by simp ultimately show "tuple_ABC = tuple_ABC'" by argo qed lemma tuple_BC_in_hom: shows "Maps.in_hom tuple_BC (Maps.MkIde (src TTfgh.p\<^sub>0)) (Maps.MkIde (src Tgh.\\.p\<^sub>0))" proof show 1: "Maps.arr tuple_BC" using csq(1) by simp have 2: "Maps.commutative_square \\tab\<^sub>0 g\\ \\tab\<^sub>1 h\\ \\Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1\\ \\TTfgh.p\<^sub>0\\" by (metis Tfg.S\<^sub>0_def Tfg.span_legs_eq(3) Tgh.S\<^sub>1_def Tgh.span_legs_eq(4) csq(1) prj_char(2) prj_char(3)) show "Maps.dom tuple_BC = Maps.MkIde (src TTfgh.p\<^sub>0)" proof - have "Maps.dom tuple_BC' = Maps.dom \\Tfg.\\.p\<^sub>0 \ TTfgh.p\<^sub>1\\" using 2 Maps.tuple_simps by simp also have "... = Chn (Span.hcomp (Span.hcomp (SPN f) (SPN g)) (SPN h))" using Maps.dom_char by (metis SPN_fgh.prj_simps(5) prj_char(2)) also have "... = Maps.MkIde (src TTfgh.p\<^sub>0)" using 1 fg gh Maps.dom_char csq(1) prj_char(3) tuple_ABC_eq_ABC'(1) Maps.Dom.simps(1) Maps.tuple_simps(2) SPN_fgh.prj_simps(3,5-6) by presburger finally have "Maps.dom tuple_BC' = Maps.MkIde (src TTfgh.p\<^sub>0)" by blast thus ?thesis using tuple_ABC_eq_ABC' by simp qed show "Maps.cod tuple_BC = Maps.MkIde (src Tgh.\\.p\<^sub>0)" proof - have "Maps.cod tuple_BC' = Maps.pbdom \\tab\<^sub>0 g\\ \\tab\<^sub>1 h\\" using 1 2 fg gh Maps.tuple_in_hom by blast also have "... = Maps.MkIde (src Tgh.\\.p\<^sub>0)" using 1 2 fg gh Maps.pbdom_def by (metis (no_types, lifting) SPN.preserves_ide SPN_fgh.\\.are_identities(2) SPN_fgh.\\.composable Span.chine_hcomp_ide_ide Tfg.S\<^sub>0_def Tfg.span_legs_eq(3) Tgh.S\<^sub>1_def Tgh.chine_hcomp_SPN_SPN Tgh.span_legs_eq(4) g.is_ide) finally show ?thesis using tuple_ABC_eq_ABC' by simp qed qed lemma tuple_ABC_in_hom: shows "Maps.in_hom tuple_ABC (Maps.MkIde (src TTfgh.p\<^sub>0)) (Maps.MkIde (src TfTgh.p\<^sub>0))" proof show 1: "Maps.arr tuple_ABC" using SPN_fgh.chine_assoc_def SPN_fgh.chine_assoc_in_hom by auto show "Maps.dom tuple_ABC = Maps.MkIde (src TTfgh.p\<^sub>0)" proof - have "Maps.dom tuple_ABC = Maps.dom SPN_fgh.chine_assoc" by (simp add: SPN_fgh.chine_assoc_def) also have "... = Chn ((SPN f \ SPN g) \ SPN h)" using SPN_fgh.chine_assoc_in_hom by blast also have "... = Maps.MkIde (src TTfgh.p\<^sub>0)" by (metis (lifting) Maps.Dom.simps(1) Maps.dom_char SPN_fgh.prj_simps(3) SPN_fgh.prj_simps(6) prj_char(3)) finally show ?thesis by blast qed show "Maps.cod tuple_ABC = Maps.MkIde (src TfTgh.p\<^sub>0)" proof - have "Maps.cod tuple_ABC = Maps.cod SPN_fgh.chine_assoc" by (simp add: SPN_fgh.chine_assoc_def) also have 1: "... = Chn (SPN f \ SPN g \ SPN h)" using SPN_fgh.chine_assoc_in_hom by blast also have "... = Maps.MkIde (src TfTgh.p\<^sub>0)" by (metis (lifting) Maps.Dom.simps(1) Maps.cod_char Maps.seq_char SPN_fgh.prj_chine_assoc(1) SPN_fgh.prj_simps(1) TfTgh.leg1_in_hom(1) TfTgh_TfTgh.u_in_hom 1 in_hhomE prj_char(4) src_hcomp) finally show ?thesis by argo qed qed lemma Chn_RHS_eq: shows "Chn RHS = \\TfHgh_HfHgh.chine\\ \ \\TfTgh_TfHgh.chine\\ \ tuple_ABC'" proof - have "Chn RHS = Chn (\.map (f, g \ h)) \ Chn (SPN f \ \.map (g, h)) \ Chn (Span.assoc (SPN f) (SPN g) (SPN h))" proof - have "Chn RHS = Chn (\.map (f, g \ h)) \ Chn ((SPN f \ \.map (g, h)) \ Span.assoc (SPN f) (SPN g) (SPN h))" using arr_RHS Span.vcomp_eq Span.Chn_vcomp by blast also have "... = Chn (\.map (f, g \ h)) \ Chn (SPN f \ \.map (g, h)) \ Chn (Span.assoc (SPN f) (SPN g) (SPN h))" proof - have "Span.seq (SPN f \ \.map (g, h)) (Span.assoc (SPN f) (SPN g) (SPN h))" using arr_RHS by auto thus ?thesis using fg gh Span.vcomp_eq [of "SPN f \ \.map (g, h)" "Span.assoc (SPN f) (SPN g) (SPN h)"] by simp qed finally show ?thesis by blast qed moreover have "Chn (\.map (f, g \ h)) = \\TfHgh_HfHgh.chine\\" using arr_RHS fg gh \.map_simp_ide VV.ide_char VV.arr_char CMP_def TfHgh.cmp_def by simp moreover have "Chn (SPN f \ \.map (g, h)) = Span.chine_hcomp (SPN f) (CMP g h)" using fg gh Span.hcomp_def \.map_simp_ide VV.ide_char VV.arr_char SPN.FF_def by simp moreover have "Chn (Span.assoc (SPN f) (SPN g) (SPN h)) = tuple_ABC" using fg gh Span.\_ide VVV.ide_char VVV.arr_char VV.ide_char VV.arr_char SPN_fgh.chine_assoc_def Span.\_def by simp moreover have "Span.chine_hcomp (SPN f) (CMP g h) = \\TfTgh_TfHgh.chine\\" proof - have "Span.chine_hcomp (SPN f) (CMP g h) = Maps.tuple (\\f.chine\\ \ \\TfTgh.p\<^sub>1\\) \\tab\<^sub>0 f\\ \\tab\<^sub>1 (g \ h)\\ (\\Tgh_Hgh.chine\\ \ \\TfTgh.p\<^sub>0\\)" proof - interpret f: span_in_category Maps.comp \\Leg0 = Maps.MkArr (src (tab\<^sub>0 f)) (trg g) \tab\<^sub>0 f\, Leg1 = Maps.MkArr (src (tab\<^sub>0 f)) (trg f) \tab\<^sub>1 f\\\ using f.determines_span by (simp add: Tfg.composable) interpret f: arrow_of_tabulations_in_maps V H \ \ src trg f f.tab \tab\<^sub>0 f\ \tab\<^sub>1 f\ f f.tab \tab\<^sub>0 f\ \tab\<^sub>1 f\ f using f.is_arrow_of_tabulations_in_maps by simp have "f.apex = Maps.CLS f.chine" proof (intro Maps.arr_eqI) show "Maps.arr f.apex" by simp show "Maps.arr \\f.chine\\" using Maps.CLS_in_hom f.is_map by blast show "Maps.Dom f.apex = Maps.Dom \\f.chine\\" using f.apex_def Tfg.RS_simps(2) Tfg.R\<^sub>0_def Tfg.composable by auto show "Maps.Cod f.apex = Maps.Cod \\f.chine\\" using f.apex_def Tfg.RS_simps(2) Tfg.R\<^sub>0_def Tfg.composable by auto show "Maps.Map f.apex = Maps.Map \\f.chine\\" proof - have "Maps.Map f.apex = \src (tab\<^sub>0 f)\" using f.apex_def Maps.dom_char Tfg.RS_simps(2) Tfg.R\<^sub>0_def Tfg.composable by auto also have "... = \f.chine\" proof (intro iso_class_eqI) have "f.is_induced_map (src (tab\<^sub>0 f))" using f.apex_is_induced_by_cell comp_cod_arr by auto thus "src (tab\<^sub>0 f) \ f.chine" using f.induced_map_unique f.chine_is_induced_map by simp qed also have "... = Maps.Map \\f.chine\\" by simp finally show ?thesis by simp qed qed thus ?thesis unfolding Span.chine_hcomp_def using fg gh CMP_def Tgh.\\.prj_char Span.hcomp_def isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>1 g" Tgh.\\.p\<^sub>1 "tab\<^sub>1 g \ Tgh.\\.p\<^sub>1"] Tgh.cmp_def TfTgh.prj_char by simp qed also have "... = Maps.tuple \\f.chine \ TfTgh.p\<^sub>1\\ \\tab\<^sub>0 f\\ \\tab\<^sub>1 (g \ h)\\ \\Tgh_Hgh.chine \ TfTgh.p\<^sub>0\\" using isomorphic_reflexive TfHgh.composable f.is_map TfHgh.composable Tgh_Hgh.is_map Maps.comp_CLS [of f.chine TfTgh.p\<^sub>1 "f.chine \ TfTgh.p\<^sub>1"] Maps.comp_CLS [of Tgh_Hgh.chine TfTgh.p\<^sub>0 "Tgh_Hgh.chine \ TfTgh.p\<^sub>0"] by auto also have "... = \\TfTgh_TfHgh.chine\\" using TfTgh_TfHgh.CLS_chine by simp finally show ?thesis by blast qed ultimately have "Chn RHS =\\TfHgh_HfHgh.chine\\ \ \\TfTgh_TfHgh.chine\\ \ tuple_ABC" by simp also have "... = \\TfHgh_HfHgh.chine\\ \ \\TfTgh_TfHgh.chine\\ \ tuple_ABC'" using tuple_ABC_eq_ABC' by simp finally show ?thesis by simp qed interpretation g\<^sub>0h\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg \tab\<^sub>1 h\ \tab\<^sub>0 g\ using gh by unfold_locales auto interpretation f\<^sub>0g\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg \tab\<^sub>1 g\ \tab\<^sub>0 f\ using fg by unfold_locales auto interpretation f\<^sub>0gh\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg \tab\<^sub>1 g \ Tgh.\\.p\<^sub>1\ \tab\<^sub>0 f\ using fg gh Tgh.\\.leg1_is_map by unfold_locales auto interpretation fg\<^sub>0h\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \ \ src trg \tab\<^sub>1 h\ \tab\<^sub>0 g \ Tfg.p\<^sub>0\ using TTfgh.r\<^sub>0s\<^sub>1_is_cospan by simp lemma src_tab_eq: shows "(\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_\ \ TTfgh_TfTgh.the_\ = TTfgh.tab" proof - have "TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_\ \ TTfgh_TfTgh.the_\ = (\[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ TTfgh.tab" unfolding TTfgh.tab_def using TTfgh_TfTgh.chine_is_induced_map TTfgh.tab_def TTfgh_TfTgh.\_simps(4) by auto moreover have "iso (\[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0)" by (simp add: fg gh) moreover have "inv (\[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) = \\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0" using fg gh by simp ultimately show ?thesis using TTfgh_TfTgh.\_simps(1) invert_side_of_triangle(1) [of "TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_\ \ TTfgh_TfTgh.the_\" "\[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0" TTfgh.tab] by argo qed text \ We need to show that the associativity isomorphism (defined in terms of tupling) coincides with \TTfgh_TfTgh.chine\ (defined in terms of tabulations). In order to do this, we need to know how the latter commutes with projections. That is the purpose of the following lemma. Unfortunately, it requires some lengthy calculations, which I haven't seen any way to avoid. \ lemma prj_chine: shows "\\TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine\\ = \\Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1\\" and "\\Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\ = \\Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1\\" and "\\Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\ = \\TTfgh.p\<^sub>0\\" proof - have 1: "ide TfTgh.p\<^sub>1" by (simp add: TfTgh.composable) have 2: "ide TTfgh_TfTgh.chine" by simp have 3: "src TfTgh.p\<^sub>1 = trg TTfgh_TfTgh.chine" using TTfgh_TfTgh.chine_in_hom(1) by simp have 4: "src (tab\<^sub>1 f) = trg TfTgh.p\<^sub>1" using TfTgh.leg1_simps(2) by blast text \ The required isomorphisms will each be established via \T2\, using the equation \src_tab_eq\ (associativities omitted from diagram): $$ \begin{array}{l} \xymatrix{ && \xtwocell[dddd]{}\omit{^{\rm the\_}\nu} & \scriptstyle{{\rm TTfgh}.{\rm apex}} \ar[dd]^{{\rm chine}} \ar[dddlll]_{{\rm TfTgh}.p_1} \ar[dddrrr]^{{\rm TfTgh}.p_0} & \xtwocell[dddd]{}\omit{^{\rm the\_}\theta} \\ &&&&& \\ &&& \scriptstyle{{\rm TfTgh.apex}} \ar[ddll]_{{\rm TfTgh}.p_1} \ar[dr]^{{\rm TfTgh}.p_0} && \\ \scriptstyle{f.{\rm apex}} \ar[dd]_{f.{\rm tab}_1} && \dtwocell\omit{^<-7>{f_0gh_1.\phi}} && \scriptstyle{{\rm Tgh.apex}} \ar[dl]_{{\rm Tgh}.p_1} \ar[dr]^{{\rm Tgh}.p_0} \ddtwocell\omit{^{g_0h_1.\phi}} && \scriptstyle{h.{\rm apex}} \ar[dd]^{h.{\rm tab}_0} \\ & \scriptstyle{f.{\rm apex}} \ar[dl]_{f.{\rm tab}_1} \ar[dr]^{f.{\rm tab}_0} \dtwocell\omit{^f.{\rm tab}} && \scriptstyle{g.{\rm apex}} \ar[dl]_{g.{\rm tab}_1} \ar[dr]^{g.{\rm tab}_0} \dtwocell\omit{^g.{\rm tab}} && \scriptstyle{h.{\rm apex}} \ar[dl]_{h.{\rm tab}_1} \ar[dr]^{h.{\rm tab}_0} \dtwocell\omit{^h.{\rm tab}} \\ \scriptstyle{{\rm trg}~f} && \scriptstyle{{\rm src}~f = {\rm trg}~g} \ar[ll]^{f} && \scriptstyle{{\rm src}~g = {\rm trg}~h} \ar[ll]^{g} && \scriptstyle{{\rm src}~h} \ar[ll]^{h} } \\ \\ \hspace{7cm}= \\ \\ \xymatrix{ &&& \scriptstyle{{\rm TTfgh.apex}} \ar[dl]_{{\rm TTfgh}.p_1} \ar[ddrr]^{{\rm TTfgh}.p_0} && \\ && \scriptstyle{{\rm Tfg.apex}} \ar[dl]_{{\rm Tfg}.p_1} \ar[dr]^{{\rm Tfg}.p_0} \ddtwocell\omit{^{f_0g_1.\phi}} & \dtwocell\omit{^<-7>{fg_0h_1.\phi}} &&& \\ & \scriptstyle{f.{\rm apex}} \ar[dl]_{f.{\rm tab}_1} \ar[dr]^{f.{\rm tab}_0} \dtwocell\omit{^f.{\rm tab}} && \scriptstyle{g.{\rm apex}} \ar[dl]_{g.{\rm tab}_1} \ar[dr]^{g.{\rm tab}_0} \dtwocell\omit{^g.{\rm tab}} && \scriptstyle{h.{\rm apex}} \ar[dl]_{h.{\rm tab}_1} \ar[dr]^{h.{\rm tab}_0} \dtwocell\omit{^h.{\rm tab}} \\ \scriptstyle{{\rm trg}~f} && \scriptstyle{{\rm src}~f = {\rm trg}~g} \ar[ll]^{f} && \scriptstyle{{\rm src}~g = {\rm trg}~h} \ar[ll]^{g} && \scriptstyle{{\rm src}~h} \ar[ll]^{h} } \end{array} $$ There is a sequential dependence between the proofs, such as we have already seen for \horizontal_composite_of_arrows_of_tabulations_in_maps.prj_chine\. \ define u\<^sub>f where "u\<^sub>f = g \ h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0" define w\<^sub>f where "w\<^sub>f = Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1" define w\<^sub>f' where "w\<^sub>f' = TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine" define \\<^sub>f where "\\<^sub>f = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" define \\<^sub>f' where "\\<^sub>f' = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]" define \\<^sub>f where "\\<^sub>f = \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" have w\<^sub>f: "ide w\<^sub>f" using w\<^sub>f_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto have w\<^sub>f_is_map: "is_left_adjoint w\<^sub>f" using w\<^sub>f_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by (simp add: left_adjoints_compose) have w\<^sub>f': "ide w\<^sub>f'" unfolding w\<^sub>f'_def by simp have w\<^sub>f'_is_map: "is_left_adjoint w\<^sub>f'" unfolding w\<^sub>f'_def using 3 TTfgh_TfTgh.is_map f\<^sub>0gh\<^sub>1.leg1_is_map by (simp add: left_adjoints_compose) have \\<^sub>f: "\\\<^sub>f : tab\<^sub>0 f \ w\<^sub>f \ u\<^sub>f\" proof (unfold \\<^sub>f_def w\<^sub>f_def u\<^sub>f_def, intro comp_in_homI) show "\\\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] : tab\<^sub>0 f \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1 \ (tab\<^sub>0 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\" using f\<^sub>0g\<^sub>1.leg1_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by auto show "\f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1 : (tab\<^sub>0 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1 \ (tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using f\<^sub>0g\<^sub>1.\_in_hom(2) Tfg.\\.T0.antipar(1) by (intro hcomp_in_vhom, auto) show "\(g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 : (tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ ((g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using Tfg.\\.T0.antipar(1) by (intro hcomp_in_vhom, auto) show "\\[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1 : ((g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ (g \ tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by (intro hcomp_in_vhom, auto) show "\\[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] : (g \ tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ g \ (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto show "\g \ fg\<^sub>0h\<^sub>1.\ : g \ (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ g \ tab\<^sub>1 h \ TTfgh.p\<^sub>0\" using fg\<^sub>0h\<^sub>1.\_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by (intro hcomp_in_vhom, auto) show "\g \ h.tab \ TTfgh.p\<^sub>0 : g \ tab\<^sub>1 h \ TTfgh.p\<^sub>0 \ g \ (h \ tab\<^sub>0 h) \ TTfgh.p\<^sub>0\" using gh fg\<^sub>0h\<^sub>1.\_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by (intro hcomp_in_vhom, auto) show "\g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] : g \ (h \ tab\<^sub>0 h) \ TTfgh.p\<^sub>0 \ g \ h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0\" using gh fg\<^sub>0h\<^sub>1.\_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by (intro hcomp_in_vhom, auto) qed have \\<^sub>f': "\\\<^sub>f' : tab\<^sub>0 f \ w\<^sub>f' \ u\<^sub>f\" proof (unfold \\<^sub>f'_def w\<^sub>f'_def u\<^sub>f_def, intro comp_in_homI) show "\\\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] : tab\<^sub>0 f \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine \ (tab\<^sub>0 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine\" using "1" "2" "3" "4" assoc'_in_hom(2) f.ide_u f.leg1_simps(3) by auto show "\f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine : (tab\<^sub>0 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine \ ((tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.\_in_hom(2) by (intro hcomp_in_vhom, auto) show "\((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine : ((tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ (((g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by (intro hcomp_in_vhom, auto) show "\(\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine : (((g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ ((g \ tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by (intro hcomp_in_vhom, auto) show "\((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine : ((g \ tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ ((g \ tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan g\<^sub>0h\<^sub>1.\_in_hom(2) by (intro hcomp_in_vhom, auto) show "\((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine : ((g \ tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ ((g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by (intro hcomp_in_vhom, auto) show "\can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) : ((g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ g \ h \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by auto show "\g \ h \ TTfgh_TfTgh.the_\ : g \ h \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ g \ h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0\" using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan TTfgh_TfTgh.the_\_in_hom by (intro hcomp_in_vhom, auto) qed have \\<^sub>f: "\\\<^sub>f : tab\<^sub>1 f \ w\<^sub>f \ tab\<^sub>1 f \ w\<^sub>f'\" proof (unfold \\<^sub>f_def w\<^sub>f_def w\<^sub>f'_def, intro comp_in_homI) show "\\\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] : tab\<^sub>1 f \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1 \ (tab\<^sub>1 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\" using TTfgh.leg1_in_hom(2) assoc'_in_hom by auto show "\TTfgh_TfTgh.the_\ : (tab\<^sub>1 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1 \ (tab\<^sub>1 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine\" using TTfgh_TfTgh.the_\_in_hom TTfgh_TfTgh.the_\_props by simp show "\\[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] : (tab\<^sub>1 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine \ tab\<^sub>1 f \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine\" using 1 2 3 4 by auto qed have iso_\\<^sub>f: "iso \\<^sub>f" unfolding \\<^sub>f_def using 1 2 3 4 \\<^sub>f \\<^sub>f_def isos_compose apply (intro isos_compose) apply (metis TTfgh.composable TTfgh.leg1_in_hom(2) Tfg.\\.T0.antipar(2) Tfg.\\.T0.ide_right Tfg.\\.leg1_in_hom(2) Tfg_Hfg.u_simps(3) f.T0.antipar(2) f.T0.ide_right f.ide_leg1 f\<^sub>0g\<^sub>1.cospan g.ide_leg1 h.ide_leg1 h.leg1_simps(4) hcomp_in_vhomE ide_hcomp iso_assoc' tab\<^sub>1_simps(1)) using TTfgh_TfTgh.the_\_props(2) f.ide_leg1 iso_assoc by blast+ have u\<^sub>f: "ide u\<^sub>f" using \\<^sub>f ide_cod by blast have w\<^sub>f_in_hhom: "in_hhom w\<^sub>f (src u\<^sub>f) (src (tab\<^sub>0 f))" using u\<^sub>f w\<^sub>f u\<^sub>f_def w\<^sub>f_def by simp have w\<^sub>f'_in_hhom: "in_hhom w\<^sub>f' (src u\<^sub>f) (src (tab\<^sub>0 f))" using u\<^sub>f w\<^sub>f' w\<^sub>f'_def u\<^sub>f_def by simp have 5: "\!\. \\ : w\<^sub>f \ w\<^sub>f'\ \ \\<^sub>f = tab\<^sub>1 f \ \ \ \\<^sub>f = \\<^sub>f' \ (tab\<^sub>0 f \ \)" proof - have eq\<^sub>f: "f.composite_cell w\<^sub>f \\<^sub>f = f.composite_cell w\<^sub>f' \\<^sub>f' \ \\<^sub>f" proof - text \ I don't see any alternative here to just grinding out the calculation. The idea is to bring \f.composite_cell w\<^sub>f \\<^sub>f\ into a form in which \src_tab_eq\ can be applied to eliminate \\\<^sub>f\ in favor of \\\<^sub>f'\. \ have "f.composite_cell w\<^sub>f \\<^sub>f = ((f \ g \ \[h, tab\<^sub>0 h, fg\<^sub>0h\<^sub>1.p\<^sub>0]) \ (f \ g \ h.tab \ fg\<^sub>0h\<^sub>1.p\<^sub>0) \ (f \ g \ fg\<^sub>0h\<^sub>1.\) \ (f \ \[g, tab\<^sub>0 g \ f\<^sub>0g\<^sub>1.p\<^sub>0, fg\<^sub>0h\<^sub>1.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, f\<^sub>0g\<^sub>1.p\<^sub>0] \ fg\<^sub>0h\<^sub>1.p\<^sub>1) \ (f \ (g.tab \ f\<^sub>0g\<^sub>1.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ fg\<^sub>0h\<^sub>1.p\<^sub>1) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, f\<^sub>0g\<^sub>1.p\<^sub>1, fg\<^sub>0h\<^sub>1.p\<^sub>1])) \ \[f, tab\<^sub>0 f, f\<^sub>0g\<^sub>1.p\<^sub>1 \ fg\<^sub>0h\<^sub>1.p\<^sub>1] \ (f.tab \ f\<^sub>0g\<^sub>1.p\<^sub>1 \ fg\<^sub>0h\<^sub>1.p\<^sub>1)" unfolding w\<^sub>f_def \\<^sub>f_def using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps Tgh.composable whisker_left by simp (* 12 sec, 30 sec cpu *) also have "... = (f \ g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (f \ g \ h.tab \ TTfgh.p\<^sub>0) \ (f \ g \ fg\<^sub>0h\<^sub>1.\) \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" using comp_assoc by simp also have "... = (\[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \\<^sup>-\<^sup>1[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (f \ g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \ (f \ g \ h.tab \ TTfgh.p\<^sub>0) \ (f \ g \ fg\<^sub>0h\<^sub>1.\) \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" proof - have "(\[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \\<^sup>-\<^sup>1[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0]) \ (f \ g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) = f \ g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' by simp thus ?thesis using comp_assoc by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (f \ g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \ (f \ g \ h.tab \ TTfgh.p\<^sub>0) \ (f \ g \ fg\<^sub>0h\<^sub>1.\) \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (\\<^sup>-\<^sup>1[f, g, (h \ tab\<^sub>0 h) \ TTfgh.p\<^sub>0] \ (f \ g \ h.tab \ TTfgh.p\<^sub>0)) \ (f \ g \ fg\<^sub>0h\<^sub>1.\) \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc assoc'_naturality [of f g "\[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]"] by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>1 h \ TTfgh.p\<^sub>0] \ (f \ g \ fg\<^sub>0h\<^sub>1.\)) \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc assoc'_naturality [of f g "h.tab \ TTfgh.p\<^sub>0"] by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ ((f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1]) \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc assoc'_naturality [of f g fg\<^sub>0h\<^sub>1.\] by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \[f, tab\<^sub>0 f \ Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" proof - have "(f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] = \[f, tab\<^sub>0 f \ Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1] = \(\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\Tfg.p\<^sub>1\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\Tfg.p\<^sub>1\<^bold>\ \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>]\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps \'_def \_def by simp also have "... = \\<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\ \<^bold>\ \<^bold>\Tfg.p\<^sub>1\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\Tfg.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\Tfg.p\<^sub>1\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>]\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps by (intro E.eval_eqI, simp_all) also have "... = \[f, tab\<^sub>0 f \ Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps \'_def \_def by simp finally show ?thesis by blast qed thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ \[f, tab\<^sub>1 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" proof - have "(f \ f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \[f, tab\<^sub>0 f \ Tfg.p\<^sub>1, TTfgh.p\<^sub>1] = \[f, tab\<^sub>1 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\_in_hom assoc_naturality [of f f\<^sub>0g\<^sub>1.\ TTfgh.p\<^sub>1] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1)" proof - have "(f \ (g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ \[f, tab\<^sub>1 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] = \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\_in_hom assoc_naturality [of f "g.tab \ Tfg.p\<^sub>0" TTfgh.p\<^sub>1] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ ((f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f.tab Tfg.p\<^sub>1 TTfgh.p\<^sub>1] comp_assoc by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1)) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1)) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ ((f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1)) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) = (f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr whisker_right comp_assoc_assoc' whisker_left [of f "\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]" "\[g, tab\<^sub>0 g, Tfg.p\<^sub>0]"] by simp thus ?thesis using comp_assoc by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ ((f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ (((\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1)) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1)) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ ((f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "((\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1)) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) = (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1" using fg fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' whisker_right [of TTfgh.p\<^sub>1 "\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0]" "\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0]"] by simp thus ?thesis using comp_assoc by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ (\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ ((f \ (g.tab \ Tfg.p\<^sub>0)) \ TTfgh.p\<^sub>1) \ ((f \ f\<^sub>0g\<^sub>1.\) \ TTfgh.p\<^sub>1) \ (\[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ TTfgh.p\<^sub>1) \ ((f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1)) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ (\\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ (\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1)) \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ (f \ (g.tab \ Tfg.p\<^sub>0)) \ (f \ f\<^sub>0g\<^sub>1.\) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps whisker_right comp_assoc by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \[f \ g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ (f \ (g.tab \ Tfg.p\<^sub>0)) \ (f \ f\<^sub>0g\<^sub>1.\) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "\\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1] \ (f \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ \[f, (g \ tab\<^sub>0 g) \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ ((f \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ TTfgh.p\<^sub>1) \ (\[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\g\<^bold>\, (\<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tfg.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tfg.p\<^sub>0\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tfg.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\f\<^bold>\, (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 g\<^bold>\) \<^bold>\ \<^bold>\Tfg.p\<^sub>0\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tfg.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tfg.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh.p\<^sub>1\<^bold>\)\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps \'_def \_def by simp also have "... = \\<^bold>\\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tfg.p\<^sub>0\<^bold>\, \<^bold>\TTfgh.p\<^sub>1\<^bold>\\<^bold>]\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps by (intro E.eval_eqI, auto) also have "... = \[f \ g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps \'_def \_def by simp finally show ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((\[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \\<^sup>-\<^sup>1[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0]) \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \[f \ g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ (f \ (g.tab \ Tfg.p\<^sub>0)) \ (f \ f\<^sub>0g\<^sub>1.\) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(\[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \\<^sup>-\<^sup>1[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0]) \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) = ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' by simp thus ?thesis by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ ((f \ g) \ h.tab \ TTfgh.p\<^sub>0) \ ((f \ g) \ fg\<^sub>0h\<^sub>1.\) \ \[f \ g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \ Tfg.p\<^sub>0] \ (f \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \ (f \ (g.tab \ Tfg.p\<^sub>0)) \ (f \ f\<^sub>0g\<^sub>1.\) \ \[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \ (f.tab \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1)) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ TTfgh.tab \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using TTfgh.tab_def Tfg.\\.tab_def by simp also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ (f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ (g \ h.tab \ Tgh.p\<^sub>0) \ (g \ g\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ TTfgh_TfTgh.the_\) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using src_tab_eq TfTgh.tab_def Tgh.\\.tab_def comp_assoc by simp text \Now we have to make this look like \f.composite_cell w\<^sub>f' \\<^sub>f' \ \\<^sub>f\.\ also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\)) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ TTfgh_TfTgh.the_\) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ (g \ h.tab \ Tgh.p\<^sub>0) \ (g \ g\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 = (f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0)" using fg gh whisker_right whisker_left by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ ((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ (f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine = (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ ((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine)" (* using fg gh whisker_right [of TTfgh_TfTgh.chine] by auto (* 2 min *) *) proof - have "arr (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ (f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using fg gh by (intro seqI' comp_in_homI) auto moreover have "arr ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ (f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ (f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ (f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ (f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ (f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ (f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ (f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr ((f \ f\<^sub>0gh\<^sub>1.\) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast moreover have "arr (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ (f.tab \ TfTgh.p\<^sub>1))" using calculation by blast ultimately show ?thesis using whisker_right [of TTfgh_TfTgh.chine] TTfgh_TfTgh.is_ide by presburger qed thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ (((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] = (f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps comp_arr_dom comp_assoc_assoc' by simp thus ?thesis by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ (((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "((f.tab \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f.tab TfTgh.p\<^sub>1 TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ (\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ ((\\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine]) \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine)) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(\\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine]) \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) = f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' by simp thus ?thesis using comp_assoc by simp qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ ((\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine]) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f \ TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, tab\<^sub>0 f \ TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" proof - have "(\[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f \ tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] = \(\<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\TfTgh.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\TfTgh.p\<^sub>1\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\TfTgh.p\<^sub>1\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using \'_def \_def by simp also have "... = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\tab\<^sub>0 f\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>1\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\tab\<^sub>0 f\<^bold>\, \<^bold>\TfTgh.p\<^sub>1\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>])\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps by (intro E.eval_eqI, auto) also have "... = \\<^sup>-\<^sup>1[f, tab\<^sub>0 f \ TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" using \'_def \_def by simp finally show ?thesis by simp qed thus ?thesis using comp_assoc by presburger qed also have "... = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f \ TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - (* * This one can't be shortcut with a straight coherence-based proof, * due to the presence of f\<^sub>0gh\<^sub>1.\, g\<^sub>0h\<^sub>1.\, h.tab, with associativities that * do not respect their domain and codomain. * * I also tried to avoid distributing the "f \" in advance, in order to * reduce the number of associativity proof steps, but it then becomes * less automatic to prove the necessary "arr" facts to do the proof. * So unfortunately the mindless grind seems to be the path of least * resistance. *) have "((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ f\<^sub>0gh\<^sub>1.\) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, tab\<^sub>0 f \ TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] = ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f f\<^sub>0gh\<^sub>1.\ TTfgh_TfTgh.chine] comp_assoc by simp also have "... = ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, ((g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - have "((f \ (g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, ((g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f "(g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0" TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - have "((f \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, ((g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f "\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0" TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = ((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - have "((f \ (g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f "(g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0" TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = (((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - have "((f \ (g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[f, (g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f "(g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0" TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)" proof - (* OK, we can perhaps shortcut the last few steps... *) have "((f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \ TTfgh_TfTgh.chine) \ ((f \ \\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((f \ (g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \((\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ ((\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ ((\<^bold>\f\<^bold>\ \<^bold>\ (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, (\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps \'_def \_def by simp also have "... = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ ((\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)\" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps apply (intro E.eval_eqI) by simp_all also have "... = \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps \'_def \_def by simp finally show ?thesis using comp_assoc by presburger qed finally show ?thesis using comp_assoc by presburger qed thus ?thesis using comp_assoc by presburger qed also have "... = (\[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\)) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "(f \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ (f \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) = (f \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" using fg gh whisker_left by simp (* 15 sec elapsed, 30 sec cpu *) thus ?thesis using comp_assoc by presburger qed also have "... = (f \ g \ h \ TTfgh_TfTgh.the_\) \ (\[f, g, h \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ \[f \ g, h, ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)) \ (f \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "\[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ \[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) = \[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (\[f \ g, h, tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ (((f \ g) \ h) \ TTfgh_TfTgh.the_\)) \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" proof - have "(\\<^sup>-\<^sup>1[f, g, h] \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ ((f \ g \ h) \ TTfgh_TfTgh.the_\) = \\<^sup>-\<^sup>1[f, g, h] \ TTfgh_TfTgh.the_\" using fg gh comp_arr_dom comp_cod_arr interchange [of "\\<^sup>-\<^sup>1[f, g, h]" "f \ g \ h" "tab\<^sub>0 h \ TTfgh.p\<^sub>0" TTfgh_TfTgh.the_\] by simp also have "... = (((f \ g) \ h) \ TTfgh_TfTgh.the_\) \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh comp_arr_dom comp_cod_arr interchange [of "(f \ g) \ h" "\\<^sup>-\<^sup>1[f, g, h]" TTfgh_TfTgh.the_\ "((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine"] by simp finally show ?thesis using comp_assoc by presburger qed also have "... = (\[f, g, h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0] \ ((f \ g) \ h \ TTfgh_TfTgh.the_\)) \ \[f \ g, h, ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh assoc_naturality [of "f \ g" h TTfgh_TfTgh.the_\] comp_assoc by simp also have "... = (f \ g \ h \ TTfgh_TfTgh.the_\) \ \[f, g, h \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ \[f \ g, h, ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using fg gh assoc_naturality [of f g "h \ TTfgh_TfTgh.the_\"] comp_assoc by simp finally show ?thesis using comp_assoc by presburger qed also have "... = ((f \ g \ h \ TTfgh_TfTgh.the_\) \ (f \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ (f \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "\[f, g, h \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ \[f \ g, h, ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) = f \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" proof - have "\[f, g, h \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ \[f \ g, h, ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g, h] \ ((tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \[f \ g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[f, g \ h, (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[f, (g \ h) \ (tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f \ \[g \ h, tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \ TTfgh_TfTgh.chine) \ (f \ (\\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \ Tgh.p\<^sub>0] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f \ ((g \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) = \\<^bold>\\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\, ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\f\<^bold>\, (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\ (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ ((\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)\" using \'_def \_def by simp also have "... = can (\<^bold>\f\<^bold>\ \<^bold>\ (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) (\<^bold>\f\<^bold>\ \<^bold>\ (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\))" using fg gh apply (unfold can_def) apply (intro E.eval_eqI) by simp_all also have "... = f \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" using fg gh whisker_can_left_0 by simp finally show ?thesis by blast qed thus ?thesis using comp_assoc by presburger qed also have "... = (f \ (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \ \[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine] \ (f.tab \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]" proof - have "((f \ g \ h \ TTfgh_TfTgh.the_\) \ (f \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ (f \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) = f \ (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]" proof - have 1: "arr ((g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" using fg gh apply (intro seqI' comp_in_homI) by auto moreover have 2: "arr (can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])" using calculation by blast ultimately show ?thesis using whisker_left f.ide_base by presburger qed thus ?thesis using comp_assoc by presburger qed also have "... = f.composite_cell w\<^sub>f' \\<^sub>f' \ \\<^sub>f" unfolding w\<^sub>f'_def \\<^sub>f'_def \\<^sub>f_def using comp_assoc by presburger finally show ?thesis by blast qed show ?thesis using w\<^sub>f w\<^sub>f' \\<^sub>f \\<^sub>f' \\<^sub>f f.T2 [of w\<^sub>f w\<^sub>f' \\<^sub>f u\<^sub>f \\<^sub>f' \\<^sub>f] eq\<^sub>f by fast qed obtain \\<^sub>f where \\<^sub>f: "\\\<^sub>f : w\<^sub>f \ w\<^sub>f'\ \ \\<^sub>f = tab\<^sub>1 f \ \\<^sub>f \ \\<^sub>f = \\<^sub>f' \ (tab\<^sub>0 f \ \\<^sub>f)" using 5 by auto show "\\TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine\\ = \\Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1\\" proof - have "iso \\<^sub>f" using \\<^sub>f BS3 w\<^sub>f_is_map w\<^sub>f'_is_map by blast hence "isomorphic w\<^sub>f w\<^sub>f'" using \\<^sub>f isomorphic_def isomorphic_symmetric by auto thus ?thesis using w\<^sub>f w\<^sub>f_def w\<^sub>f'_def Maps.CLS_eqI isomorphic_symmetric by auto qed text \ On to the next equation: \[ \\\Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\ = \\Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1\\\. \] We have to make use of the equation \\\<^sub>f = \\<^sub>f' \ (tab\<^sub>0 f \ \\<^sub>f)\ in this part, similarly to how the equation \src_tab_eq\ was used to replace \TTfgh.tab\ in the first part. \ define u\<^sub>g where "u\<^sub>g = h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0" define w\<^sub>g where "w\<^sub>g = Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1" define w\<^sub>g' where "w\<^sub>g' = Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" define \\<^sub>g where "\\<^sub>g = \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" define \\<^sub>g' where "\\<^sub>g' = (h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" define \\<^sub>g where "\\<^sub>g = \[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (tab\<^sub>0 f \ \\<^sub>f) \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" have u\<^sub>g: "ide u\<^sub>g" unfolding u\<^sub>g_def by simp have w\<^sub>g: "ide w\<^sub>g" unfolding w\<^sub>g_def using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto have w\<^sub>g_is_map: "is_left_adjoint w\<^sub>g" unfolding w\<^sub>g_def using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps left_adjoints_compose by simp have w\<^sub>g': "ide w\<^sub>g'" unfolding w\<^sub>g'_def by simp have w\<^sub>g'_is_map: "is_left_adjoint w\<^sub>g'" unfolding w\<^sub>g'_def using TTfgh_TfTgh.is_map left_adjoints_compose by simp have \\<^sub>g: "\\\<^sub>g : tab\<^sub>0 g \ w\<^sub>g \ u\<^sub>g\" using w\<^sub>g_def u\<^sub>g_def \\<^sub>g_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps fg\<^sub>0h\<^sub>1.\_in_hom by auto have \\<^sub>g': "\\\<^sub>g' : tab\<^sub>0 g \ w\<^sub>g' \ u\<^sub>g\" unfolding w\<^sub>g'_def u\<^sub>g_def \\<^sub>g'_def by (intro comp_in_homI) auto have w\<^sub>g_in_hhom: "in_hhom w\<^sub>g (src u\<^sub>g) (src (tab\<^sub>0 g))" unfolding w\<^sub>g_def u\<^sub>g_def by auto have w\<^sub>g'_in_hhom: "in_hhom w\<^sub>g' (src u\<^sub>g) (src (tab\<^sub>0 g))" unfolding w\<^sub>g'_def u\<^sub>g_def by auto have \\<^sub>g: "\\\<^sub>g : tab\<^sub>1 g \ w\<^sub>g \ tab\<^sub>1 g \ w\<^sub>g'\" proof (unfold \\<^sub>g_def w\<^sub>g_def, intro comp_in_homI) (* auto can solve this, but it's too slow *) show "\\\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] : tab\<^sub>1 g \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1 \ (tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto show "\inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1 : (tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ (tab\<^sub>0 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\_in_hom f\<^sub>0g\<^sub>1.\_uniqueness(2) by (intro hcomp_in_vhom) auto show "\\[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] : (tab\<^sub>0 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1 \ tab\<^sub>0 f \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps \\<^sub>f w\<^sub>f_def w\<^sub>f'_def by auto show "\tab\<^sub>0 f \ \\<^sub>f : tab\<^sub>0 f \ Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1 \ tab\<^sub>0 f \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps \\<^sub>f w\<^sub>f_def w\<^sub>f'_def by auto show "\\\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] : tab\<^sub>0 f \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine \ (tab\<^sub>0 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine\" by auto show "\f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine : (tab\<^sub>0 f \ TfTgh.p\<^sub>1) \ TTfgh_TfTgh.chine \ ((tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine\" using f\<^sub>0gh\<^sub>1.\_in_hom by (intro hcomp_in_vhom) auto show "\\[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] : ((tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine \ (tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\" by auto show "\\[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] : (tab\<^sub>1 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine \ tab\<^sub>1 g \ w\<^sub>g'\" using w\<^sub>g'_def by auto qed have eq\<^sub>g: "g.composite_cell w\<^sub>g \\<^sub>g = g.composite_cell w\<^sub>g' \\<^sub>g' \ \\<^sub>g" proof - have "g.composite_cell w\<^sub>g \\<^sub>g = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1] \ (g.tab \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1)" unfolding w\<^sub>g_def \\<^sub>g_def by simp also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ ((g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1]) \ (g.tab \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1)" using fg gh f\<^sub>0g\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps whisker_left comp_assoc by simp also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ (\[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]))) \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1] \ (g.tab \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1)" proof - have "(\[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) = g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' by simp thus ?thesis by (simp add: comp_assoc) qed also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ \[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1]) \ (g.tab \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1)" using comp_assoc by presburger also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ (\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (g.tab \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1))" using fg gh f\<^sub>0g\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc pentagon' invert_opposite_sides_of_square [of "\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1" "(\\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ (g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1])" "\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" "\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1]"] by simp also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps assoc'_naturality [of g.tab Tfg.p\<^sub>0 TTfgh.p\<^sub>1] by simp also have "... = (g \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \ (g \ h.tab \ TTfgh.p\<^sub>0) \ (g \ fg\<^sub>0h\<^sub>1.\) \ \[g, tab\<^sub>0 g \ Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ (\[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \ TTfgh.p\<^sub>1) \ ((g.tab \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ (f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" proof - have "(f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] = ((f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (\\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1)) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using comp_assoc by presburger also have "... = ((f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ ((tab\<^sub>0 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1) \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1)) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps whisker_right comp_assoc_assoc' by simp also have "... = ((f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1)) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\_uniqueness comp_cod_arr by simp also have "... = ((tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" proof - have "(f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) = f\<^sub>0g\<^sub>1.\ \ inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1" using f\<^sub>0g\<^sub>1.\_uniqueness whisker_right by simp also have "... = (tab\<^sub>1 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1" using f\<^sub>0g\<^sub>1.\_uniqueness comp_arr_inv' by simp finally show ?thesis by simp qed also have "... = \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr by simp finally have "(f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] = \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sub>f \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" unfolding \\<^sub>f_def using comp_assoc by presburger also have "... = \\<^sub>f' \ (tab\<^sub>0 f \ \\<^sub>f) \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" using \\<^sub>f comp_assoc by simp also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (tab\<^sub>0 f \ \\<^sub>f) \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" unfolding \\<^sub>f'_def using comp_assoc by presburger also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (\\<^sup>-\<^sup>1[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine)) \ \\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \ (tab\<^sub>0 f \ \\<^sub>f) \ \[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \ (inv f\<^sub>0g\<^sub>1.\ \ TTfgh.p\<^sub>1) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" proof - have "(\\<^sup>-\<^sup>1[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ (f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine) = f\<^sub>0gh\<^sub>1.\ \ TTfgh_TfTgh.chine" using f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_cod_arr comp_arr_dom comp_assoc_assoc' by simp thus ?thesis using comp_assoc by fastforce qed also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sub>g" unfolding \\<^sub>g_def using comp_assoc by presburger also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \\<^sub>g" proof - have "(((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of "(g.tab \ Tgh.p\<^sub>1)" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sub>g" proof - have "((g.tab \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of g.tab Tgh.p\<^sub>1 "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp thus ?thesis using comp_assoc by presburger qed also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ ((\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ \\<^sub>g" proof - have "(\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) = g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" using comp_cod_arr comp_assoc_assoc' by simp thus ?thesis using comp_assoc g\<^sub>0h\<^sub>1.\_in_hom by simp qed also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ ((\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ \\<^sub>g" proof - have "(\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) = (\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((g \ (tab\<^sub>0 g \ Tgh.p\<^sub>1)) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_assoc comp_assoc_assoc' by simp also have "... = (\\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_cod_arr comp_assoc_assoc' by simp also have "... = (((g \ (tab\<^sub>0 g \ Tgh.p\<^sub>1)) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine)" using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps whisker_right comp_assoc_assoc' by simp also have "... = (\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine" using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_cod_arr by simp finally show ?thesis by presburger qed also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ ((((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (\\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ \\<^sub>g" using comp_assoc by presburger also have "... = (g \ h \ TTfgh_TfTgh.the_\) \ (can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (g \ (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (\[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sub>g" proof - have "(((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g \ (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" proof - have "(((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ (((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ ((((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" using comp_assoc by presburger also have "... = ((((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" proof - have "(((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g \ tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of "g \ g\<^sub>0h\<^sub>1.\" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ ((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" proof - have "(((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g \ tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of "g \ h.tab \ Tgh.p\<^sub>0" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ (((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g, tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" proof - have "((g \ g\<^sub>0h\<^sub>1.\) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g, tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of g g\<^sub>0h\<^sub>1.\ "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp thus ?thesis using comp_assoc by presburger qed also have "... = \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g \ (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" proof - have "((g \ h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[g, tab\<^sub>1 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g \ (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps assoc'_naturality [of g "h.tab \ Tgh.p\<^sub>0" "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp thus ?thesis using comp_assoc by presburger qed finally show ?thesis by simp qed thus ?thesis using comp_assoc by presburger qed also have "... = ((g \ h \ TTfgh_TfTgh.the_\) \ (g \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ (g \ (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine])) \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sub>g" proof - have "can (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = g \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" proof - have "\\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = can (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) ((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" proof - have "\\<^sup>-\<^sup>1[g \ (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps canI_associator_0 \'_def \_def by simp also have "... = can (((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) ((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" unfolding can_def using gh apply (intro E.eval_eqI) by simp_all finally show ?thesis by blast qed moreover have "\\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = can ((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\g\<^bold>\ \<^bold>\ ((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" proof - have "\\<^sup>-\<^sup>1[g, (h \ tab\<^sub>0 h) \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps canI_associator_0 \'_def \_def by simp also have "... = can ((\<^bold>\g\<^bold>\ \<^bold>\ (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\g\<^bold>\ \<^bold>\ ((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" unfolding can_def using gh apply (intro E.eval_eqI) by simp_all finally show ?thesis by blast qed ultimately show ?thesis using gh whisker_can_left_0 by simp qed moreover have "\[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" proof - have "\[g, tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[g \ tab\<^sub>0 g \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ ((\[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \ TfTgh.p\<^sub>0) \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[(g \ tab\<^sub>0 g) \ Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g \ tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] = \\<^bold>\\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>1\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 g\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>1\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tgh.p\<^sub>1\<^bold>\\<^bold>] \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[(\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 g\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>1\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tgh.p\<^sub>1\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\g\<^bold>\, \<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tgh.p\<^sub>1\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using gh g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps \'_def \_def by simp also have "... = \\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\tab\<^sub>0 g\<^bold>\, \<^bold>\Tgh.p\<^sub>1\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" apply (intro E.eval_eqI) by simp_all also have "... = g \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" using gh g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps \'_def \_def by simp finally show ?thesis by simp qed ultimately show ?thesis using comp_assoc by presburger qed also have "... = (g \ (h \ TTfgh_TfTgh.the_\) \ (can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g.tab \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sub>g" using gh whisker_left by auto (* 11 sec *) also have "... = g.composite_cell w\<^sub>g' \\<^sub>g' \ \\<^sub>g" unfolding w\<^sub>g'_def \\<^sub>g'_def using comp_assoc by presburger finally show ?thesis by blast qed have 6: "\!\. \\ : w\<^sub>g \ w\<^sub>g'\ \ \\<^sub>g = tab\<^sub>1 g \ \ \ \\<^sub>g = \\<^sub>g' \ (tab\<^sub>0 g \ \)" using w\<^sub>g w\<^sub>g' \\<^sub>g \\<^sub>g' \\<^sub>g eq\<^sub>g g.T2 [of w\<^sub>g w\<^sub>g' \\<^sub>g u\<^sub>g \\<^sub>g' \\<^sub>g] by blast obtain \\<^sub>g where \\<^sub>g: "\\\<^sub>g : w\<^sub>g \ w\<^sub>g'\ \ \\<^sub>g = tab\<^sub>1 g \ \\<^sub>g \ \\<^sub>g = \\<^sub>g' \ (tab\<^sub>0 g \ \\<^sub>g)" using 6 by auto show "\\Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\ = \\Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1\\" proof - have "iso \\<^sub>g" using \\<^sub>g BS3 w\<^sub>g_is_map w\<^sub>g'_is_map by blast hence "isomorphic w\<^sub>g w\<^sub>g'" using \\<^sub>g isomorphic_def isomorphic_symmetric by auto thus ?thesis using w\<^sub>g w\<^sub>g' w\<^sub>g_def w\<^sub>g'_def Maps.CLS_eqI by auto qed text \Now the last equation: similar, but somewhat simpler.\ define u\<^sub>h where "u\<^sub>h = tab\<^sub>0 h \ TTfgh.p\<^sub>0" define w\<^sub>h where "w\<^sub>h = TTfgh.p\<^sub>0" define w\<^sub>h' where "w\<^sub>h' = Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" define \\<^sub>h where "\\<^sub>h = tab\<^sub>0 h \ TTfgh.p\<^sub>0" define \\<^sub>h' where "\\<^sub>h' = TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" define \\<^sub>h where "\\<^sub>h = \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" have u\<^sub>h: "ide u\<^sub>h" unfolding u\<^sub>h_def by simp have w\<^sub>h: "ide w\<^sub>h" unfolding w\<^sub>h_def by simp have w\<^sub>h_is_map: "is_left_adjoint w\<^sub>h" unfolding w\<^sub>h_def by simp have w\<^sub>h': "ide w\<^sub>h'" unfolding w\<^sub>h'_def by simp have w\<^sub>h'_is_map: "is_left_adjoint w\<^sub>h'" unfolding w\<^sub>h'_def using g\<^sub>0h\<^sub>1.p\<^sub>0_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps TTfgh_TfTgh.is_map left_adjoints_compose by simp have \\<^sub>h: "\\\<^sub>h : tab\<^sub>0 h \ w\<^sub>h \ u\<^sub>h\" unfolding \\<^sub>h_def w\<^sub>h_def u\<^sub>h_def by auto have \\<^sub>h': "\\\<^sub>h' : tab\<^sub>0 h \ w\<^sub>h' \ u\<^sub>h\" unfolding \\<^sub>h'_def w\<^sub>h'_def u\<^sub>h_def using g\<^sub>0h\<^sub>1.p\<^sub>0_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps by (intro comp_in_homI) auto have \\<^sub>h: "\\\<^sub>h : tab\<^sub>1 h \ w\<^sub>h \ tab\<^sub>1 h \ w\<^sub>h'\" proof (unfold \\<^sub>h_def w\<^sub>h_def w\<^sub>h'_def, intro comp_in_homI) (* auto can solve this, but it's too slow *) show "\inv fg\<^sub>0h\<^sub>1.\ : tab\<^sub>1 h \ TTfgh.p\<^sub>0 \ (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.\_uniqueness by blast show "\\[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] : (tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1 \ tab\<^sub>0 g \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto show "\tab\<^sub>0 g \ \\<^sub>g : tab\<^sub>0 g \ Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1 \ tab\<^sub>0 g \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\" using \\<^sub>g w\<^sub>g_def w\<^sub>g'_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto show "\\\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] : tab\<^sub>0 g \ Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine \ (tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto show "\g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine : (tab\<^sub>0 g \ Tgh.p\<^sub>1) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine \ (tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by force show "\\[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] : (tab\<^sub>1 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine \ tab\<^sub>1 h \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto qed have eq\<^sub>h: "h.composite_cell w\<^sub>h \\<^sub>h = h.composite_cell w\<^sub>h' \\<^sub>h' \ \\<^sub>h" proof - text \ Once again, the strategy is to form the subexpression \[ \\[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]\ \] which is equal to \\\<^sub>g\, so that we can make use of the equation \\\<^sub>g = \\<^sub>g' \ (tab\<^sub>0 g \ \\<^sub>g)\. \ have "h.composite_cell w\<^sub>h \\<^sub>h = (h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0)" unfolding w\<^sub>h_def \\<^sub>h_def by simp also have "... = \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0)" proof - have "(h \ tab\<^sub>0 h \ TTfgh.p\<^sub>0) \ \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] = \[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]" using comp_cod_arr by simp thus ?thesis using comp_assoc by metis qed also have "... = (\[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \ (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" proof - have "(h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ (\\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ inv fg\<^sub>0h\<^sub>1.\ = (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ ((tab\<^sub>0 g \ Tfg.p\<^sub>0) \ TTfgh.p\<^sub>1) \ inv fg\<^sub>0h\<^sub>1.\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc_assoc' by simp also have "... = (h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ inv fg\<^sub>0h\<^sub>1.\" using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps fg\<^sub>0h\<^sub>1.\_uniqueness comp_cod_arr by simp also have "... = (h.tab \ TTfgh.p\<^sub>0) \ (tab\<^sub>1 h \ TTfgh.p\<^sub>0)" using comp_arr_inv' fg\<^sub>0h\<^sub>1.\_uniqueness by simp also have "... = h.tab \ TTfgh.p\<^sub>0" using comp_arr_dom fg\<^sub>0h\<^sub>1.p\<^sub>0_simps by simp finally have "(h.tab \ TTfgh.p\<^sub>0) \ fg\<^sub>0h\<^sub>1.\ \ (\\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \ inv fg\<^sub>0h\<^sub>1.\ = h.tab \ TTfgh.p\<^sub>0" by simp thus ?thesis using comp_assoc by simp qed also have "... = \\<^sub>g \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" unfolding \\<^sub>g_def by simp also have "... = (\\<^sub>g' \ (tab\<^sub>0 g \ \\<^sub>g)) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" using \\<^sub>g by simp also have "... = (h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" unfolding \\<^sub>g'_def using comp_assoc by presburger also have "... = (h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ ((\\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" proof - have "(\\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) = (h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" using comp_cod_arr comp_assoc_assoc' by simp thus ?thesis using comp_assoc by simp qed also have "... = (h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (\[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ ((h.tab \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" using comp_assoc by presburger also have "... = (h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" using assoc_naturality [of h.tab Tgh.p\<^sub>0 "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] comp_assoc by simp also have "... = (h \ TTfgh_TfTgh.the_\) \ can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ ((\\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine)) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" proof - have "(\\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) = h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" using comp_cod_arr comp_assoc_assoc' by simp thus ?thesis using comp_assoc by simp qed also have "... = (h \ TTfgh_TfTgh.the_\) \ (can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (\\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine])) \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" using comp_assoc by presburger also have "... = ((h \ TTfgh_TfTgh.the_\) \ (h \ can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\))) \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" proof - have "can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (\\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) = can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" using \'_def \_def by simp also have "... = can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ can (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" proof - have "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\h\<^bold>\, \<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\ = can (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" unfolding can_def apply (intro E.eval_eqI) by simp_all thus ?thesis by simp qed also have "... = can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" by simp also have "... = h \ can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" using whisker_can_left_0 by simp finally have "can (\<^bold>\h\<^bold>\ \<^bold>\ ((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (((\<^bold>\h\<^bold>\ \<^bold>\ \<^bold>\tab\<^sub>0 h\<^bold>\) \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) \ (\\<^sup>-\<^sup>1[h \ tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) = h \ can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)" by simp thus ?thesis using comp_assoc by presburger qed also have "... = (h \ TTfgh_TfTgh.the_\ \ can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\)) \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" using whisker_left [of h] comp_assoc by simp also have "... = (h \ TTfgh_TfTgh.the_\ \ \\<^sup>-\<^sup>1[tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]) \ \[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (h.tab \ Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (g\<^sub>0h\<^sub>1.\ \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine) \ \\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine] \ (tab\<^sub>0 g \ \\<^sub>g) \ \[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \ inv fg\<^sub>0h\<^sub>1.\" proof - have "can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\, \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\tab\<^sub>0 h\<^bold>\, \<^bold>\Tgh.p\<^sub>0\<^bold>\, \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\\<^bold>]\" unfolding can_def apply (intro E.eval_eqI) by auto also have "... = \\<^sup>-\<^sup>1[tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" using \'_def \_def by simp finally have "can (((\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\) \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) (\<^bold>\tab\<^sub>0 h\<^bold>\ \<^bold>\ \<^bold>\Tgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TfTgh.p\<^sub>0\<^bold>\ \<^bold>\ \<^bold>\TTfgh_TfTgh.chine\<^bold>\) = \\<^sup>-\<^sup>1[tab\<^sub>0 h \ Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \ \\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine]" by simp thus ?thesis by simp qed also have "... = h.composite_cell w\<^sub>h' \\<^sub>h' \ \\<^sub>h" unfolding w\<^sub>h'_def \\<^sub>h'_def \\<^sub>h_def using comp_assoc by presburger finally show ?thesis by simp qed have 7: "\!\. \\ : w\<^sub>h \ w\<^sub>h'\ \ \\<^sub>h = tab\<^sub>1 h \ \ \ \\<^sub>h = \\<^sub>h' \ (tab\<^sub>0 h \ \)" using w\<^sub>h w\<^sub>h' \\<^sub>h \\<^sub>h' \\<^sub>h eq\<^sub>h h.T2 [of w\<^sub>h w\<^sub>h' \\<^sub>h u\<^sub>h \\<^sub>h' \\<^sub>h] by blast obtain \\<^sub>h where \\<^sub>h: "\\\<^sub>h : w\<^sub>h \ w\<^sub>h'\ \ \\<^sub>h = tab\<^sub>1 h \ \\<^sub>h \ \\<^sub>h = \\<^sub>h' \ (tab\<^sub>0 h \ \\<^sub>h)" using 7 by auto show "\\Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\ = \\TTfgh.p\<^sub>0\\" proof - have "iso \\<^sub>h" using \\<^sub>h BS3 w\<^sub>h_is_map w\<^sub>h'_is_map by blast hence "isomorphic w\<^sub>h w\<^sub>h'" using \\<^sub>h isomorphic_def isomorphic_symmetric by auto thus ?thesis using w\<^sub>h w\<^sub>h' w\<^sub>h_def w\<^sub>h'_def Maps.CLS_eqI [of w\<^sub>h w\<^sub>h'] by simp qed qed text \ Finally, we can show that @{term TTfgh_TfTgh.chine} is given by tupling. \ lemma CLS_chine: shows "\\TTfgh_TfTgh.chine\\ = tuple_ABC" proof - have "tuple_ABC = SPN_fgh.chine_assoc" using SPN_fgh.chine_assoc_def by simp also have "... = \\TTfgh_TfTgh.chine\\" proof (intro Maps.arr_eqI) show "Maps.arr SPN_fgh.chine_assoc" using SPN_fgh.chine_assoc_in_hom by auto show "Maps.arr \\TTfgh_TfTgh.chine\\" using Maps.CLS_in_hom TTfgh_TfTgh.is_map by blast show "Maps.Dom SPN_fgh.chine_assoc = Maps.Dom \\TTfgh_TfTgh.chine\\" using SPN_fgh.chine_assoc_def Maps.dom_char tuple_ABC_in_hom TTfgh_TfTgh.chine_in_hom by fastforce show "Maps.Cod SPN_fgh.chine_assoc = Maps.Cod \\TTfgh_TfTgh.chine\\" proof - have "Maps.Cod SPN_fgh.chine_assoc = Maps.Cod tuple_ABC" using SPN_fgh.chine_assoc_def by simp also have "... = src (prj\<^sub>0 (tab\<^sub>1 g \ prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)) (tab\<^sub>0 f))" by (metis (lifting) Maps.Dom.simps(1) Maps.seq_char SPN_fgh.prj_chine_assoc(1) SPN_fgh.prj_simps(1) calculation f\<^sub>0gh\<^sub>1.leg1_simps(3) prj_char(4)) also have "... = Maps.Cod \\TTfgh_TfTgh.chine\\" using Maps.cod_char TTfgh_TfTgh.chine_in_hom by simp finally show ?thesis by blast qed show "Maps.Map SPN_fgh.chine_assoc = Maps.Map \\TTfgh_TfTgh.chine\\" proof - have 0: "Chn (Span.hcomp (SPN f) (Span.hcomp (SPN g) (SPN h))) = Maps.MkIde (src TfTgh.p\<^sub>0)" using fg gh by (metis (mono_tags, lifting) Maps.in_homE Maps.seqE SPN_fgh.prj_chine_assoc(1) SPN_fgh.prj_simps(1) SPN_fgh.prj_simps(13) calculation tuple_ABC_in_hom) have "tuple_ABC = \\TTfgh_TfTgh.chine\\" proof (intro Maps.prj_joint_monic [of SPN_fgh.\.leg0 "SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1" tuple_ABC "\\TTfgh_TfTgh.chine\\"]) show "Maps.cospan SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" using SPN_fgh.\\.dom.is_span SPN_fgh.\\.leg1_composite SPN_fgh.cospan_\\ by auto show "Maps.seq SPN_fgh.Prj\<^sub>1 tuple_ABC" using 0 tuple_ABC_in_hom SPN_fgh.prj_in_hom(4) by auto show "Maps.seq SPN_fgh.Prj\<^sub>1 \\TTfgh_TfTgh.chine\\" proof show "Maps.in_hom \\TTfgh_TfTgh.chine\\ \\src TTfgh_TfTgh.chine\\ \\trg TTfgh_TfTgh.chine\\" using fg gh TTfgh_TfTgh.chine_in_hom Maps.CLS_in_hom TTfgh_TfTgh.is_map by blast show "Maps.in_hom SPN_fgh.Prj\<^sub>1 \\trg TTfgh_TfTgh.chine\\ SPN_fgh.\.apex" proof show "Maps.cospan SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" using SPN_fgh.prj_in_hom(4) by blast show "\\trg TTfgh_TfTgh.chine\\ = Maps.pbdom SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" proof - have "\\trg TTfgh_TfTgh.chine\\ = Maps.MkIde (src TfTgh.p\<^sub>0)" by simp also have "... = Maps.pbdom SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" using 0 Maps.pbdom_def SPN_fgh.chine_composite(2) by presburger finally show ?thesis by blast qed show "SPN_fgh.\.apex = Maps.dom SPN_fgh.\.leg0" using SPN_f.dom.apex_def by blast qed qed have 2: "Maps.commutative_square SPN_fgh.\.leg0 SPN_fgh.\.leg1 SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0" proof show "Maps.cospan SPN_fgh.\.leg0 SPN_fgh.\.leg1" using SPN_fgh.\\.legs_form_cospan(1) by simp show "Maps.span SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0" using SPN_fgh.prj_simps(2-3,5-6) by presburger show "Maps.dom SPN_fgh.\.leg0 = Maps.cod SPN_fgh.Prj\<^sub>0\<^sub>1" using SPN_fgh.prj_simps(8) SPN_g.dom.is_span SPN_g.dom.leg_simps(2) by auto show "SPN_fgh.\.leg0 \ SPN_fgh.Prj\<^sub>0\<^sub>1 = SPN_fgh.\.leg1 \ SPN_fgh.Prj\<^sub>0" by (metis (no_types, lifting) Maps.cod_comp Maps.comp_assoc Maps.pullback_commutes' SPN_fgh.\\.dom.leg_simps(1) SPN_fgh.\\.leg0_composite SPN_fgh.cospan_\\) qed have 1: "Maps.commutative_square SPN_fgh.\.leg0 (Maps.comp SPN_fgh.\.leg1 SPN_fgh.\\.prj\<^sub>1) SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC" proof show "Maps.cospan SPN_fgh.\.leg0 (Maps.comp SPN_fgh.\.leg1 SPN_fgh.\\.prj\<^sub>1)" using fg gh SPN_fgh.prj_simps(10) by blast show "Maps.span SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC" using fg gh csq(2) by blast show "Maps.dom SPN_fgh.\.leg0 = Maps.cod SPN_fgh.Prj\<^sub>1\<^sub>1" using fg gh SPN_f.dom.leg_simps(2) SPN_fgh.prj_simps(7) by auto show "SPN_fgh.\.leg0 \ SPN_fgh.Prj\<^sub>1\<^sub>1 = (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ tuple_BC" using 2 fg gh Maps.comp_assoc csq(2) Maps.prj_tuple [of SPN_fgh.\.leg0 SPN_fgh.\.leg1 SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0] by blast qed show "SPN_fgh.Prj\<^sub>1 \ tuple_ABC = SPN_fgh.Prj\<^sub>1 \ Maps.CLS TTfgh_TfTgh.chine" proof - have "SPN_fgh.Prj\<^sub>1 \ tuple_ABC = SPN_fgh.Prj\<^sub>1\<^sub>1" using csq(2) by simp also have "... = \\Tfg.p\<^sub>1 \ TTfgh.p\<^sub>1\\" using prj_char by simp also have "... = \\TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine\\" using prj_chine(1) by simp also have "... = \\TfTgh.p\<^sub>1\\ \ \\TTfgh_TfTgh.chine\\" proof - have "is_left_adjoint TfTgh.p\<^sub>1" by (simp add: fg) moreover have "is_left_adjoint TTfgh_TfTgh.chine" using TTfgh_TfTgh.is_map by simp moreover have "TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine \ TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine" using fg gh isomorphic_reflexive by simp ultimately show ?thesis using Maps.comp_CLS [of TfTgh.p\<^sub>1 TTfgh_TfTgh.chine "TfTgh.p\<^sub>1 \ TTfgh_TfTgh.chine"] by simp qed also have "... = SPN_fgh.Prj\<^sub>1 \ Maps.CLS TTfgh_TfTgh.chine" using prj_char by simp finally show ?thesis by blast qed show "Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ tuple_ABC = Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\" proof - have "Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ tuple_ABC = tuple_BC" using csq(2) Maps.prj_tuple [of SPN_fgh.\.leg0 "SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1" SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC] by simp also have "... = Maps.comp (Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (Maps.comp SPN_fgh.\.leg1 SPN_fgh.\\.prj\<^sub>1)) \\TTfgh_TfTgh.chine\\" proof (intro Maps.prj_joint_monic [of SPN_fgh.\.leg0 SPN_fgh.\.leg1 tuple_BC "Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\"]) show "Maps.cospan SPN_fgh.\.leg0 SPN_fgh.\.leg1" using SPN_fgh.\\.legs_form_cospan(1) by simp show "Maps.seq SPN_fgh.\\.prj\<^sub>1 tuple_BC" proof show "Maps.in_hom tuple_BC (Maps.MkIde (src TTfgh.p\<^sub>0)) (Maps.MkIde (src Tgh.p\<^sub>0))" using tuple_BC_in_hom by simp show "Maps.in_hom SPN_fgh.\\.prj\<^sub>1 (Maps.MkIde (src Tgh.p\<^sub>0)) SPN_fgh.\.apex" proof - have "Maps.pbdom SPN_fgh.\.leg0 SPN_fgh.\.leg1 = Maps.MkIde (src Tgh.p\<^sub>0)" using fg gh Maps.pbdom_def by (metis (no_types, lifting) SPN.preserves_ide SPN_fgh.\\.are_identities(2) SPN_fgh.\\.composable Span.chine_hcomp_ide_ide Tgh.chine_hcomp_SPN_SPN g.is_ide) thus ?thesis using SPN_fgh.\\.prj_in_hom(1) by simp qed qed show "Maps.seq SPN_fgh.\\.prj\<^sub>1 (Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\)" proof show "Maps.in_hom SPN_fgh.\\.prj\<^sub>1 (Maps.pbdom SPN_fgh.\.leg0 SPN_fgh.\.leg1) SPN_fgh.\.apex" using SPN_fgh.\\.prj_in_hom(1) by simp show "Maps.in_hom (Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\) \\src TTfgh_TfTgh.chine\\ (Maps.pbdom SPN_fgh.\.leg0 SPN_fgh.\.leg1)" proof show "Maps.in_hom \\TTfgh_TfTgh.chine\\ \\src TTfgh_TfTgh.chine\\ \\trg TTfgh_TfTgh.chine\\" using fg gh TTfgh_TfTgh.chine_in_hom Maps.CLS_in_hom TTfgh_TfTgh.is_map by blast show "Maps.in_hom (Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)) \\trg TTfgh_TfTgh.chine\\ (Maps.pbdom SPN_fgh.\.leg0 SPN_fgh.\.leg1)" proof show "Maps.cospan SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" using SPN_fgh.prj_in_hom(4) by blast show "\\trg TTfgh_TfTgh.chine\\ = Maps.pbdom SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" proof - have "\\trg TTfgh_TfTgh.chine\\ = Maps.MkIde (src TfTgh.p\<^sub>0)" by simp also have "... = Maps.pbdom SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" using 0 Maps.pbdom_def SPN_fgh.chine_composite(2) by presburger finally show ?thesis by blast qed show "Maps.pbdom SPN_fgh.\.leg0 SPN_fgh.\.leg1 = Maps.dom (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1)" using fg gh Maps.pbdom_def SPN_fgh.\\.apex_composite SPN_fgh.\\.dom.apex_def SPN_fgh.\\.dom.is_span SPN_fgh.\\.leg1_composite by presburger qed qed qed show "SPN_fgh.\\.prj\<^sub>0 \ tuple_BC = SPN_fgh.\\.prj\<^sub>0 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\" proof - have "SPN_fgh.\\.prj\<^sub>0 \ tuple_BC = SPN_fgh.Prj\<^sub>0" using csq(1) by simp also have "... = SPN_fgh.\\.prj\<^sub>0 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\" proof - have "SPN_fgh.\\.prj\<^sub>0 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\ = \\Tgh.p\<^sub>0\\ \ \\TfTgh.p\<^sub>0\\ \ \\TTfgh_TfTgh.chine\\" using fg gh Tgh.\\.prj_char TfTgh.prj_char(1) isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>1 g" "prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)" "tab\<^sub>1 g \ Tgh.p\<^sub>1"] by simp also have "... = \\Tgh.p\<^sub>0\\ \ \\TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\" using fg gh TTfgh_TfTgh.is_map isomorphic_reflexive Maps.comp_CLS [of TfTgh.p\<^sub>0 TTfgh_TfTgh.chine "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp also have "... = \\Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\" using fg gh TTfgh_TfTgh.is_map left_adjoints_compose isomorphic_reflexive Maps.comp_CLS [of Tgh.p\<^sub>0 "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" "Tgh.p\<^sub>0 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp also have "... = \\TTfgh.p\<^sub>0\\" using prj_chine(3) by simp also have "... = SPN_fgh.Prj\<^sub>0" using prj_char by simp finally show ?thesis by argo qed finally show ?thesis by blast qed show "SPN_fgh.\\.prj\<^sub>1 \ tuple_BC = SPN_fgh.\\.prj\<^sub>1 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\" proof - have "SPN_fgh.\\.prj\<^sub>1 \ tuple_BC = SPN_fgh.Prj\<^sub>0\<^sub>1" using csq(1) by simp also have "... = SPN_fgh.\\.prj\<^sub>1 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\" proof - have "SPN_fgh.\\.prj\<^sub>1 \ Maps.PRJ\<^sub>0 SPN_fgh.\.leg0 (SPN_fgh.\.leg1 \ SPN_fgh.\\.prj\<^sub>1) \ \\TTfgh_TfTgh.chine\\ = \\Tgh.p\<^sub>1\\ \ \\TfTgh.p\<^sub>0\\ \ \\TTfgh_TfTgh.chine\\" using fg gh Tgh.\\.prj_char TfTgh.prj_char(1) isomorphic_reflexive Maps.comp_CLS [of "tab\<^sub>1 g" "prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)" "tab\<^sub>1 g \ Tgh.p\<^sub>1"] by simp also have "... = \\Tgh.p\<^sub>1\\ \ \\TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\" using fg gh TTfgh_TfTgh.is_map isomorphic_reflexive Maps.comp_CLS [of TfTgh.p\<^sub>0 TTfgh_TfTgh.chine "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp also have "... = \\Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine\\" using fg gh TTfgh_TfTgh.is_map left_adjoints_compose isomorphic_reflexive Maps.comp_CLS [of Tgh.p\<^sub>1 "TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine" "Tgh.p\<^sub>1 \ TfTgh.p\<^sub>0 \ TTfgh_TfTgh.chine"] by simp also have "... = \\Tfg.p\<^sub>0 \ TTfgh.p\<^sub>1\\" using prj_chine(2) by simp also have "... = SPN_fgh.Prj\<^sub>0\<^sub>1" using prj_char by simp finally show ?thesis by argo qed finally show ?thesis by blast qed qed finally show ?thesis by simp qed qed thus ?thesis using SPN_fgh.chine_assoc_def by simp qed qed finally show ?thesis by simp qed text \ At long last, we can show associativity coherence for \SPN\. \ lemma assoc_coherence: shows "LHS = RHS" proof (intro Span.arr_eqI) show "Span.par LHS RHS" using par_LHS_RHS by blast show "Chn LHS = Chn RHS" proof - have "Chn LHS = \\HHfgh_HfHgh.chine\\ \ \\THfgh_HHfgh.chine\\ \ \\TTfgh_THfgh.chine\\" using Chn_LHS_eq by simp also have "... = \\HHfgh_HfHgh.chine \ THfgh_HHfgh.chine \ TTfgh_THfgh.chine\\" proof - have "\\THfgh_HHfgh.chine\\ \ \\TTfgh_THfgh.chine\\ = \\THfgh_HHfgh.chine \ TTfgh_THfgh.chine\\" using fg gh isomorphic_reflexive HHfgh_HfHgh.is_map THfgh_HHfgh.is_map TTfgh_THfgh.is_map left_adjoints_compose Maps.comp_CLS [of THfgh_HHfgh.chine TTfgh_THfgh.chine "THfgh_HHfgh.chine \ TTfgh_THfgh.chine"] by simp moreover have "\\HHfgh_HfHgh.chine\\ \ \\THfgh_HHfgh.chine \ TTfgh_THfgh.chine\\ = \\HHfgh_HfHgh.chine \ THfgh_HHfgh.chine \ TTfgh_THfgh.chine\\" proof - have "ide (HHfgh_HfHgh.chine \ THfgh_HHfgh.chine \ TTfgh_THfgh.chine)" proof - have "ide (THfgh_HHfgh.chine \ TTfgh_THfgh.chine)" using fg gh HHfgh_HfHgh.is_map THfgh_HHfgh.is_map TTfgh_THfgh.is_map left_adjoint_is_ide left_adjoints_compose by auto moreover have "src HHfgh_HfHgh.chine = trg (THfgh_HHfgh.chine \ TTfgh_THfgh.chine)" using fg gh HHfgh_HfHgh.chine_in_hom \_def by auto ultimately show ?thesis by simp qed thus ?thesis using fg gh isomorphic_reflexive HHfgh_HfHgh.is_map THfgh_HHfgh.is_map TTfgh_THfgh.is_map left_adjoints_compose Maps.comp_CLS [of HHfgh_HfHgh.chine "THfgh_HHfgh.chine \ TTfgh_THfgh.chine" "HHfgh_HfHgh.chine \ THfgh_HHfgh.chine \ TTfgh_THfgh.chine"] by auto qed ultimately show ?thesis by argo qed also have "... = \\TfHgh_HfHgh.chine \ TfTgh_TfHgh.chine \ TTfgh_TfTgh.chine\\" proof - interpret A: vertical_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg \(f \ g) \ h\ TTfgh.tab \tab\<^sub>0 h \ TTfgh.p\<^sub>0\ \(tab\<^sub>1 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\ \f \ g \ h\ TfTgh.tab \(tab\<^sub>0 h \ Tgh.p\<^sub>0) \ TfTgh.p\<^sub>0\ \tab\<^sub>1 f \ TfTgh.p\<^sub>1\ \f \ g \ h\ TfHgh.\\.tab \tab\<^sub>0 (g \ h) \ TfHgh.\\.p\<^sub>0\ \tab\<^sub>1 f \ TfHgh.\\.p\<^sub>1\ \\[f, g, h]\ \f \ g \ h\ .. interpret B: vertical_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg \(f \ g) \ h\ TTfgh.tab \tab\<^sub>0 h \ TTfgh.p\<^sub>0\ \(tab\<^sub>1 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\ \f \ g \ h\ TfHgh.\\.tab \tab\<^sub>0 (g \ h) \ TfHgh.\\.p\<^sub>0\ \tab\<^sub>1 f \ TfHgh.\\.p\<^sub>1\ \f \ g \ h\ HfHgh.tab \tab\<^sub>0 (f \ g \ h)\ \tab\<^sub>1 (f \ g \ h)\ \\[f, g, h]\ \f \ g \ h\ using fg gh by unfold_locales auto interpret C: vertical_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg \(f \ g) \ h\ THfgh.\\.tab \tab\<^sub>0 h \ THfgh.\\.p\<^sub>0\ \tab\<^sub>1 (f \ g) \ THfgh.\\.p\<^sub>1\ \(f \ g) \ h\ HHfgh.tab \tab\<^sub>0 ((f \ g) \ h)\ \tab\<^sub>1 ((f \ g) \ h)\ \f \ g \ h\ HfHgh.tab \tab\<^sub>0 (f \ g \ h)\ \tab\<^sub>1 (f \ g \ h)\ \(f \ g) \ h\ \\[f, g, h]\ using fg gh by unfold_locales auto interpret D: vertical_composite_of_arrows_of_tabulations_in_maps V H \ \ src trg \(f \ g) \ h\ TTfgh.tab \tab\<^sub>0 h \ TTfgh.p\<^sub>0\ \(tab\<^sub>1 f \ Tfg.p\<^sub>1) \ TTfgh.p\<^sub>1\ \(f \ g) \ h\ THfgh.\\.tab \tab\<^sub>0 h \ THfgh.\\.p\<^sub>0\ \tab\<^sub>1 (f \ g) \ THfgh.\\.p\<^sub>1\ \f \ g \ h\ HfHgh.tab \tab\<^sub>0 (f \ g \ h)\ \tab\<^sub>1 (f \ g \ h)\ \(f \ g) \ h\ \\[f, g, h]\ using fg gh by unfold_locales auto have "HHfgh_HfHgh.chine \ THfgh_HHfgh.chine \ TTfgh_THfgh.chine \ D.chine" proof - have "D.chine \ D.\.chine \ TTfgh_THfgh.chine" using D.chine_char by simp also have "... \ C.chine \ TTfgh_THfgh.chine" using fg gh comp_arr_dom isomorphic_reflexive by simp also have "... \ (C.\.chine \ THfgh_HHfgh.chine) \ TTfgh_THfgh.chine" using C.chine_char hcomp_isomorphic_ide by simp also have "... \ (HHfgh_HfHgh.chine \ THfgh_HHfgh.chine) \ TTfgh_THfgh.chine" proof - have "C.\.chine = HHfgh_HfHgh.chine" using fg gh comp_arr_dom comp_cod_arr \_def by simp hence "isomorphic C.\.chine HHfgh_HfHgh.chine" using isomorphic_reflexive by simp thus ?thesis using hcomp_isomorphic_ide by simp qed also have "... \ HHfgh_HfHgh.chine \ THfgh_HHfgh.chine \ TTfgh_THfgh.chine" proof - have "ide HHfgh_HfHgh.chine \ ide THfgh_HHfgh.chine \ ide TTfgh_THfgh.chine" by simp moreover have "src HHfgh_HfHgh.chine = trg THfgh_HHfgh.chine \ src THfgh_HHfgh.chine = trg TTfgh_THfgh.chine" using fg gh HHfgh_HfHgh.chine_in_hom THfgh_HHfgh.chine_in_hom TTfgh_THfgh.chine_in_hom \_def by auto ultimately show ?thesis using fg gh iso_assoc isomorphic_def assoc_in_hom [of HHfgh_HfHgh.chine THfgh_HHfgh.chine TTfgh_THfgh.chine] by auto qed finally show ?thesis using isomorphic_symmetric by blast qed also have "... \ B.chine" proof - have "D.chine = B.chine" using fg gh comp_arr_dom comp_cod_arr by simp thus ?thesis using isomorphic_reflexive by simp qed also have "... \ TfHgh_HfHgh.chine \ TfTgh_TfHgh.chine \ TTfgh_TfTgh.chine" proof - have "B.chine \ TfHgh_HfHgh.chine \ B.\.chine" using B.chine_char by simp also have "... \ TfHgh_HfHgh.chine \ A.chine" using fg gh comp_cod_arr isomorphic_reflexive by simp also have "... \ TfHgh_HfHgh.chine \ TfTgh_TfHgh.chine \ TTfgh_TfTgh.chine" using A.chine_char hcomp_ide_isomorphic by simp finally show ?thesis by blast qed finally have "HHfgh_HfHgh.chine \ THfgh_HHfgh.chine \ TTfgh_THfgh.chine \ TfHgh_HfHgh.chine \ TfTgh_TfHgh.chine \ TTfgh_TfTgh.chine" by blast thus ?thesis using fg gh Maps.CLS_eqI isomorphic_implies_hpar(1) by blast qed also have "... = \\TfHgh_HfHgh.chine\\ \ \\TfTgh_TfHgh.chine\\ \ \\TTfgh_TfTgh.chine\\" using fg gh isomorphic_reflexive TfTgh_TfHgh.is_map TfHgh_HfHgh.is_map TTfgh_TfTgh.is_map left_adjoints_compose Maps.comp_CLS [of TfTgh_TfHgh.chine TTfgh_TfTgh.chine "TfTgh_TfHgh.chine \ TTfgh_TfTgh.chine"] Maps.comp_CLS [of TfHgh_HfHgh.chine "TfTgh_TfHgh.chine \ TTfgh_TfTgh.chine" "TfHgh_HfHgh.chine \ TfTgh_TfHgh.chine \ TTfgh_TfTgh.chine"] by simp also have "... = Chn RHS" using Chn_RHS_eq CLS_chine tuple_ABC_eq_ABC'(2) by simp finally show ?thesis by blast qed qed end subsubsection "SPN is an Equivalence Pseudofunctor" context bicategory_of_spans begin interpretation Maps: maps_category V H \ \ src trg .. interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 .. no_notation Fun.comp (infixl "\" 55) notation Span.vcomp (infixr "\" 55) notation Span.hcomp (infixr "\" 53) notation Maps.comp (infixr "\" 55) notation isomorphic (infix "\" 50) interpretation SPN: "functor" V Span.vcomp SPN using SPN_is_functor by simp interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN using SPN_is_weak_arrow_of_homs by simp interpretation HoSPN_SPN: composite_functor VV.comp Span.VV.comp Span.vcomp SPN.FF \\\\. Span.hcomp (fst \\) (snd \\)\ .. interpretation SPNoH: composite_functor VV.comp V Span.vcomp \\\\. fst \\ \ snd \\\ SPN .. interpretation \: transformation_by_components VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map \\rs. CMP (fst rs) (snd rs)\ using compositor_is_natural_transformation by simp interpretation \: natural_isomorphism VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map \.map using compositor_is_natural_isomorphism by simp abbreviation \ where "\ \ \.map" interpretation SPN: pseudofunctor V H \ \ src trg Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \ proof show "\f g h. \ ide f; ide g; ide h; src f = trg g; src g = trg h \ \ SPN \[f, g, h] \ \ (f \ g, h) \ (\ (f, g) \ SPN h) = \ (f, g \ h) \ (SPN f \ \ (g, h)) \ Span.assoc (SPN f) (SPN g) (SPN h)" proof - fix f g h assume f: "ide f" and g: "ide g" and h: "ide h" assume fg: "src f = trg g" and gh: "src g = trg h" interpret fgh: three_composable_identities_in_bicategory_of_spans V H \ \ src trg f g h using f g h fg gh by (unfold_locales, simp) show "fgh.LHS = fgh.RHS" using fgh.assoc_coherence by simp qed qed lemma SPN_is_pseudofunctor: shows "pseudofunctor V H \ \ src trg Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \" .. interpretation SPN: equivalence_pseudofunctor V H \ \ src trg Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \ proof show "\\ \'. \par \ \'; SPN \ = SPN \'\ \ \ = \'" proof - fix \ \' assume par: "par \ \'" assume eq: "SPN \ = SPN \'" interpret dom_\: identity_in_bicategory_of_spans V H \ \ src trg \dom \\ using par apply unfold_locales by auto interpret cod_\: identity_in_bicategory_of_spans V H \ \ src trg \cod \\ using par apply unfold_locales by auto interpret \: arrow_of_tabulations_in_maps V H \ \ src trg \dom \\ \tab_of_ide (dom \)\ \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \tab_of_ide (cod \)\ \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (cod \)\ \ using par apply unfold_locales by auto interpret \: arrow_in_bicategory_of_spans V H \ \ src trg \dom \\ \cod \\ \ using par apply unfold_locales by auto interpret \': arrow_of_tabulations_in_maps V H \ \ src trg \dom \\ \tab_of_ide (dom \)\ \tab\<^sub>0 (dom \)\ \tab\<^sub>1 (dom \)\ \cod \\ \tab_of_ide (cod \)\ \tab\<^sub>0 (cod \)\ \tab\<^sub>1 (cod \)\ \' using par apply unfold_locales by auto interpret \': arrow_in_bicategory_of_spans V H \ \ src trg \dom \\ \cod \\ \' using par apply unfold_locales by auto have "\.chine \ \'.chine" using par eq SPN_def spn_def Maps.CLS_eqI \.is_ide by auto hence "\.\ = \'.\" using \.\_naturality \'.\_naturality by (metis \.\_simps(4) \'.\_simps(4) \.chine_is_induced_map \'.chine_is_induced_map \.induced_map_preserved_by_iso) thus "\ = \'" using par \.\_in_terms_of_\ \'.\_in_terms_of_\ by metis qed show "\a'. Span.obj a' \ \a. obj a \ Span.equivalent_objects (SPN.map\<^sub>0 a) a'" proof - fix a' assume a': "Span.obj a'" let ?a = "Maps.Dom (Chn a')" have a: "obj ?a" using a' Span.obj_char Span.ide_char Maps.ide_char by blast moreover have "Span.equivalent_objects (SPN.map\<^sub>0 ?a) a'" proof - have "SPN.map\<^sub>0 ?a = a'" proof (intro Span.arr_eqI) have "Chn (SPN.map\<^sub>0 ?a) = Chn (Span.src (SPN ?a))" using a a' by auto also have "... = Maps.MkIde (Maps.Dom (Chn a'))" proof - have "Maps.arr \\tab\<^sub>0 (dom (Maps.Dom (Chn a')))\\" proof - have "is_left_adjoint (tab\<^sub>0 (dom (Maps.Dom (Chn a'))))" using a by auto thus ?thesis using Maps.CLS_in_hom by auto qed moreover have "arr (Maps.Dom (Chn a'))" using a by auto moreover have "Span.arr (SPN (Maps.Dom (Chn a')))" using a a' SPN_in_hom by auto ultimately show ?thesis using a a' SPN_def Span.src_def Maps.cod_char obj_simps(2) by simp qed also have "... = Chn a'" using a' Maps.MkIde_Dom Span.obj_char Span.ide_char by simp finally show "Chn (SPN.map\<^sub>0 ?a) = Chn a'" by simp show "Span.par (SPN.map\<^sub>0 (Maps.Dom (Chn a'))) a'" using a a' Span.obj_char apply (intro conjI) using SPN.map\<^sub>0_simps(1) Span.obj_def apply blast apply simp apply (metis (no_types, lifting) SPN.map\<^sub>0_def SPN.preserves_arr Span.obj_src \Chn (SPN.map\<^sub>0 (Maps.Dom (Chn a'))) = Chn a'\ obj_def) by (metis (no_types, lifting) SPN.map\<^sub>0_def SPN.preserves_arr Span.obj_src \Chn (SPN.map\<^sub>0 (Maps.Dom (Chn a'))) = Chn a'\ obj_def) qed thus ?thesis using Span.equivalent_objects_reflexive by (simp add: a') qed ultimately show "\a. obj a \ Span.equivalent_objects (SPN.map\<^sub>0 a) a'" by auto qed show "\a b g. \obj a; obj b; Span.in_hhom g (SPN.map\<^sub>0 a) (SPN.map\<^sub>0 b); Span.ide g\ \ \f. \f : a \ b\ \ ide f \ Span.isomorphic (SPN f) g" proof - fix a b g assume a: "obj a" and b: "obj b" and g_in_hhom: "Span.in_hhom g (SPN.map\<^sub>0 a) (SPN.map\<^sub>0 b)" and ide_g: "Span.ide g" have arr_a: "arr a" using a by auto have arr_b: "arr b" using b by auto show "\f. \f : a \ b\ \ ide f \ Span.isomorphic (SPN f) g" proof - interpret g: arrow_of_spans Maps.comp g using ide_g Span.ide_char by blast interpret g: identity_arrow_of_spans Maps.comp g using ide_g Span.ide_char by unfold_locales auto interpret REP_leg0: map_in_bicategory V H \ \ src trg \Maps.REP g.leg0\ using Maps.REP_in_Map [of g.leg0] by unfold_locales auto have 0: "\Maps.REP g.leg0 : src (Maps.REP g.apex) \ Maps.Cod g.leg0\" using g.dom.leg_in_hom Maps.REP_in_hhom by (metis (no_types, lifting) Maps.Dom_cod Maps.REP_simps(2) Maps.arr_cod g.dom.leg_simps(1)) have 1: "\Maps.REP g.leg1 : src (Maps.REP g.apex) \ Maps.Cod g.leg1\" using g.dom.leg_in_hom Maps.REP_in_hhom by (metis (no_types, lifting) Maps.Dom_cod Maps.REP_simps(2) Maps.arr_cod g.dom.leg_simps(3)) let ?f = "Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*" have f_in_hhom: "\?f : a \ b\" proof show "\Maps.REP g.leg1 : src (Maps.REP g.apex) \ b\" proof - have "\Maps.REP g.leg1 : src (Maps.REP g.apex) \ Maps.Cod g.leg1\" using 1 by simp moreover have "Maps.Cod g.leg1 = b" proof - have "src (Maps.REP g.dtrg) = src (Maps.REP (Leg0 (Dom (SPN.map\<^sub>0 b))))" using g_in_hhom Span.trg_def [of g] by auto also have "... = src (Maps.REP (Maps.cod \\tab\<^sub>0 b\\))" using b arr_b SPN.map\<^sub>0_def Span.src_def SPN_in_hom by auto also have "... = src (Maps.REP \\trg (tab\<^sub>0 b)\\)" using b Maps.CLS_in_hom [of "tab\<^sub>0 b"] by force also have "... = src (Maps.REP \\b\\)" using b by fastforce also have "... = b" using b obj_simps by auto finally show ?thesis by simp qed ultimately show ?thesis by argo qed show "\(Maps.REP g.leg0)\<^sup>* : a \ src (Maps.REP g.apex)\" proof - have "\Maps.REP g.leg0 : src (Maps.REP g.apex) \ a\" proof - have "src (Maps.REP g.dsrc) = src (Maps.REP (Leg0 (Dom (SPN.map\<^sub>0 a))))" using g_in_hhom Span.src_def [of g] by auto also have "... = src (Maps.REP (Maps.cod \\tab\<^sub>0 a\\))" using a arr_a SPN.map\<^sub>0_def Span.src_def SPN_in_hom by auto also have "... = src (Maps.REP \\trg (tab\<^sub>0 a)\\)" using a Maps.CLS_in_hom [of "tab\<^sub>0 a"] by force also have "... = src (Maps.REP \\a\\)" using a by fastforce also have "... = a" using a obj_simps by auto finally show ?thesis by fast qed thus ?thesis using REP_leg0.antipar REP_leg0.ide_right apply (intro in_hhomI) by auto qed qed moreover have ide_f: "ide ?f" using REP_leg0.antipar f_in_hhom by fastforce moreover have "Span.isomorphic (SPN ?f) g" proof - have SPN_f_eq: "SPN ?f = \Chn = \\spn ?f\\, Dom = \Leg0 = \\tab\<^sub>0 ?f\\, Leg1 = \\tab\<^sub>1 ?f\\\, Cod = \Leg0 = \\tab\<^sub>0 ?f\\, Leg1 = \\tab\<^sub>1 ?f\\\\" using calculation(1) SPN_def [of ?f] REP_leg0.antipar by auto text \ We need an invertible arrow of spans from \SPN f\ to \g\. There exists a tabulation \(REP g.leg0, \, REP g.leg1)\ of \f\. There is also a tabulation \(tab\<^sub>0 f, \', tab\<^sub>1 f) of f\. As these are tabulations of the same arrow, they are equivalent. This yields an equivalence map which is an arrow of spans from \(tab\<^sub>0 f, tab\<^sub>1 f)\ to \(REP g.leg0, \, REP g.leg1)\. Its isomorphism class is an invertible arrow of spans in maps from \(CLS (tab\<^sub>0 f), CLS (tab\<^sub>1 f))\ to \(g.leg0, g.leg1)\. \ interpret f: identity_in_bicategory_of_spans V H \ \ src trg ?f using ide_f apply unfold_locales by auto interpret f: arrow_of_tabulations_in_maps V H \ \ src trg ?f f.tab \tab\<^sub>0 ?f\ \tab\<^sub>1 ?f\ ?f f.tab \tab\<^sub>0 ?f\ \tab\<^sub>1 ?f\ ?f using f.is_arrow_of_tabulations_in_maps by simp interpret g: span_of_maps V H \ \ src trg \Maps.REP g.leg0\ \Maps.REP g.leg1\ using Span.arr_char by (unfold_locales, blast+) have 2: "src (Maps.REP g.leg0) = src (Maps.REP g.leg1)" using 0 1 by fastforce hence "\\. tabulation (\) (\) \ \ src trg ?f \ (Maps.REP g.leg0) (Maps.REP g.leg1)" using BS2' [of "Maps.REP g.leg0" "Maps.REP g.leg1" ?f] isomorphic_reflexive Span.arr_char by auto hence "tabulation V H \ \ src trg ?f (REP_leg0.trnr\<^sub>\ (Maps.REP g.leg1) ?f) (Maps.REP g.leg0) (Maps.REP g.leg1)" using 2 REP_leg0.canonical_tabulation [of "Maps.REP g.leg1"] by auto hence 3: "\w w' \ \ \ \ \' \'. equivalence_in_bicategory (\) (\) \ \ src trg w' w \ \ \ \w : src (tab\<^sub>0 ?f) \ src (Maps.REP g.leg0)\ \ \w' : src (Maps.REP g.leg0) \ src (tab\<^sub>0 ?f)\ \ \\ : Maps.REP g.leg0 \ w \ tab\<^sub>0 ?f\ \ \\ : tab\<^sub>1 ?f \ Maps.REP g.leg1 \ w\ \ iso \ \ f.tab = (?f \ \) \ \[?f, Maps.REP g.leg0, w] \ (REP_leg0.trnr\<^sub>\ (Maps.REP g.leg1) ?f \ w) \ \ \ \\' : tab\<^sub>0 ?f \ w' \ Maps.REP g.leg0\ \ \\' : Maps.REP g.leg1 \ tab\<^sub>1 ?f \ w'\ \ iso \' \ REP_leg0.trnr\<^sub>\ (Maps.REP g.leg1) ?f = (?f \ \') \ \[?f, tab\<^sub>0 ?f, w'] \ (f.tab \ w') \ \'" using f.apex_unique_up_to_equivalence [of "REP_leg0.trnr\<^sub>\ (Maps.REP g.leg1) ?f" "Maps.REP g.leg0" "Maps.REP g.leg1"] by simp obtain w w' \ \ \ \ \' \' where 4: "equivalence_in_bicategory (\) (\) \ \ src trg w' w \ \ \ \w : src (tab\<^sub>0 ?f) \ src (Maps.REP g.leg0)\ \ \w' : src (Maps.REP g.leg0) \ src (tab\<^sub>0 ?f)\ \ \\ : Maps.REP g.leg0 \ w \ tab\<^sub>0 ?f\ \ \\ : tab\<^sub>1 ?f \ Maps.REP g.leg1 \ w\ \ iso \ \ f.tab = (?f \ \) \ \[?f, Maps.REP g.leg0, w] \ (REP_leg0.trnr\<^sub>\ (Maps.REP g.leg1) ?f \ w) \ \ \ \\' : tab\<^sub>0 ?f \ w' \ Maps.REP g.leg0\ \ \\' : Maps.REP g.leg1 \ tab\<^sub>1 ?f \ w'\ \ iso \' \ REP_leg0.trnr\<^sub>\ (Maps.REP g.leg1) ?f = (?f \ \') \ \[?f, tab\<^sub>0 ?f, w'] \ (f.tab \ w') \ \'" using 3 by meson hence w\\: "equivalence_map w \ \w : src (tab\<^sub>0 ?f) \ src (Maps.REP g.leg0)\ \ \\ : Maps.REP g.leg0 \ w \ tab\<^sub>0 ?f\ \ \\ : tab\<^sub>1 ?f \ Maps.REP g.leg1 \ w\ \ iso \" using equivalence_map_def quasi_inverses_def quasi_inverses_symmetric by meson let ?W = "\Chn = \\w\\, Dom = Dom (SPN ?f), Cod = Dom g\" have W_in_hom: "Span.in_hom ?W (SPN ?f) g" proof have "arrow_of_spans Maps.comp ?W" proof interpret Dom_W: span_in_category Maps.comp \Dom ?W\ proof (unfold_locales, intro conjI) show "Maps.arr (Leg0 (Dom ?W))" apply (intro Maps.arrI) apply auto by (metis f.base_simps(2) f.satisfies_T0 f.u_in_hom src_hcomp) show "Maps.arr (Leg1 (Dom ?W))" using 1 apply (intro Maps.arrI) apply auto proof - let ?f = "tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)" assume 1: "\Maps.REP g.leg1 : Maps.Dom g.apex \ Maps.Cod g.leg1\" have "\?f : src (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)) \ Maps.Cod g.leg1\ \ is_left_adjoint ?f \ \tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\ = \?f\" using 1 by simp thus "\f. \f : src (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)) \ Maps.Cod g.leg1\ \ is_left_adjoint f \ \tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\ = \f\" by blast qed show "Maps.dom (Leg0 (Dom ?W)) = Maps.dom (Leg1 (Dom ?W))" proof - have "Maps.dom (Leg0 (Dom ?W)) = Maps.MkIde (src (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)))" using Maps.dom_char apply simp by (metis (no_types, lifting) Maps.CLS_in_hom Maps.in_homE f.base_simps(2) f.satisfies_T0 f.u_simps(3) hcomp_simps(1)) also have "... = Maps.dom (Leg1 (Dom ?W))" using Maps.dom_char Maps.CLS_in_hom f.leg1_is_map f_in_hhom apply simp by (metis (no_types, lifting) Maps.in_homE Maps.REP_simps(3) f.base_simps(2) f.leg1_is_map f.leg1_simps(3) f.leg1_simps(4) g.dom.leg_simps(3) trg_hcomp) finally show ?thesis by blast qed qed show "Maps.span (Leg0 (Dom ?W)) (Leg1 (Dom ?W))" using Dom_W.span_in_category_axioms Dom_W.is_span by blast interpret Cod_W: span_in_category Maps.comp \Cod ?W\ apply unfold_locales by fastforce show "Maps.span (Leg0 (Cod ?W)) (Leg1 (Cod ?W))" by fastforce show "Maps.in_hom (Chn ?W) Dom_W.apex Cod_W.apex" proof show 1: "Maps.arr (Chn ?W)" using w\\ Maps.CLS_in_hom [of w] equivalence_is_adjoint by auto show "Maps.dom (Chn ?W) = Dom_W.apex" proof - have "Maps.dom (Chn ?W) = Maps.MkIde (src w)" using 1 w\\ Maps.dom_char by simp also have "... = Dom_W.apex" proof - have "src w = src (tab\<^sub>0 ?f)" using w\\ by blast thus ?thesis using Dom_W.apex_def Maps.arr_char Maps.dom_char apply simp by (metis (no_types, lifting) f.base_simps(2) f.satisfies_T0 f.u_in_hom hcomp_simps(1)) qed finally show ?thesis by fastforce qed show "Maps.cod (Chn ?W) = Cod_W.apex" proof - have "Maps.cod (Chn ?W) = Maps.MkIde (trg w)" using 1 w\\ Maps.cod_char by simp also have "... = Cod_W.apex" proof - have "trg w = src (Maps.REP g.leg0)" using w\\ by blast thus ?thesis using Cod_W.apex_def Maps.arr_char Maps.cod_char apply simp using g.dom.apex_def Maps.dom_char Maps.REP_simps(2) g.dom.is_span by presburger qed finally show ?thesis by fastforce qed qed show "Cod_W.leg0 \ Chn ?W = Dom_W.leg0" proof - have "Cod_W.leg0 \ Chn ?W = g.leg0 \ \\w\\" by simp also have "... = \\Maps.REP g.leg0\\ \ \\w\\" using g.dom.leg_simps(1) Maps.CLS_REP [of g.leg0] by simp also have "... = \\Maps.REP g.leg0 \ w\\" proof - have "is_left_adjoint (Maps.REP g.leg0)" by fast moreover have "is_left_adjoint w" using w\\ equivalence_is_adjoint by simp moreover have "Maps.REP g.leg0 \ w \ Maps.REP g.leg0 \ w" using w\\ isomorphic_reflexive Maps.REP_in_hhom by (metis (no_types, lifting) REP_leg0.ide_left adjoint_pair_antipar(1) calculation(2) ide_hcomp in_hhomE) ultimately show ?thesis using w\\ Maps.comp_CLS isomorphic_reflexive equivalence_is_adjoint by blast qed also have "... = \\tab\<^sub>0 ?f\\" proof - have "iso \" proof - have "is_left_adjoint (Maps.REP g.leg0 \ w)" using w\\ equivalence_is_adjoint Maps.REP_in_hhom by (simp add: g.leg0_is_map in_hhom_def left_adjoints_compose) moreover have "is_left_adjoint (tab\<^sub>0 ?f)" by simp ultimately show ?thesis using w\\ BS3 by blast qed thus ?thesis using w\\ Maps.CLS_eqI equivalence_is_adjoint by (meson isomorphic_def isomorphic_implies_hpar(1)) qed finally show ?thesis by fastforce qed show "Cod_W.leg1 \ Chn ?W = Dom_W.leg1" proof - have "Cod_W.leg1 \ Chn ?W = g.leg1 \ \\w\\" by simp also have "... = \\Maps.REP g.leg1\\ \ \\w\\" using g.dom.leg_simps(3) Maps.CLS_REP by presburger also have "... = \\Maps.REP g.leg1 \ w\\" proof - have "is_left_adjoint (Maps.REP g.leg1)" by fast moreover have "is_left_adjoint w" using w\\ equivalence_is_adjoint by simp moreover have "Maps.REP g.leg1 \ w \ Maps.REP g.leg1 \ w" using w\\ isomorphic_reflexive Maps.REP_in_hhom by (metis (no_types, lifting) "2" calculation(2) g.dom.is_span hcomp_ide_isomorphic Maps.ide_REP in_hhomE right_adjoint_determines_left_up_to_iso) ultimately show ?thesis using w\\ Maps.comp_CLS isomorphic_reflexive equivalence_is_adjoint by blast qed also have "... = \\tab\<^sub>1 ?f\\" proof - have "ide (Maps.REP g.leg1 \ w)" using 2 w\\ equivalence_map_is_ide by auto moreover have "Maps.REP g.leg1 \ w \ tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)" using w\\ equivalence_is_adjoint f.leg1_is_map right_adjoint_determines_left_up_to_iso adjoint_pair_preserved_by_iso by (meson adjoint_pair_antipar(2) ide_in_hom(2) ide_is_iso) ultimately show ?thesis using Maps.CLS_eqI by blast qed finally show ?thesis by fastforce qed qed thus W: "Span.arr ?W" using Span.arr_char by blast interpret Dom_W: span_in_category Maps.comp \\Leg0 = Maps.MkArr (src (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*))) (src (Maps.REP g.leg0)\<^sup>*) (iso_class (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*))), Leg1 = Maps.MkArr (src (tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*))) (Maps.Cod g.leg1) (iso_class (tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)))\\ using W Span.arr_char by (simp add: arrow_of_spans_def) interpret Cod_W: span_in_category Maps.comp \Cod ?W\ using W Span.arr_char by (simp add: arrow_of_spans_def) show "Span.dom ?W = SPN ?f" proof - have "Span.dom ?W = \Chn = Dom_W.apex, Dom = \Leg0 = \\tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\\, Leg1 = \\tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\\\, Cod = \Leg0 = \\tab\<^sub>0 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\\, Leg1 = \\tab\<^sub>1 (Maps.REP g.leg1 \ (Maps.REP g.leg0)\<^sup>*)\\\\" using 0 W Span.dom_char by simp also have "... = SPN ?f" using SPN_def Dom_W.apex_def Maps.dom_char Dom_W.is_span iso_class_eqI spn_ide apply simp using ide_f by blast finally show ?thesis by blast qed show "Span.cod ?W = g" using 0 W Span.cod_char Cod_W.apex_def by simp qed moreover have "Span.iso ?W" proof - have "Maps.iso \\w\\" proof - have "Maps.arr \\w\\ \ w \ Maps.Map \\w\\ \ equivalence_map w" proof (intro conjI) show 1: "Maps.arr \\w\\" using w\\ Maps.CLS_in_hom equivalence_is_adjoint by blast show "equivalence_map w" using w\\ by blast show "w \ Maps.Map \\w\\" using 1 w\\ equivalence_is_adjoint Maps.arr_char by (simp add: equivalence_map_is_ide ide_in_iso_class) qed thus ?thesis using Maps.iso_char' by blast qed thus ?thesis using w\\ W_in_hom Span.iso_char by auto qed ultimately show ?thesis using Span.isomorphic_def by blast qed ultimately show ?thesis by blast qed qed show "\r s \. \ide r; ide s; src r = src s; trg r = trg s; Span.in_hom \ (SPN r) (SPN s)\ \ \\. \\ : r \ s\ \ SPN \ = \" proof - fix r s \ assume r: "ide r" and s: "ide s" assume src_eq: "src r = src s" and trg_eq: "trg r = trg s" assume \: "Span.in_hom \ (SPN r) (SPN s)" interpret \: arrow_of_spans Maps.comp \ using \ Span.arr_char by auto interpret r: identity_in_bicategory_of_spans V H \ \ src trg r using r by unfold_locales auto interpret s: identity_in_bicategory_of_spans V H \ \ src trg s using s by unfold_locales auto interpret s: arrow_of_tabulations_in_maps V H \ \ src trg s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 s\ s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 s\ s using s.is_arrow_of_tabulations_in_maps by simp have \_dom_leg0_eq: "\.dom.leg0 = \\tab\<^sub>0 r\\" using \ Span.dom_char SPN_def [of r] by auto have \_dom_leg1_eq: "\.dom.leg1 = \\tab\<^sub>1 r\\" using \ Span.dom_char SPN_def [of r] by auto have \_cod_leg0_eq: "\.cod.leg0 = \\tab\<^sub>0 s\\" using \ Span.cod_char SPN_def [of s] by auto have \_cod_leg1_eq: "\.cod.leg1 = \\tab\<^sub>1 s\\" using \ Span.cod_char SPN_def [of s] by auto have 1: "tab\<^sub>0 s \ Maps.REP \.chine \ tab\<^sub>0 r" proof - have "tab\<^sub>0 s \ Maps.REP \.chine \ Maps.REP \.cod.leg0 \ Maps.REP \.chine" proof - have "Maps.REP \.cod.leg0 \ tab\<^sub>0 s" using \_cod_leg0_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS s.satisfies_T0 by presburger moreover have "src (tab\<^sub>0 s) = trg (Maps.REP \.chine)" using Maps.seq_char [of "\\tab\<^sub>0 s\\" \.chine] \.cod.leg_simps(1) \.leg0_commutes \_cod_leg0_eq by auto ultimately show ?thesis using hcomp_isomorphic_ide [of "tab\<^sub>0 s" "Maps.REP \.cod.leg0" "Maps.REP \.chine"] isomorphic_symmetric by fastforce qed also have "... \ Maps.REP \.dom.leg0" proof - have "\\Maps.REP \.cod.leg0\\ \ \\Maps.REP \.chine\\ = \\Maps.REP \.dom.leg0\\" using \.leg0_commutes Maps.CLS_REP \.chine_simps(1) \.cod.leg_simps(1) \.dom.leg_simps(1) by presburger hence "\\Maps.REP \.cod.leg0 \ Maps.REP \.chine\\ = \\Maps.REP \.dom.leg0\\" using Maps.comp_CLS [of "Maps.REP \.cod.leg0" "Maps.REP \.chine" "Maps.REP \.cod.leg0 \ Maps.REP \.chine"] isomorphic_reflexive by (metis (no_types, lifting) Maps.seq_char Maps.REP_in_hhom(2) Maps.REP_simps(2-3) \.chine_in_hom \.cod.leg_in_hom(1) \.dom.leg_simps(1) \.leg0_commutes ide_hcomp Maps.ide_REP) thus ?thesis using Maps.CLS_eqI Maps.seq_char Maps.ide_REP by (meson calculation isomorphic_implies_ide(2)) qed also have "... \ tab\<^sub>0 r" using \_dom_leg0_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS r.satisfies_T0 by presburger finally show ?thesis by blast qed obtain \ where \: "\\ : tab\<^sub>0 s \ Maps.REP \.chine \ tab\<^sub>0 r\ \ iso \" using 1 by blast have 2: "tab\<^sub>1 s \ Maps.REP \.chine \ tab\<^sub>1 r" proof - have "tab\<^sub>1 s \ Maps.REP \.chine \ Maps.REP \.cod.leg1 \ Maps.REP \.chine" proof - have "Maps.REP \.cod.leg1 \ tab\<^sub>1 s" using \_cod_leg1_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS s.leg1_is_map by presburger moreover have "src (Maps.REP \.cod.leg1) = trg (Maps.REP \.chine)" using Maps.seq_char by auto ultimately show ?thesis using hcomp_isomorphic_ide [of "Maps.REP \.cod.leg1" "tab\<^sub>1 s" "Maps.REP \.chine"] isomorphic_symmetric by fastforce qed also have "... \ Maps.REP \.dom.leg1" proof - have "\\Maps.REP \.cod.leg1\\ \ \\Maps.REP \.chine\\ = \\Maps.REP \.dom.leg1\\" using \.leg1_commutes Maps.CLS_REP \.chine_simps(1) \.cod.leg_simps(3) \.dom.leg_simps(3) by presburger hence "\\Maps.REP \.cod.leg1 \ Maps.REP \.chine\\ = \\Maps.REP \.dom.leg1\\" using Maps.comp_CLS [of "Maps.REP \.cod.leg1" "Maps.REP \.chine" "Maps.REP \.cod.leg1 \ Maps.REP \.chine"] isomorphic_reflexive by (metis (no_types, lifting) Maps.seq_char Maps.REP_in_hhom(2) Maps.REP_simps(2) Maps.REP_simps(3) \.chine_in_hom \.cod.leg_in_hom(2) \.dom.leg_simps(3) \.leg1_commutes ide_hcomp Maps.ide_REP) thus ?thesis using Maps.CLS_eqI Maps.seq_char Maps.ide_REP by (meson calculation isomorphic_implies_ide(2)) qed also have "... \ tab\<^sub>1 r" using \_dom_leg1_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS r.leg1_is_map by presburger finally show ?thesis by blast qed obtain \ where \: "\\ : tab\<^sub>1 r \ tab\<^sub>1 s \ Maps.REP \.chine\ \ iso \" using 2 isomorphic_symmetric by blast define \ where "\ \ (s \ \) \ \[s, tab\<^sub>0 s, Maps.REP \.chine] \ (s.tab \ Maps.REP \.chine) \ \" have \: "\\ : tab\<^sub>1 r \ s \ tab\<^sub>0 r\" proof (unfold \_def, intro comp_in_homI) show "\\ : tab\<^sub>1 r \ tab\<^sub>1 s \ Maps.REP \.chine\" using \ by simp show 3: "\s.tab \ Maps.REP \.chine : tab\<^sub>1 s \ Maps.REP \.chine \ (s \ tab\<^sub>0 s) \ Maps.REP \.chine\" apply (intro hcomp_in_vhom) apply auto using "1" by fastforce show "\\[s, tab\<^sub>0 s, Maps.REP \.chine] : (s \ tab\<^sub>0 s) \ Maps.REP \.chine \ s \ tab\<^sub>0 s \ Maps.REP \.chine\" using s assoc_in_hom [of s "tab\<^sub>0 s" "Maps.REP \.chine"] by (metis (no_types, lifting) Maps.ide_REP 3 \.chine_simps(1) hcomp_in_vhomE ideD(2) ideD(3) s.ide_u s.tab_simps(2) s.u_simps(3)) show "\s \ \ : s \ tab\<^sub>0 s \ Maps.REP \.chine \ s \ tab\<^sub>0 r\" using 1 s \ isomorphic_implies_hpar(4) src_eq by auto qed define \ where "\ \ r.T0.trnr\<^sub>\ s \ \ inv (r.T0.trnr\<^sub>\ r r.tab)" have \: "\\ : r \ s\" proof (unfold \_def, intro comp_in_homI) show "\inv (r.T0.trnr\<^sub>\ r r.tab) : r \ tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*\" using r.yields_isomorphic_representation by fastforce show "\r.T0.trnr\<^sub>\ s \ : tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>* \ s\" using s \ src_eq r.T0.adjoint_transpose_right(2) [of s "tab\<^sub>1 r"] by auto qed interpret \: arrow_in_bicategory_of_spans V H \ \ src trg r s \ using \ by unfold_locales auto interpret \: arrow_of_tabulations_in_maps V H \ \ src trg r r.tab \tab\<^sub>0 r\ \tab\<^sub>1 r\ s s.tab \tab\<^sub>0 s\ \tab\<^sub>1 s\ \ using \.is_arrow_of_tabulations_in_maps by simp have \_eq: "\ = \.\" proof - have "r.T0.trnr\<^sub>\ s \ \ inv (r.T0.trnr\<^sub>\ r r.tab) = r.T0.trnr\<^sub>\ s \.\ \ inv (r.T0.trnr\<^sub>\ r r.tab)" using \ \_def \.\_in_terms_of_\ by auto hence "r.T0.trnr\<^sub>\ s \ = r.T0.trnr\<^sub>\ s \.\" using r s \ r.T0.adjoint_transpose_right(2) r.yields_isomorphic_representation iso_inv_iso iso_is_retraction retraction_is_epi epiE by (metis \.in_hom \_def arrI) thus ?thesis using \ \.\_in_hom(2) src_eq r.T0.adjoint_transpose_right(6) bij_betw_imp_inj_on [of "r.T0.trnr\<^sub>\ s" "hom (tab\<^sub>1 r) (s \ tab\<^sub>0 r)" "hom (tab\<^sub>1 r \ (tab\<^sub>0 r)\<^sup>*) s"] inj_on_def [of "r.T0.trnr\<^sub>\ s" "hom (tab\<^sub>1 r) (s \ tab\<^sub>0 r)"] by simp qed have "\.is_induced_map (Maps.REP \.chine)" using \ \ \_eq \_def \.is_induced_map_iff \.chine_simps(1) Maps.ide_REP by blast hence 3: "Maps.REP \.chine \ \.chine" using \.chine_is_induced_map \.induced_map_unique by simp have "SPN \ = \" proof (intro Span.arr_eqI) show "Span.par (SPN \) \" using \ \ SPN_in_hom by (metis (no_types, lifting) SPN.preserves_cod SPN.preserves_dom Span.in_homE in_homE) show "Chn (SPN \) = \.chine" proof - have "Chn (SPN \) = \\spn \\\" using \ SPN_def spn_def by auto also have "... = \\\.chine\\" using \ spn_def by fastforce also have "... = \\Maps.REP \.chine\\" using 3 isomorphic_symmetric Maps.CLS_eqI iso_class_eqI isomorphic_implies_hpar(3) isomorphic_implies_hpar(4) by auto also have "... = \.chine" using Maps.CLS_REP \.chine_simps(1) by blast finally show ?thesis by blast qed qed thus "\\. \\ : r \ s\ \ SPN \ = \" using \ by auto qed qed theorem SPN_is_equivalence_pseudofunctor: shows "equivalence_pseudofunctor V H \ \ src trg Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \" .. text \ We have completed the proof of the second half of the main result (CKS Theorem 4): \B\ is biequivalent (via \SPN\) to \Span(Maps(B))\. \ corollary shows "equivalent_bicategories V H \ \ src trg Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg" using SPN_is_equivalence_pseudofunctor equivalent_bicategories_def by blast end end diff --git a/thys/Bicategory/Coherence.thy b/thys/Bicategory/Coherence.thy --- a/thys/Bicategory/Coherence.thy +++ b/thys/Bicategory/Coherence.thy @@ -1,3957 +1,3957 @@ (* Title: Coherence Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) section "Coherence" theory Coherence imports Bicategory begin text \ \sloppypar In this section, we generalize to bicategories the proof of the Coherence Theorem that we previously gave for monoidal categories (see \MonoidalCategory.evaluation_map.coherence\ in @{session MonoidalCategory}). As was the case for the previous proof, the current proof takes a syntactic approach. First we define a formal ``bicategorical language'' of terms constructed using syntactic operators that correspond to the various operations (vertical and horizontal composition, associators and unitors) found in a bicategory. Terms of the language are classified as formal ``arrows'', ``identities'', or ``objects'' according to the syntactic operators used in their formation. A class of terms called ``canonical'' is also defined in this way. Functions that map ``arrows'' to their ``domain'' and ``codomain'', and to their ``source'' and ``target'' are defined by recursion on the structure of terms. Next, we define a notion of ``normal form'' for terms in this language and we give a recursive definition of a function that maps terms to their normal forms. Normalization moves vertical composition inside of horizontal composition and ``flattens'' horizontal composition by associating all horizontal compositions to the right. In addition, normalization deletes from a term any horizontal composites involving an arrow and its source or target, replacing such composites by just the arrow itself. We then define a ``reduction function'' that maps each identity term \t\ to a ``canonical'' term \t\<^bold>\\ that connects \t\ with its normal form. The definition of reduction is also recursive, but it is somewhat more complex than normalization in that it involves two mutually recursive functions: one that applies to any identity term and another that applies only to terms that are the horizontal composite of two identity terms. The next step is to define an ``evaluation function'' that evaluates terms in a given bicategory (which is left as an unspecified parameter). We show that evaluation respects bicategorical structure: the domain, codomain, source, and target mappings on terms correspond under evaluation to the actual domain, codomain, source and target mappings on the given bicategory, the vertical and horizontal composition on terms correspond to the actual vertical and horizontal composition of the bicategory, and unit and associativity terms evaluate to the actual unit and associativity isomorphisms of the bicategory. In addition, ``object terms'' evaluate to objects (\emph{i.e.}~0-cells), ``identity terms'' evaluate to identities (\emph{i.e.}~1-cells), ``arrow terms'' evaluate to arrows (\emph{i.e.}~2-cells), and ``canonical terms'' evaluate to canonical isomorphisms. A term is defined to be ``coherent'' if, roughly speaking, it is a formal arrow whose evaluation commutes with the evaluations of the reductions to normal form of its domain and codomain. We then prove the Coherence Theorem, expressed in the form: ``every arrow is coherent.'' This implies a more classical version of the Coherence Theorem, which says that: ``syntactically parallel arrows with the same normal form have equal evaluations''. \ subsection "Bicategorical Language" text \ For the most part, the definition of the ``bicategorical language'' of terms is a straightforward generalization of the ``monoidal language'' that we used for monoidal categories. Some modifications are required, however, due to the fact that horizontal composition in a bicategory is a partial operation, whereas the the tensor product in a monoidal category is well-defined for all pairs of arrows. One difference is that we have found it necessary to introduce a new class of primitive terms whose elements represent ``formal objects'', so that there is some way to identify the source and target of what would otherwise be an empty horizontal composite. This was not an issue for monoidal categories, because the totality of horizontal composition meant that there was no need for syntactically defined sources and targets. Another difference is what we have chosen for the ``generators'' of the language and how they are used to form primitive terms. For monoidal categories, we supposed that we were given a category \C\ and the syntax contained a constructor to form a primitive term corresponding to each arrow of \C\. We assumed a category as the given data, rather than something less structured, such as a graph, because we were primarily interested in the tensor product and the associators and unitors, and were relatively uninterested in the strictly associative and unital composition of the underlying category. For bicategories, we also take the vertical composition as given for the same reasons; however, this is not yet sufficient due to the fact that horizontal composition in a bicategory is a partial operation, in contrast to the tensor product in a monoidal category, which is defined for all pairs of arrows. To deal with this issue, for bicategories we assume that source and target mappings are also given, so that the given data forms a category with ``horizontal homs''. The given source and target mappings are extended to all terms and used to define when two terms are ``formally horizontally composable''. \ locale bicategorical_language = category V + horizontal_homs V src trg for V :: "'a comp" (infixr "\" 55) and src :: "'a \ 'a" and trg :: "'a \ 'a" begin text \ Constructor \Prim\<^sub>0\ is used to construct ``formal objects'' and \Prim\ is used to construct primitive terms that are not formal objects. \ datatype (discs_sels) 't "term" = Prim\<^sub>0 't ("\<^bold>\_\<^bold>\\<^sub>0") | Prim 't ("\<^bold>\_\<^bold>\") | Hcomp "'t term" "'t term" (infixr "\<^bold>\" 53) | Vcomp "'t term" "'t term" (infixr "\<^bold>\" 55) | Lunit "'t term" ("\<^bold>\\<^bold>[_\<^bold>]") | Lunit' "'t term" ("\<^bold>\\<^sup>-\<^sup>1\<^bold>[_\<^bold>]") | Runit "'t term" ("\<^bold>\\<^bold>[_\<^bold>]") | Runit' "'t term" ("\<^bold>\\<^sup>-\<^sup>1\<^bold>[_\<^bold>]") | Assoc "'t term" "'t term" "'t term" ("\<^bold>\\<^bold>[_, _, _\<^bold>]") | Assoc' "'t term" "'t term" "'t term" ("\<^bold>\\<^sup>-\<^sup>1\<^bold>[_, _, _\<^bold>]") text \ We define formal domain, codomain, source, and target functions on terms. \ primrec Src :: "'a term \ 'a term" where "Src \<^bold>\\\<^bold>\\<^sub>0 = \<^bold>\\\<^bold>\\<^sub>0" | "Src \<^bold>\\\<^bold>\ = \<^bold>\src \\<^bold>\\<^sub>0" | "Src (t \<^bold>\ u) = Src u" | "Src (t \<^bold>\ u) = Src t" | "Src \<^bold>\\<^bold>[t\<^bold>] = Src t" | "Src \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Src t" | "Src \<^bold>\\<^bold>[t\<^bold>] = Src t" | "Src \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Src t" | "Src \<^bold>\\<^bold>[t, u, v\<^bold>] = Src v" | "Src \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = Src v" primrec Trg :: "'a term \ 'a term" where "Trg \<^bold>\\\<^bold>\\<^sub>0 = \<^bold>\\\<^bold>\\<^sub>0" | "Trg \<^bold>\\\<^bold>\ = \<^bold>\trg \\<^bold>\\<^sub>0" | "Trg (t \<^bold>\ u) = Trg t" | "Trg (t \<^bold>\ u) = Trg t" | "Trg \<^bold>\\<^bold>[t\<^bold>] = Trg t" | "Trg \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Trg t" | "Trg \<^bold>\\<^bold>[t\<^bold>] = Trg t" | "Trg \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Trg t" | "Trg \<^bold>\\<^bold>[t, u, v\<^bold>] = Trg t" | "Trg \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = Trg t" primrec Dom :: "'a term \ 'a term" where "Dom \<^bold>\\\<^bold>\\<^sub>0 = \<^bold>\\\<^bold>\\<^sub>0" | "Dom \<^bold>\\\<^bold>\ = \<^bold>\dom \\<^bold>\" | "Dom (t \<^bold>\ u) = Dom t \<^bold>\ Dom u" | "Dom (t \<^bold>\ u) = Dom u" | "Dom \<^bold>\\<^bold>[t\<^bold>] = Trg t \<^bold>\ Dom t" | "Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Dom t" | "Dom \<^bold>\\<^bold>[t\<^bold>] = Dom t \<^bold>\ Src t" | "Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Dom t" | "Dom \<^bold>\\<^bold>[t, u, v\<^bold>] = (Dom t \<^bold>\ Dom u) \<^bold>\ Dom v" | "Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = Dom t \<^bold>\ (Dom u \<^bold>\ Dom v)" primrec Cod :: "'a term \ 'a term" where "Cod \<^bold>\\\<^bold>\\<^sub>0 = \<^bold>\\\<^bold>\\<^sub>0" | "Cod \<^bold>\\\<^bold>\ = \<^bold>\cod \\<^bold>\" | "Cod (t \<^bold>\ u) = Cod t \<^bold>\ Cod u" | "Cod (t \<^bold>\ u) = Cod t" | "Cod \<^bold>\\<^bold>[t\<^bold>] = Cod t" | "Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Trg t \<^bold>\ Cod t" | "Cod \<^bold>\\<^bold>[t\<^bold>] = Cod t" | "Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Cod t \<^bold>\ Src t" | "Cod \<^bold>\\<^bold>[t, u, v\<^bold>] = Cod t \<^bold>\ (Cod u \<^bold>\ Cod v)" | "Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = (Cod t \<^bold>\ Cod u) \<^bold>\ Cod v" text \ A term is a ``formal arrow'' if it is constructed from primitive arrows in such a way that horizontal and vertical composition are applied only to formally composable pairs of terms. The definitions of ``formal identity'' and ``formal object'' follow a similar pattern. \ primrec Arr :: "'a term \ bool" where "Arr \<^bold>\\\<^bold>\\<^sub>0 = obj \" | "Arr \<^bold>\\\<^bold>\ = arr \" | "Arr (t \<^bold>\ u) = (Arr t \ Arr u \ Src t = Trg u)" | "Arr (t \<^bold>\ u) = (Arr t \ Arr u \ Dom t = Cod u)" | "Arr \<^bold>\\<^bold>[t\<^bold>] = Arr t" | "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Arr t" | "Arr \<^bold>\\<^bold>[t\<^bold>] = Arr t" | "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Arr t" | "Arr \<^bold>\\<^bold>[t, u, v\<^bold>] = (Arr t \ Arr u \ Arr v \ Src t = Trg u \ Src u = Trg v)" | "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = (Arr t \ Arr u \ Arr v \ Src t = Trg u \ Src u = Trg v)" primrec Ide :: "'a term \ bool" where "Ide \<^bold>\\\<^bold>\\<^sub>0 = obj \" | "Ide \<^bold>\\\<^bold>\ = ide \" | "Ide (t \<^bold>\ u) = (Ide t \ Ide u \ Src t = Trg u)" | "Ide (t \<^bold>\ u) = False" | "Ide \<^bold>\\<^bold>[t\<^bold>] = False" | "Ide \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = False" | "Ide \<^bold>\\<^bold>[t\<^bold>] = False" | "Ide \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = False" | "Ide \<^bold>\\<^bold>[t, u, v\<^bold>] = False" | "Ide \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = False" primrec Obj :: "'a term \ bool" where "Obj \<^bold>\\\<^bold>\\<^sub>0 = obj \" | "Obj \<^bold>\\\<^bold>\ = False" | "Obj (t \<^bold>\ u) = False" | "Obj (t \<^bold>\ u) = False" | "Obj \<^bold>\\<^bold>[t\<^bold>] = False" | "Obj \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = False" | "Obj \<^bold>\\<^bold>[t\<^bold>] = False" | "Obj \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = False" | "Obj \<^bold>\\<^bold>[t, u, v\<^bold>] = False" | "Obj \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = False" abbreviation HSeq :: "'a term \ 'a term \ bool" where "HSeq t u \ Arr t \ Arr u \ Src t = Trg u" abbreviation VSeq :: "'a term \ 'a term \ bool" where "VSeq t u \ Arr t \ Arr u \ Dom t = Cod u" abbreviation HPar :: "'a term => 'a term \ bool" where "HPar t u \ Arr t \ Arr u \ Src t = Src u \ Trg t = Trg u" abbreviation VPar :: "'a term => 'a term \ bool" where "VPar t u \ Arr t \ Arr u \ Dom t = Dom u \ Cod t = Cod u" abbreviation HHom :: "'a term \ 'a term \ 'a term set" where "HHom a b \ { t. Arr t \ Src t = a \ Trg t = b }" abbreviation VHom :: "'a term \ 'a term \ 'a term set" where "VHom f g \ { t. Arr t \ Dom t = f \ Cod t = g }" lemma is_Prim0_Src: shows "is_Prim\<^sub>0 (Src t)" by (induct t; simp) lemma is_Prim0_Trg: shows "is_Prim\<^sub>0 (Trg t)" by (induct t; simp) lemma Src_Src [simp]: shows "Arr t \ Src (Src t) = Src t" by (induct t) auto lemma Trg_Trg [simp]: shows "Arr t \ Trg (Trg t) = Trg t" by (induct t) auto lemma Src_Trg [simp]: shows "Arr t \ Src (Trg t) = Trg t" by (induct t) auto lemma Trg_Src [simp]: shows "Arr t \ Trg (Src t) = Src t" by (induct t) auto lemma Dom_Src [simp]: shows "Arr t \ Dom (Src t) = Src t" by (induct t) auto lemma Dom_Trg [simp]: shows "Arr t \ Dom (Trg t) = Trg t" by (induct t) auto lemma Cod_Src [simp]: shows "Arr t \ Cod (Src t) = Src t" by (induct t) auto lemma Cod_Trg [simp]: shows "Arr t \ Cod (Trg t) = Trg t" by (induct t) auto lemma Src_Dom_Cod: shows "Arr t \ Src (Dom t) = Src t \ Src (Cod t) = Src t" using src_dom src_cod by (induct t) auto lemma Src_Dom [simp]: shows "Arr t \ Src (Dom t) = Src t" using Src_Dom_Cod by blast lemma Src_Cod [simp]: shows "Arr t \ Src (Cod t) = Src t" using Src_Dom_Cod by blast lemma Trg_Dom_Cod: shows "Arr t \ Trg (Dom t) = Trg t \ Trg (Cod t) = Trg t" using trg_dom trg_cod by (induct t) auto lemma Trg_Dom [simp]: shows "Arr t \ Trg (Dom t) = Trg t" using Trg_Dom_Cod by blast lemma Trg_Cod [simp]: shows "Arr t \ Trg (Cod t) = Trg t" using Trg_Dom_Cod by blast lemma VSeq_implies_HPar: shows "VSeq t u \ HPar t u" using Src_Dom [of t] Src_Cod [of u] Trg_Dom [of t] Trg_Cod [of u] by auto lemma Dom_Dom [simp]: shows "Arr t \ Dom (Dom t) = Dom t" by (induct t, auto) lemma Cod_Cod [simp]: shows "Arr t \ Cod (Cod t) = Cod t" by (induct t, auto) lemma Dom_Cod [simp]: shows "Arr t \ Dom (Cod t) = Cod t" by (induct t, auto) lemma Cod_Dom [simp]: shows "Arr t \ Cod (Dom t) = Dom t" by (induct t, auto) lemma Obj_implies_Ide (*[simp]*): shows "Obj t \ Ide t" by (induct t) auto lemma Ide_implies_Arr [simp]: shows "Ide t \ Arr t" by (induct t, auto) lemma Dom_Ide: shows "Ide t \ Dom t = t" by (induct t, auto) lemma Cod_Ide: shows "Ide t \ Cod t = t" by (induct t, auto) lemma Obj_Src [simp]: shows "Arr t \ Obj (Src t)" by (induct t) auto lemma Obj_Trg [simp]: shows "Arr t \ Obj (Trg t)" by (induct t) auto lemma Ide_Dom [simp]: shows "Arr t \ Ide (Dom t)" using Obj_implies_Ide by (induct t, auto) lemma Ide_Cod [simp]: shows "Arr t \ Ide (Cod t)" using Obj_implies_Ide by (induct t, auto) lemma Arr_in_Hom: assumes "Arr t" shows "t \ HHom (Src t) (Trg t)" and "t \ VHom (Dom t) (Cod t)" proof - have 1: "Arr t \ t \ HHom (Src t) (Trg t) \ t \ VHom (Dom t) (Cod t)" by (induct t, auto) show "t \ HHom (Src t) (Trg t)" using assms 1 by simp show "t \ VHom (Dom t) (Cod t)" using assms 1 by simp qed lemma Ide_in_Hom: assumes "Ide t" shows "t \ HHom (Src t) (Trg t)" and "t \ VHom t t" proof - have 1: "Ide t \ t \ HHom (Src t) (Trg t) \ t \ VHom t t" by (induct t, auto) show "t \ HHom (Src t) (Trg t)" using assms 1 by simp show "t \ VHom t t" using assms 1 by simp qed lemma Obj_in_Hom: assumes "Obj t" shows "t \ HHom t t" and "t \ VHom t t" proof - have 1: "Obj t \ t \ HHom t t \ t \ VHom t t" by (induct t, auto) show "t \ HHom t t" using assms 1 by simp show "t \ VHom t t" using assms 1 by simp qed text \ A formal arrow is ``canonical'' if the only primitive arrows used in its construction are objects and identities. \ primrec Can :: "'a term \ bool" where "Can \<^bold>\\\<^bold>\\<^sub>0 = obj \" | "Can \<^bold>\\\<^bold>\ = ide \" | "Can (t \<^bold>\ u) = (Can t \ Can u \ Src t = Trg u)" | "Can (t \<^bold>\ u) = (Can t \ Can u \ Dom t = Cod u)" | "Can \<^bold>\\<^bold>[t\<^bold>] = Can t" | "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Can t" | "Can \<^bold>\\<^bold>[t\<^bold>] = Can t" | "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Can t" | "Can \<^bold>\\<^bold>[t, u, v\<^bold>] = (Can t \ Can u \ Can v \ Src t = Trg u \ Src u = Trg v)" | "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = (Can t \ Can u \ Can v \ Src t = Trg u \ Src u = Trg v)" lemma Ide_implies_Can: shows "Ide t \ Can t" by (induct t, auto) lemma Can_implies_Arr: shows "Can t \ Arr t" by (induct t, auto) text \ Canonical arrows can be formally inverted. \ primrec Inv :: "'a term \ 'a term" where "Inv \<^bold>\\\<^bold>\\<^sub>0 = \<^bold>\\\<^bold>\\<^sub>0" | "Inv \<^bold>\\\<^bold>\ = \<^bold>\inv \\<^bold>\" | "Inv (t \<^bold>\ u) = (Inv t \<^bold>\ Inv u)" | "Inv (t \<^bold>\ u) = (Inv u \<^bold>\ Inv t)" | "Inv \<^bold>\\<^bold>[t\<^bold>] = \<^bold>\\<^sup>-\<^sup>1\<^bold>[Inv t\<^bold>]" | "Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = \<^bold>\\<^bold>[Inv t\<^bold>]" | "Inv \<^bold>\\<^bold>[t\<^bold>] = \<^bold>\\<^sup>-\<^sup>1\<^bold>[Inv t\<^bold>]" | "Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = \<^bold>\\<^bold>[Inv t\<^bold>]" | "Inv \<^bold>\\<^bold>[t, u, v\<^bold>] = \<^bold>\\<^sup>-\<^sup>1\<^bold>[Inv t, Inv u, Inv v\<^bold>]" | "Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = \<^bold>\\<^bold>[Inv t, Inv u, Inv v\<^bold>]" lemma Src_Inv [simp]: shows "Can t \ Src (Inv t) = Src t" using Can_implies_Arr VSeq_implies_HPar apply (induct t, auto) by metis lemma Trg_Inv [simp]: shows "Can t \ Trg (Inv t) = Trg t" using Can_implies_Arr VSeq_implies_HPar apply (induct t, auto) by metis lemma Dom_Inv [simp]: shows "Can t \ Dom (Inv t) = Cod t" by (induct t, auto) lemma Cod_Inv [simp]: shows "Can t \ Cod (Inv t) = Dom t" by (induct t, auto) lemma Inv_preserves_Ide: shows "Ide t \ Ide (Inv t)" using Ide_implies_Can by (induct t, auto) lemma Can_Inv [simp]: shows "Can t \ Can (Inv t)" by (induct t, auto) lemma Inv_in_Hom [intro]: assumes "Can t" shows "Inv t \ HHom (Src t) (Trg t)" and "Inv t \ VHom (Cod t) (Dom t)" using assms Can_Inv Can_implies_Arr by simp_all lemma Inv_Ide [simp]: assumes "Ide a" shows "Inv a = a" using assms by (induct a, auto) lemma Inv_Inv [simp]: assumes "Can t" shows "Inv (Inv t) = t" using assms by (induct t, auto) subsection "Normal Terms" text \ We call a term ``normal'' if it is either a formal object or it is constructed from primitive arrows using only horizontal composition associated to the right. Essentially, such terms are (typed) composable sequences of arrows of @{term V}, where the empty list is represented by a formal object and \\<^bold>\\ is used as the list constructor. \ fun Nml :: "'a term \ bool" where "Nml \<^bold>\\\<^bold>\\<^sub>0 = obj \" | "Nml \<^bold>\\\<^bold>\ = arr \" | "Nml (\<^bold>\\\<^bold>\ \<^bold>\ u) = (arr \ \ Nml u \ \ is_Prim\<^sub>0 u \ \<^bold>\src \\<^bold>\\<^sub>0 = Trg u)" | "Nml _ = False" lemma Nml_HcompD: assumes "Nml (t \<^bold>\ u)" shows "\<^bold>\un_Prim t\<^bold>\ = t" and "arr (un_Prim t)" and "Nml t" and "Nml u" and "\ is_Prim\<^sub>0 u" and "\<^bold>\src (un_Prim t)\<^bold>\\<^sub>0 = Trg u" and "Src t = Trg u" proof - have 1: "t = \<^bold>\un_Prim t\<^bold>\ \ arr (un_Prim t) \ Nml t \ Nml u \ \ is_Prim\<^sub>0 u \ \<^bold>\src (un_Prim t)\<^bold>\\<^sub>0 = Trg u" using assms by (cases t; simp; cases u; simp) show "\<^bold>\un_Prim t\<^bold>\ = t" using 1 by simp show "arr (un_Prim t)" using 1 by simp show "Nml t" using 1 by simp show "Nml u" using 1 by simp show "\ is_Prim\<^sub>0 u" using 1 by simp show "\<^bold>\src (un_Prim t)\<^bold>\\<^sub>0 = Trg u" using 1 by simp show "Src t = Trg u" using assms apply (cases t) by simp_all qed lemma Nml_implies_Arr: shows "Nml t \ Arr t" by (induct t, auto simp add: Nml_HcompD) lemma Nml_Src [simp]: shows "Nml t \ Nml (Src t)" apply (induct t, simp_all) using Nml_HcompD by metis lemma Nml_Trg [simp]: shows "Nml t \ Nml (Trg t)" apply (induct t, simp_all) using Nml_HcompD by metis lemma Nml_Dom [simp]: shows "Nml t \ Nml (Dom t)" proof (induct t, simp_all add: Nml_HcompD) fix u v assume I1: "Nml (Dom u)" assume I2: "Nml (Dom v)" assume uv: "Nml (u \<^bold>\ v)" show "Nml (Dom u \<^bold>\ Dom v)" proof - have 1: "is_Prim (Dom u) \ arr (un_Prim (Dom u)) \ Dom u = \<^bold>\dom (un_Prim u)\<^bold>\" using uv by (cases u; simp; cases v, simp_all) have 2: "Nml v \ \ is_Prim\<^sub>0 v \ \ is_Vcomp v \ \ is_Lunit' v \ \ is_Runit' v" using uv by (cases u, simp_all; cases v, simp_all) have "arr (dom (un_Prim u))" using 1 by fastforce moreover have "Nml (Dom v) \ \ is_Prim\<^sub>0 v" using 2 I2 by (cases v, simp_all) moreover have "\<^bold>\src (dom (un_Prim u))\<^bold>\\<^sub>0 = Trg (Dom v)" proof - have "Trg (Dom v) = Src (Dom u)" using uv Nml_implies_Arr by fastforce also have "... = \<^bold>\src (dom (un_Prim u))\<^bold>\\<^sub>0" using 1 by fastforce finally show ?thesis by argo qed moreover have "\ is_Prim\<^sub>0 (Dom v)" using 2 by (cases v, simp_all) ultimately show ?thesis using 1 2 by simp qed qed lemma Nml_Cod [simp]: shows "Nml t \ Nml (Cod t)" proof (induct t, simp_all add: Nml_HcompD) fix u v assume I1: "Nml (Cod u)" assume I2: "Nml (Cod v)" assume uv: "Nml (u \<^bold>\ v)" show "Nml (Cod u \<^bold>\ Cod v)" proof - have 1: "is_Prim (Cod u) \ arr (un_Prim (Cod u)) \ Cod u = \<^bold>\cod (un_Prim u)\<^bold>\" using uv by (cases u; simp; cases v, simp_all) have 2: "Nml v \ \ is_Prim\<^sub>0 v \ \ is_Vcomp v \ \ is_Lunit' v \ \ is_Runit' v" using uv by (cases u; simp; cases v, simp_all) have "arr (cod (un_Prim u))" using 1 by fastforce moreover have "Nml (Cod v) \ \ is_Prim\<^sub>0 v" using 2 I2 by (cases v, simp_all) moreover have "\<^bold>\src (cod (un_Prim u))\<^bold>\\<^sub>0 = Trg (Cod v)" proof - have "Trg (Cod v) = Src (Cod u)" using uv Nml_implies_Arr by fastforce also have "... = \<^bold>\src (cod (un_Prim u))\<^bold>\\<^sub>0" using 1 by fastforce finally show ?thesis by argo qed moreover have "\ is_Prim\<^sub>0 (Cod v)" using 2 by (cases v; simp) ultimately show ?thesis using 1 2 by simp qed qed lemma Nml_Inv [simp]: assumes "Can t" and "Nml t" shows "Nml (Inv t)" proof - have "Can t \ Nml t \ Nml (Inv t)" apply (induct t, simp_all) proof - fix u v assume I1: "Nml u \ Nml (Inv u)" assume I2: "Nml v \ Nml (Inv v)" assume uv: "Can u \ Can v \ Src u = Trg v \ Nml (u \<^bold>\ v)" show "Nml (Inv u \<^bold>\ Inv v)" proof - have u: "Arr u \ Can u" using uv Can_implies_Arr by blast have v: "Arr v \ Can v" using uv Can_implies_Arr by blast have 1: "\<^bold>\un_Prim u\<^bold>\ = u \ ide (un_Prim u) \ Nml u \ Nml v \ \ is_Prim\<^sub>0 v \ \<^bold>\src (un_Prim u)\<^bold>\\<^sub>0 = Trg v" using uv Nml_HcompD [of u v] apply simp using uv by (cases u, simp_all) have 2: "\<^bold>\un_Prim (Inv u)\<^bold>\ = Inv u \ arr (un_Prim (Inv u)) \ Nml (Inv u)" using 1 by (cases u; simp) moreover have "Nml (Inv v) \ \ is_Prim\<^sub>0 (Inv v)" using 1 I2 by (cases v, simp_all) moreover have "\<^bold>\src (un_Prim (Inv u))\<^bold>\\<^sub>0 = Trg (Inv v)" using 1 2 v by (cases u, simp_all) ultimately show ?thesis by (cases u, simp_all) qed qed thus ?thesis using assms by blast qed text \ The following function defines a horizontal composition for normal terms. If such terms are regarded as lists, this is just (typed) list concatenation. \ fun HcompNml (infixr "\<^bold>\\<^bold>\\<^bold>\" 53) where "\<^bold>\\\<^bold>\\<^sub>0 \<^bold>\\<^bold>\\<^bold>\ u = u" | "t \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\\<^bold>\\<^sub>0 = t" | "\<^bold>\\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u = \<^bold>\\\<^bold>\ \<^bold>\ u" | "(t \<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" | "t \<^bold>\\<^bold>\\<^bold>\ u = undefined" lemma HcompNml_Prim [simp]: assumes "\ is_Prim\<^sub>0 t" shows "\<^bold>\\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ t = \<^bold>\\\<^bold>\ \<^bold>\ t" using assms by (cases t, simp_all) lemma HcompNml_term_Prim\<^sub>0 [simp]: assumes "Src t = Trg \<^bold>\\\<^bold>\\<^sub>0" shows "t \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\\<^bold>\\<^sub>0 = t" using assms by (cases t, simp_all) lemma HcompNml_Nml: assumes "Nml (t \<^bold>\ u)" shows "t \<^bold>\\<^bold>\\<^bold>\ u = t \<^bold>\ u" using assms HcompNml_Prim by (metis Nml_HcompD(1) Nml_HcompD(5)) lemma Nml_HcompNml: assumes "Nml t" and "Nml u" and "Src t = Trg u" shows "Nml (t \<^bold>\\<^bold>\\<^bold>\ u)" and "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom t \<^bold>\\<^bold>\\<^bold>\ Dom u" and "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t \<^bold>\\<^bold>\\<^bold>\ Cod u" and "Src (t \<^bold>\\<^bold>\\<^bold>\ u) = Src u" and "Trg (t \<^bold>\\<^bold>\\<^bold>\ u) = Trg t" proof - have 0: "\u. \ Nml t; Nml u; Src t = Trg u \ \ Nml (t \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom t \<^bold>\\<^bold>\\<^bold>\ Dom u \ Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t \<^bold>\\<^bold>\\<^bold>\ Cod u \ Src (t \<^bold>\\<^bold>\\<^bold>\ u) = Src u \ Trg (t \<^bold>\\<^bold>\\<^bold>\ u) = Trg t" proof (induct t, simp_all add: Nml_HcompD(1-4)) fix \ :: 'a and u :: "'a term" assume \: "arr \" assume u: "Nml u" assume 1: "\<^bold>\src \\<^bold>\\<^sub>0 = Trg u" show "Nml (\<^bold>\\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (\<^bold>\\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = \<^bold>\dom \\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Dom u \ Cod (\<^bold>\\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = \<^bold>\cod \\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Cod u \ Src (\<^bold>\\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = Src u \ Trg (\<^bold>\\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = \<^bold>\trg \\<^bold>\\<^sub>0" using u \ 1 by (cases u, simp_all) next fix u v w assume I1: "\x. Nml x \ Src v = Trg x \ Nml (v \<^bold>\\<^bold>\\<^bold>\ x) \ Dom (v \<^bold>\\<^bold>\\<^bold>\ x) = Dom v \<^bold>\\<^bold>\\<^bold>\ Dom x \ Cod (v \<^bold>\\<^bold>\\<^bold>\ x) = Cod v \<^bold>\\<^bold>\\<^bold>\ Cod x \ Src (v \<^bold>\\<^bold>\\<^bold>\ x) = Src x \ Trg (v \<^bold>\\<^bold>\\<^bold>\ x) = Trg v" assume I2: "\x. Nml x \ Trg u = Trg x \ Nml (w \<^bold>\\<^bold>\\<^bold>\ x) \ Dom (w \<^bold>\\<^bold>\\<^bold>\ x) = Dom w \<^bold>\\<^bold>\\<^bold>\ Dom x \ Cod (w \<^bold>\\<^bold>\\<^bold>\ x) = Cod w \<^bold>\\<^bold>\\<^bold>\ Cod x \ Src (w \<^bold>\\<^bold>\\<^bold>\ x) = Src x \ Trg (w \<^bold>\\<^bold>\\<^bold>\ x) = Trg w" assume vw: "Nml (v \<^bold>\ w)" assume u: "Nml u" assume wu: "Src w = Trg u" show "Nml ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) \ Dom ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = (Dom v \<^bold>\ Dom w) \<^bold>\\<^bold>\\<^bold>\ Dom u \ Cod ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = (Cod v \<^bold>\ Cod w) \<^bold>\\<^bold>\\<^bold>\ Cod u \ Src ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = Src u \ Trg ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = Trg v" proof - have v: "v = \<^bold>\un_Prim v\<^bold>\ \ Nml v" using vw Nml_implies_Arr Nml_HcompD by metis have w: "Nml w \ \ is_Prim\<^sub>0 w \ \<^bold>\src (un_Prim v)\<^bold>\\<^sub>0 = Trg w" using vw by (simp add: Nml_HcompD) have "is_Prim\<^sub>0 u \ ?thesis" by (cases u; simp add: vw wu) moreover have "\ is_Prim\<^sub>0 u \ ?thesis" proof - assume 1: "\ is_Prim\<^sub>0 u" have "Src v = Trg (w \<^bold>\\<^bold>\\<^bold>\ u)" using u v w I2 [of u] by (cases v, simp_all) hence "Nml (v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u) = Dom v \<^bold>\\<^bold>\\<^bold>\ Dom (w \<^bold>\\<^bold>\\<^bold>\ u) \ Cod (v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u) = Cod v \<^bold>\\<^bold>\\<^bold>\ Cod (w \<^bold>\\<^bold>\\<^bold>\ u) \ Src (v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u) = Src u \ Trg (v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u) = Trg v" using u v w I1 [of "w \<^bold>\\<^bold>\\<^bold>\ u"] I2 [of u] by argo moreover have "v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u = (v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u" using 1 by (cases u, simp_all) moreover have "(Dom v \<^bold>\ Dom w) \<^bold>\\<^bold>\\<^bold>\ Dom u = Dom v \<^bold>\\<^bold>\\<^bold>\ Dom (w \<^bold>\\<^bold>\\<^bold>\ u)" using v w u vw 1 I2 Nml_Dom HcompNml_Prim Nml_HcompD(1) Nml_HcompD(5) by (cases u, simp_all) moreover have "(Cod v \<^bold>\ Cod w) \<^bold>\\<^bold>\\<^bold>\ Cod u = Cod v \<^bold>\\<^bold>\\<^bold>\ Cod (w \<^bold>\\<^bold>\\<^bold>\ u)" using v w u vw 1 I2 Nml_HcompD(1) Nml_HcompD(5) HcompNml_Prim by (cases u, simp_all) ultimately show ?thesis by argo qed ultimately show ?thesis by blast qed next fix a u assume a: "obj a" assume u: "Nml u" assume au: "\<^bold>\a\<^bold>\\<^sub>0 = Trg u" show "Nml (Trg u \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (Trg u \<^bold>\\<^bold>\\<^bold>\ u) = Dom (Trg u) \<^bold>\\<^bold>\\<^bold>\ Dom u \ Cod (Trg u \<^bold>\\<^bold>\\<^bold>\ u) = Cod (Trg u) \<^bold>\\<^bold>\\<^bold>\ Cod u \ Src (Trg u \<^bold>\\<^bold>\\<^bold>\ u) = Src u \ Trg (Trg u \<^bold>\\<^bold>\\<^bold>\ u) = Trg (Trg u)" using au by (metis Cod.simps(1) Dom.simps(1) HcompNml.simps(1) Trg.simps(1) u) qed show "Nml (t \<^bold>\\<^bold>\\<^bold>\ u) " using assms 0 by blast show "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom t \<^bold>\\<^bold>\\<^bold>\ Dom u" using assms 0 by blast show "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t \<^bold>\\<^bold>\\<^bold>\ Cod u" using assms 0 by blast show "Src (t \<^bold>\\<^bold>\\<^bold>\ u) = Src u" using assms 0 by blast show "Trg (t \<^bold>\\<^bold>\\<^bold>\ u) = Trg t" using assms 0 by blast qed lemma HcompNml_in_Hom [intro]: assumes "Nml t" and "Nml u" and "Src t = Trg u" shows "t \<^bold>\\<^bold>\\<^bold>\ u \ HHom (Src u) (Trg t)" and "t \<^bold>\\<^bold>\\<^bold>\ u \ VHom (Dom t \<^bold>\\<^bold>\\<^bold>\ Dom u) (Cod t \<^bold>\\<^bold>\\<^bold>\ Cod u)" using assms Nml_HcompNml Nml_implies_Arr by auto lemma Src_HcompNml: assumes "Nml t" and "Nml u" and "Src t = Trg u" shows "Src (t \<^bold>\\<^bold>\\<^bold>\ u) = Src u" using assms Nml_HcompNml(4) by simp lemma Trg_HcompNml: assumes "Nml t" and "Nml u" and "Src t = Trg u" shows "Trg (t \<^bold>\\<^bold>\\<^bold>\ u) = Trg t" using assms Nml_HcompNml(5) by simp lemma Dom_HcompNml: assumes "Nml t" and "Nml u" and "Src t = Trg u" shows "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom t \<^bold>\\<^bold>\\<^bold>\ Dom u" using assms Nml_HcompNml(2) by simp lemma Cod_HcompNml: assumes "Nml t" and "Nml u" and "Src t = Trg u" shows "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t \<^bold>\\<^bold>\\<^bold>\ Cod u" using assms Nml_HcompNml(3) by simp lemma is_Hcomp_HcompNml: assumes "Nml t" and "Nml u" and "Src t = Trg u" and "\ is_Prim\<^sub>0 t" and "\ is_Prim\<^sub>0 u" shows "is_Hcomp (t \<^bold>\\<^bold>\\<^bold>\ u)" proof - have "\ \ is_Hcomp (t \<^bold>\\<^bold>\\<^bold>\ u); Nml t; Nml u; Src t = Trg u; \ is_Prim\<^sub>0 t; \ is_Prim\<^sub>0 u \ \ False" proof (induct t, simp_all add: Nml_HcompD) fix a assume a: "obj a" assume u: "Nml u" assume 1: "\ is_Hcomp u" assume 2: "\ is_Prim\<^sub>0 (Trg u)" show "False" using u 1 2 by (cases u; simp) next fix v w assume I2: "\ is_Hcomp (w \<^bold>\\<^bold>\\<^bold>\ u) \ False" assume vw: "Nml (v \<^bold>\ w)" assume u: "Nml u" assume 1: "\ is_Hcomp ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u)" assume 2: "\ is_Prim\<^sub>0 u" assume 3: "Src w = Trg u" show "False" proof - have v: "v = \<^bold>\un_Prim v\<^bold>\" using vw Nml_HcompD by force have w: "Nml w \ \ is_Prim\<^sub>0 w \ \<^bold>\src (un_Prim v)\<^bold>\\<^sub>0 = Trg w" using vw Nml_HcompD [of v w] by blast have "(v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u = v \<^bold>\ (w \<^bold>\\<^bold>\\<^bold>\ u)" proof - have "(v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u = \<^bold>\un_Prim v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (w \<^bold>\\<^bold>\\<^bold>\ u)" using u v 2 by (cases u, simp_all) also have "... = \<^bold>\un_Prim v\<^bold>\ \<^bold>\ (w \<^bold>\\<^bold>\\<^bold>\ u)" using u w I2 by fastforce also have "... = v \<^bold>\ (w \<^bold>\\<^bold>\\<^bold>\ u)" using v by simp finally show ?thesis by simp qed thus ?thesis using 1 by simp qed qed thus ?thesis using assms by blast qed text \ The following function defines the ``dimension'' of a term, which is the number of inputs (or outputs) when the term is regarded as an interconnection matrix. For normal terms, this is just the length of the term when regarded as a list of arrows of @{term C}. This function is used as a ranking of terms in the subsequent associativity proof. \ primrec dim :: "'a term \ nat" where "dim \<^bold>\\\<^bold>\\<^sub>0 = 0" | "dim \<^bold>\\\<^bold>\ = 1" | "dim (t \<^bold>\ u) = (dim t + dim u)" | "dim (t \<^bold>\ u) = dim t" | "dim \<^bold>\\<^bold>[t\<^bold>] = dim t" | "dim \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = dim t" | "dim \<^bold>\\<^bold>[t\<^bold>] = dim t" | "dim \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = dim t" | "dim \<^bold>\\<^bold>[t, u, v\<^bold>] = dim t + dim u + dim v" | "dim \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = dim t + dim u + dim v" lemma HcompNml_assoc: assumes "Nml t" and "Nml u" and "Nml v" and "Src t = Trg u" and "Src u = Trg v" shows "(t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - have "\n t u v. \ dim t = n; Nml t; Nml u; Nml v; Src t = Trg u; Src u = Trg v \ \ (t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - fix n show "\t u v. \ dim t = n; Nml t; Nml u; Nml v; Src t = Trg u; Src u = Trg v \ \ (t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof (induction n rule: nat_less_induct) fix n :: nat and t :: "'a term" and u v assume I: "\mt u v. dim t = m \ Nml t \ Nml u \ Nml v \ Src t = Trg u \ Src u = Trg v \ (t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" assume dim: "dim t = n" assume t: "Nml t" assume u: "Nml u" assume v: "Nml v" assume tu: "Src t = Trg u" assume uv: "Src u = Trg v" show "(t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - have "is_Prim\<^sub>0 t \ ?thesis" by (cases t; simp) moreover have "\is_Prim\<^sub>0 t \ is_Prim\<^sub>0 u \ ?thesis" by (cases t; simp; cases u; simp) moreover have "\ is_Prim\<^sub>0 t \ \ is_Prim\<^sub>0 u \ is_Prim\<^sub>0 v \ ?thesis" proof - assume 1: "\ is_Prim\<^sub>0 t" assume 2: "\ is_Prim\<^sub>0 u" assume 3: "is_Prim\<^sub>0 v" have "\is_Prim\<^sub>0 (t \<^bold>\\<^bold>\\<^bold>\ u)" using 1 2 t u tu is_Hcomp_HcompNml [of t u] by (cases t, simp, cases u, auto) thus ?thesis using 1 2 3 tu uv by (cases v, simp, cases "t \<^bold>\\<^bold>\\<^bold>\ u", simp_all) qed moreover have "\is_Prim\<^sub>0 t \ \ is_Prim\<^sub>0 u \ \ is_Prim\<^sub>0 v \ is_Prim t \ ?thesis" using v by (cases t, simp_all, cases u, simp_all; cases v, simp_all) moreover have "\is_Prim\<^sub>0 t \ \ is_Prim\<^sub>0 u \ \ is_Prim\<^sub>0 v \ is_Hcomp t \ ?thesis" proof (cases t, simp_all) fix w :: "'a term" and x :: "'a term" assume 1: " \ is_Prim\<^sub>0 u \ \ is_Prim\<^sub>0 v" assume 2: "t = (w \<^bold>\ x)" show "((w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = (w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - have w: "w = \<^bold>\un_Prim w\<^bold>\" using t 1 2 Nml_HcompD by auto have x: "Nml x" using t w 1 2 by (metis Nml.simps(3)) have "((w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = (w \<^bold>\\<^bold>\\<^bold>\ (x \<^bold>\\<^bold>\\<^bold>\ u)) \<^bold>\\<^bold>\\<^bold>\ v" using u v w x 1 2 by (cases u, simp_all) also have "... = (w \<^bold>\ (x \<^bold>\\<^bold>\\<^bold>\ u)) \<^bold>\\<^bold>\\<^bold>\ v" using t w u 1 2 HcompNml_Prim is_Hcomp_HcompNml Nml_HcompD by (metis Src.simps(3) term.distinct_disc(3) tu) also have "... = w \<^bold>\\<^bold>\\<^bold>\ ((x \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v)" using u v w x 1 by (cases u, simp_all; cases v, simp_all) also have "... = w \<^bold>\\<^bold>\\<^bold>\ (x \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v))" proof - have "dim x < dim t" using 2 w by (cases w; simp) moreover have "Src x = Trg u \ Src u = Trg v" using tu uv 2 by auto ultimately show ?thesis using u v x dim I by simp qed also have "... = (w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - have 3: "is_Hcomp (u \<^bold>\\<^bold>\\<^bold>\ v)" using u v uv 1 is_Hcomp_HcompNml by auto obtain u' :: "'a term" and v' where uv': "u \<^bold>\\<^bold>\\<^bold>\ v = u' \<^bold>\ v'" using 3 is_Hcomp_def by blast thus ?thesis by simp qed finally show ?thesis by simp qed qed moreover have "is_Prim\<^sub>0 t \ is_Prim t \ is_Hcomp t" using t by (cases t, simp_all) ultimately show ?thesis by blast qed qed qed thus ?thesis using assms by blast qed lemma HcompNml_Trg_Nml: assumes "Nml t" shows "Trg t \<^bold>\\<^bold>\\<^bold>\ t = t" proof - have "Nml t \ Trg t \<^bold>\\<^bold>\\<^bold>\ t = t" proof (induct t, simp_all add: Nml_HcompD) fix u v assume uv: "Nml (u \<^bold>\ v)" assume I1: "Trg u \<^bold>\\<^bold>\\<^bold>\ u = u" have 1: "Nml u \ Nml v \ Src u = Trg v" using uv Nml_HcompD by blast have 2: "Trg (u \<^bold>\ v) = Trg u" using uv by auto show "Trg u \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\ v = u \<^bold>\ v" proof - have "Trg u \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\ v = Trg u \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\\<^bold>\\<^bold>\ v" using uv HcompNml_Nml by simp also have "... = (Trg u \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v" using 1 2 HcompNml_assoc Src_Trg Nml_Trg Nml_implies_Arr by simp also have "... = u \<^bold>\ v" using I1 uv HcompNml_Nml by simp finally show ?thesis by simp qed qed thus ?thesis using assms by simp qed lemma HcompNml_Nml_Src: assumes "Nml t" shows "t \<^bold>\\<^bold>\\<^bold>\ Src t = t" proof - have "Nml t \ t \<^bold>\\<^bold>\\<^bold>\ Src t = t" proof (induct t, simp_all add: Nml_HcompD) fix u v assume uv: "Nml (u \<^bold>\ v)" assume I2: "v \<^bold>\\<^bold>\\<^bold>\ Src v = v" have 1: "Nml u \ Nml v \ Src u = Trg v" using uv Nml_HcompD by blast have 2: "Src (u \<^bold>\ v) = Src v" using uv Trg_HcompNml by auto show "(u \<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ Src v = u \<^bold>\ v" proof - have "(u \<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ Src v = (u \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ Src v" using uv HcompNml_Nml by simp also have "... = u \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ Src v)" using 1 2 HcompNml_assoc Trg_Src Nml_Src Nml_implies_Arr by simp also have "... = u \<^bold>\ v" using I2 uv HcompNml_Nml by simp finally show ?thesis by simp qed qed thus ?thesis using assms by simp qed lemma HcompNml_Obj_Nml: assumes "Obj t" and "Nml u" and "Src t = Trg u" shows "t \<^bold>\\<^bold>\\<^bold>\ u = u" using assms by (cases t, simp_all add: HcompNml_Trg_Nml) lemma HcompNml_Nml_Obj: assumes "Nml t" and "Obj u" and "Src t = Trg u" shows "t \<^bold>\\<^bold>\\<^bold>\ u = t" using assms by (cases u, simp_all) lemma Ide_HcompNml: assumes "Ide t" and "Ide u" and "Nml t" and "Nml u" and "Src t = Trg u" shows "Ide (t \<^bold>\\<^bold>\\<^bold>\ u)" using assms by (metis (mono_tags, lifting) Nml_HcompNml(1) Nml_implies_Arr Dom_HcompNml Ide_Dom Ide_in_Hom(2) mem_Collect_eq) lemma Can_HcompNml: assumes "Can t" and "Can u" and "Nml t" and "Nml u" and "Src t = Trg u" shows "Can (t \<^bold>\\<^bold>\\<^bold>\ u)" proof - have "\u. \ Can t \ Nml t; Can u \ Nml u; Src t = Trg u \ \ Can (t \<^bold>\\<^bold>\\<^bold>\ u)" proof (induct t, simp_all add: HcompNml_Trg_Nml HcompNml_Nml_Src) show "\x u. ide x \ arr x \ Can u \ Nml u \ \<^bold>\src x\<^bold>\\<^sub>0 = Trg u \ Can (\<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u)" by (metis Ide.simps(1-2) Ide_implies_Can Can.simps(3) Nml.elims(2) Nml.simps(2) HcompNml.simps(12) HcompNml_Prim Ide_HcompNml Src.simps(2) term.disc(2)) show "\v w u. (\u. Nml v \ Can u \ Nml u \ Trg w = Trg u \ Can (v \<^bold>\\<^bold>\\<^bold>\ u)) \ (\ua. Nml w \ Can ua \ Nml ua \ Trg u = Trg ua \ Can (w \<^bold>\\<^bold>\\<^bold>\ ua)) \ Can v \ Can w \ Src v = Trg w \ Nml (v \<^bold>\ w) \ Can u \ Nml u \ Src w = Trg u \ Can ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u)" by (metis Nml_HcompD(3-4) HcompNml_Nml Nml_HcompNml(1) HcompNml_assoc Trg_HcompNml) qed thus ?thesis using assms by blast qed lemma Inv_HcompNml: assumes "Can t" and "Can u" and "Nml t" and "Nml u" and "Src t = Trg u" shows "Inv (t \<^bold>\\<^bold>\\<^bold>\ u) = Inv t \<^bold>\\<^bold>\\<^bold>\ Inv u" proof - have "\u. \ Can t \ Nml t; Can u \ Nml u; Src t = Trg u \ \ Inv (t \<^bold>\\<^bold>\\<^bold>\ u) = Inv t \<^bold>\\<^bold>\\<^bold>\ Inv u" proof (induct t, simp_all add: HcompNml_Trg_Nml HcompNml_Nml_Src) show "\x u. \ obj x; Can u \ Nml u; \<^bold>\x\<^bold>\\<^sub>0 = Trg u \ \ Inv u = Inv (Trg u) \<^bold>\\<^bold>\\<^bold>\ Inv u" by (metis HcompNml.simps(1) Inv.simps(1)) show "\x u. ide x \ arr x \ Can u \ Nml u \ \<^bold>\src x\<^bold>\\<^sub>0 = Trg u \ Inv (\<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = \<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Inv u" by (metis Ide.simps(2) HcompNml.simps(2) HcompNml_Prim Inv.simps(1,3) Inv_Ide Inv_Inv is_Prim\<^sub>0_def) fix v w u assume I1: "\x. Nml v \ Can x \ Nml x \ Trg w = Trg x \ Inv (v \<^bold>\\<^bold>\\<^bold>\ x) = Inv v \<^bold>\\<^bold>\\<^bold>\ Inv x" assume I2: "\x. Nml w \ Can x \ Nml x \ Trg u = Trg x \ Inv (w \<^bold>\\<^bold>\\<^bold>\ x) = Inv w \<^bold>\\<^bold>\\<^bold>\ Inv x" assume vw: "Can v \ Can w \ Src v = Trg w \ Nml (v \<^bold>\ w)" assume wu: "Src w = Trg u" assume u: "Can u \ Nml u" have v: "Can v \ Nml v" using vw Nml_HcompD by blast have w: "Can w \ Nml w" using v vw by (cases v, simp_all) show "Inv ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = (Inv v \<^bold>\ Inv w) \<^bold>\\<^bold>\\<^bold>\ Inv u" proof - have "is_Prim\<^sub>0 u \ ?thesis" apply (cases u) by simp_all moreover have "\ is_Prim\<^sub>0 u \ ?thesis" proof - assume 1: "\ is_Prim\<^sub>0 u" have "Inv ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = Inv (v \<^bold>\\<^bold>\\<^bold>\ (w \<^bold>\\<^bold>\\<^bold>\ u))" using 1 by (cases u, simp_all) also have "... = Inv v \<^bold>\\<^bold>\\<^bold>\ Inv (w \<^bold>\\<^bold>\\<^bold>\ u)" using u v w vw wu I1 Nml_HcompNml Can_HcompNml Nml_Inv Can_Inv by simp also have "... = Inv v \<^bold>\\<^bold>\\<^bold>\ (Inv w \<^bold>\\<^bold>\\<^bold>\ Inv u)" using u v w I2 Nml_HcompNml by simp also have "... = (Inv v \<^bold>\ Inv w) \<^bold>\\<^bold>\\<^bold>\ Inv u" using v 1 by (cases u, simp_all) finally show ?thesis by blast qed ultimately show ?thesis by blast qed qed thus ?thesis using assms by blast qed text \ The following function defines vertical composition for compatible normal terms, by ``pushing the composition down'' to arrows of @{text V}. \ fun VcompNml :: "'a term \ 'a term \ 'a term" (infixr "\<^bold>\\<^bold>\\<^bold>\" 55) where "\<^bold>\\\<^bold>\\<^sub>0 \<^bold>\\<^bold>\\<^bold>\ u = u" | "\<^bold>\\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\\<^bold>\ = \<^bold>\\ \ \\<^bold>\" | "(u \<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (w \<^bold>\ x) = (u \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\ v \<^bold>\\<^bold>\\<^bold>\ x)" | "t \<^bold>\\<^bold>\\<^bold>\ \<^bold>\\\<^bold>\\<^sub>0 = t" | "t \<^bold>\\<^bold>\\<^bold>\ _ = undefined \<^bold>\ undefined" text \ Note that the last clause above is not relevant to normal terms. We have chosen a provably non-normal value in order to validate associativity. \ lemma Nml_VcompNml: assumes "Nml t" and "Nml u" and "Dom t = Cod u" shows "Nml (t \<^bold>\\<^bold>\\<^bold>\ u)" and "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom u" and "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t" proof - have 0: "\u. \ Nml t; Nml u; Dom t = Cod u \ \ Nml (t \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom u \ Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t" proof (induct t, simp_all add: Nml_HcompD) show "\x u. obj x \ Nml u \ \<^bold>\x\<^bold>\\<^sub>0 = Cod u \ Nml (Cod u \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (Cod u \<^bold>\\<^bold>\\<^bold>\ u) = Dom u \ Cod (Cod u \<^bold>\\<^bold>\\<^bold>\ u) = Cod (Cod u)" by (metis Cod.simps(1) VcompNml.simps(1)) fix \ u assume \: "arr \" assume u: "Nml u" assume 1: "\<^bold>\dom \\<^bold>\ = Cod u" show "Nml (\<^bold>\\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (\<^bold>\\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = Dom u \ Cod (\<^bold>\\\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = \<^bold>\cod \\<^bold>\" using \ u 1 by (cases u, simp_all) next fix u v w assume I2: "\u. \ Nml u; Dom w = Cod u \ \ Nml (w \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (w \<^bold>\\<^bold>\\<^bold>\ u) = Dom u \ Cod (w \<^bold>\\<^bold>\\<^bold>\ u) = Cod w" assume vw: "Nml (v \<^bold>\ w)" have v: "Nml v" using vw Nml_HcompD by force have w: "Nml w" using vw Nml_HcompD by force assume u: "Nml u" assume 1: "(Dom v \<^bold>\ Dom w) = Cod u" show "Nml ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) \ Dom ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = Dom u \ Cod ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = Cod v \<^bold>\ Cod w" using u v w 1 proof (cases u, simp_all) fix x y assume 2: "u = x \<^bold>\ y" have 4: "is_Prim x \ x = \<^bold>\un_Prim x\<^bold>\ \ arr (un_Prim x) \ Nml y \ \ is_Prim\<^sub>0 y" using u 2 by (cases x, cases y, simp_all) have 5: "is_Prim v \ v = \<^bold>\un_Prim v\<^bold>\ \ arr (un_Prim v) \ Nml w \ \ is_Prim\<^sub>0 w" using v w vw by (cases v, simp_all) have 6: "dom (un_Prim v) = cod (un_Prim x) \ Dom w = Cod y" proof - have "\<^bold>\src (un_Prim v)\<^bold>\\<^sub>0 = Trg w" using vw Nml_HcompD [of v w] by simp thus "dom (un_Prim v) = cod (un_Prim x) \ Dom w = Cod y" using 1 2 4 5 apply (cases u, simp_all) by (metis Cod.simps(2) Dom.simps(2) term.simps(2)) qed have "(v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u = \<^bold>\un_Prim v \ un_Prim x\<^bold>\ \<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ y" using 2 4 5 6 VcompNml.simps(2) [of "un_Prim v" "un_Prim x"] by simp moreover have "Nml (\<^bold>\un_Prim v \ un_Prim x\<^bold>\ \<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ y)" proof - have "Nml (w \<^bold>\\<^bold>\\<^bold>\ y)" using I2 4 5 6 by simp moreover have "\ is_Prim\<^sub>0 (w \<^bold>\\<^bold>\\<^bold>\ y)" using vw 4 5 6 I2 Nml_Cod Nml_HcompD(5) is_Prim\<^sub>0_def by (metis Cod.simps(1) Cod.simps(3)) moreover have "\<^bold>\src (un_Prim v \ un_Prim x)\<^bold>\\<^sub>0 = Trg (w \<^bold>\\<^bold>\\<^bold>\ y)" using vw 4 5 6 I2 Nml_HcompD(6) Nml_implies_Arr src.as_nat_trans.is_natural_1 src.as_nat_trans.preserves_comp_2 Trg_Cod src_cod by (metis seqI) ultimately show ?thesis using 4 5 6 Nml.simps(3) [of "un_Prim v \ un_Prim x" "(w \<^bold>\\<^bold>\\<^bold>\ y)"] by simp qed ultimately show "Nml (v \<^bold>\\<^bold>\\<^bold>\ x \<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ y) \ Dom (v \<^bold>\\<^bold>\\<^bold>\ x) = Dom x \ Dom (w \<^bold>\\<^bold>\\<^bold>\ y) = Dom y \ Cod (v \<^bold>\\<^bold>\\<^bold>\ x) = Cod v \ Cod (w \<^bold>\\<^bold>\\<^bold>\ y) = Cod w" using 4 5 6 I2 by (metis (no_types, lifting) Cod.simps(2) Dom.simps(2) VcompNml.simps(2) cod_comp dom_comp seqI) qed qed show "Nml (t \<^bold>\\<^bold>\\<^bold>\ u)" using assms 0 by blast show "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom u" using assms 0 by blast show "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t" using assms 0 by blast qed lemma VcompNml_in_Hom [intro]: assumes "Nml t" and "Nml u" and "Dom t = Cod u" shows "t \<^bold>\\<^bold>\\<^bold>\ u \ HHom (Src u) (Trg u)" and "t \<^bold>\\<^bold>\\<^bold>\ u \ VHom (Dom u) (Cod t)" proof - show 1: "t \<^bold>\\<^bold>\\<^bold>\ u \ VHom (Dom u) (Cod t)" using assms Nml_VcompNml Nml_implies_Arr by simp show "t \<^bold>\\<^bold>\\<^bold>\ u \ HHom (Src u) (Trg u)" using assms 1 Src_Dom Trg_Dom Nml_implies_Arr by fastforce qed lemma Src_VcompNml: assumes "Nml t" and "Nml u" and "Dom t = Cod u" shows "Src (t \<^bold>\\<^bold>\\<^bold>\ u) = Src u" using assms VcompNml_in_Hom by simp lemma Trg_VcompNml: assumes "Nml t" and "Nml u" and "Dom t = Cod u" shows "Trg (t \<^bold>\\<^bold>\\<^bold>\ u) = Trg u" using assms VcompNml_in_Hom by simp lemma Dom_VcompNml: assumes "Nml t" and "Nml u" and "Dom t = Cod u" shows "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom u" using assms Nml_VcompNml(2) by simp lemma Cod_VcompNml: assumes "Nml t" and "Nml u" and "Dom t = Cod u" shows "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t" using assms Nml_VcompNml(3) by simp lemma VcompNml_Cod_Nml [simp]: assumes "Nml t" shows "VcompNml (Cod t) t = t" proof - have "Nml t \ Cod t \<^bold>\\<^bold>\\<^bold>\ t = t" apply (induct t) by (auto simp add: Nml_HcompD comp_cod_arr) thus ?thesis using assms by blast qed lemma VcompNml_Nml_Dom [simp]: assumes "Nml t" shows "t \<^bold>\\<^bold>\\<^bold>\ (Dom t) = t" proof - have "Nml t \ t \<^bold>\\<^bold>\\<^bold>\ Dom t = t" apply (induct t) by (auto simp add: Nml_HcompD comp_arr_dom) thus ?thesis using assms by blast qed lemma VcompNml_Ide_Nml [simp]: assumes "Nml t" and "Ide a" and "Dom a = Cod t" shows "a \<^bold>\\<^bold>\\<^bold>\ t = t" using assms Ide_in_Hom by simp lemma VcompNml_Nml_Ide [simp]: assumes "Nml t" and "Ide a" and "Dom t = Cod a" shows "t \<^bold>\\<^bold>\\<^bold>\ a = t" using assms Ide_in_Hom by auto lemma VcompNml_assoc: assumes "Nml t" and "Nml u" and "Nml v" and "Dom t = Cod u" and "Dom u = Cod v" shows "(t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - have "\u v. \ Nml t; Nml u; Nml v; Dom t = Cod u; Dom u = Cod v \ \ (t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof (induct t, simp_all) show "\x u v. obj x \ Nml u \ Nml v \ \<^bold>\x\<^bold>\\<^sub>0 = Cod u \ Dom u = Cod v \ u \<^bold>\\<^bold>\\<^bold>\ v = Cod u \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\\<^bold>\\<^bold>\ v" by (metis VcompNml.simps(1)) fix f u v assume f: "arr f" assume u: "Nml u" assume v: "Nml v" assume 1: "\<^bold>\dom f\<^bold>\ = Cod u" assume 2: "Dom u = Cod v" show "(\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = \<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" using u v f 1 2 comp_assoc apply (cases u) apply simp_all apply (cases v) by simp_all next fix u v w x assume I1: "\u v. \ Nml w; Nml u; Nml v; Dom w = Cod u; Dom u = Cod v \ \ (w \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = w \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" assume I2: "\u v. \ Nml x; Nml u; Nml v; Dom x = Cod u; Dom u = Cod v \ \ (x \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = x \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" assume wx: "Nml (w \<^bold>\ x)" assume u: "Nml u" assume v: "Nml v" assume 1: "(Dom w \<^bold>\ Dom x) = Cod u" assume 2: "Dom u = Cod v" show "((w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = (w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\\<^bold>\\<^bold>\ v" proof - have w: "Nml w" using wx Nml_HcompD by blast have x: "Nml x" using wx Nml_HcompD by blast have "is_Hcomp u" using u 1 by (cases u) simp_all thus ?thesis using u v apply (cases u, simp_all, cases v, simp_all) proof - fix u1 u2 v1 v2 assume 3: "u = (u1 \<^bold>\ u2)" assume 4: "v = (v1 \<^bold>\ v2)" show "(w \<^bold>\\<^bold>\\<^bold>\ u1) \<^bold>\\<^bold>\\<^bold>\ v1 = w \<^bold>\\<^bold>\\<^bold>\ u1 \<^bold>\\<^bold>\\<^bold>\ v1 \ (x \<^bold>\\<^bold>\\<^bold>\ u2) \<^bold>\\<^bold>\\<^bold>\ v2 = x \<^bold>\\<^bold>\\<^bold>\ u2 \<^bold>\\<^bold>\\<^bold>\ v2" proof - have "Nml u1 \ Nml u2" using u 3 Nml_HcompD by blast moreover have "Nml v1 \ Nml v2" using v 4 Nml_HcompD by blast ultimately show ?thesis using w x I1 I2 1 2 3 4 by simp qed qed qed qed thus ?thesis using assms by blast qed lemma Ide_VcompNml: assumes "Ide t" and "Ide u" and "Nml t" and "Nml u" and "Dom t = Cod u" shows "Ide (t \<^bold>\\<^bold>\\<^bold>\ u)" proof - have "\u. \ Ide t; Ide u; Nml t; Nml u; Dom t = Cod u \ \ Ide (VcompNml t u)" by (induct t, simp_all) thus ?thesis using assms by blast qed lemma Can_VcompNml: assumes "Can t" and "Can u" and "Nml t" and "Nml u" and "Dom t = Cod u" shows "Can (t \<^bold>\\<^bold>\\<^bold>\ u)" proof - have "\u. \ Can t \ Nml t; Can u \ Nml u; Dom t = Cod u \ \ Can (t \<^bold>\\<^bold>\\<^bold>\ u)" proof (induct t, simp_all) fix t u v assume I1: "\v. \ Nml t; Can v \ Nml v; Dom t = Cod v \ \ Can (t \<^bold>\\<^bold>\\<^bold>\ v)" assume I2: "\v. \ Nml u; Can v \ Nml v; Dom u = Cod v \ \ Can (u \<^bold>\\<^bold>\\<^bold>\ v)" assume tu: "Can t \ Can u \ Src t = Trg u \ Nml (t \<^bold>\ u)" have t: "Can t \ Nml t" using tu Nml_HcompD by blast have u: "Can u \ Nml u" using tu Nml_HcompD by blast assume v: "Can v \ Nml v" assume 1: "(Dom t \<^bold>\ Dom u) = Cod v" show "Can ((t \<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v)" proof - have 2: "(Dom t \<^bold>\ Dom u) = Cod v" using 1 by simp show ?thesis using v 2 proof (cases v; simp) fix w x assume wx: "v = (w \<^bold>\ x)" have "Can w \ Nml w" using v wx Nml_HcompD Can.simps(3) by blast moreover have "Can x \ Nml x" using v wx Nml_HcompD Can.simps(3) by blast moreover have "Dom t = Cod w" using 2 wx by simp moreover have ux: "Dom u = Cod x" using 2 wx by simp ultimately show "Can (t \<^bold>\\<^bold>\\<^bold>\ w) \ Can (u \<^bold>\\<^bold>\\<^bold>\ x) \ Src (t \<^bold>\\<^bold>\\<^bold>\ w) = Trg (u \<^bold>\\<^bold>\\<^bold>\ x)" using t u v tu wx I1 I2 by (metis Nml_HcompD(7) Src_VcompNml Trg_VcompNml) qed qed qed thus ?thesis using assms by blast qed lemma Inv_VcompNml: assumes "Can t" and "Can u" and "Nml t" and "Nml u" and "Dom t = Cod u" shows "Inv (t \<^bold>\\<^bold>\\<^bold>\ u) = Inv u \<^bold>\\<^bold>\\<^bold>\ Inv t" proof - have "\u. \ Can t \ Nml t; Can u \ Nml u; Dom t = Cod u \ \ Inv (t \<^bold>\\<^bold>\\<^bold>\ u) = Inv u \<^bold>\\<^bold>\\<^bold>\ Inv t" proof (induct t, simp_all) show "\x u. \ obj x; Can u \ Nml u; \<^bold>\x\<^bold>\\<^sub>0 = Cod u \ \ Inv u = Inv u \<^bold>\\<^bold>\\<^bold>\ Inv (Cod u)" by (simp add: Can_implies_Arr) show "\x u. \ ide x \ arr x; Can u \ Nml u; \<^bold>\x\<^bold>\ = Cod u \ \ Inv u = Inv u \<^bold>\\<^bold>\\<^bold>\ Inv (Cod u)" by (simp add: Can_implies_Arr) fix v w u assume vw: "Can v \ Can w \ Src v = Trg w \ Nml (v \<^bold>\ w)" have v: "Can v \ Nml w" using vw Nml_HcompD by blast have w: "Can w \ Nml w" using vw Nml_HcompD by blast assume I1: "\x. \ Nml v; Can x \ Nml x; Dom v = Cod x \ \ Inv (v \<^bold>\\<^bold>\\<^bold>\ x) = Inv x \<^bold>\\<^bold>\\<^bold>\ Inv v" assume I2: "\x. \ Nml w; Can x \ Nml x; Dom w = Cod x \ \ Inv (w \<^bold>\\<^bold>\\<^bold>\ x) = Inv x \<^bold>\\<^bold>\\<^bold>\ Inv w" assume u: "Can u \ Nml u" assume 1: "(Dom v \<^bold>\ Dom w) = Cod u" show "Inv ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = Inv u \<^bold>\\<^bold>\\<^bold>\ (Inv v \<^bold>\ Inv w)" using v 1 proof (cases w, simp_all) show "\\. obj \ \ Dom v \<^bold>\ \<^bold>\\\<^bold>\\<^sub>0 = Cod u \ w = \<^bold>\\\<^bold>\\<^sub>0 \ Can v \ Inv ((v \<^bold>\ \<^bold>\\\<^bold>\\<^sub>0) \<^bold>\\<^bold>\\<^bold>\ u) = Inv u \<^bold>\\<^bold>\\<^bold>\ (Inv v \<^bold>\ \<^bold>\\\<^bold>\\<^sub>0)" using Nml_HcompD(5) is_Prim\<^sub>0_def vw by blast show "\\. arr \ \ Dom v \<^bold>\ \<^bold>\dom \\<^bold>\ = Cod u \ w = \<^bold>\\\<^bold>\ \ Can v \ Inv ((v \<^bold>\ \<^bold>\\\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ u) = Inv u \<^bold>\\<^bold>\\<^bold>\ (Inv v \<^bold>\ \<^bold>\inv \\<^bold>\)" by (metis Ide.simps(2) Can.simps(2) Nml_HcompD(1) Dom.simps(2) Inv_Ide Dom_Inv Nml_Inv ideD(2) inv_ide VcompNml_Cod_Nml VcompNml_Nml_Dom u vw) show "\y z. Nml (y \<^bold>\ z) \ Dom v \<^bold>\ Dom y \<^bold>\ Dom z = Cod u \ w = y \<^bold>\ z \ Can v \ Inv ((v \<^bold>\ y \<^bold>\ z) \<^bold>\\<^bold>\\<^bold>\ u) = Inv u \<^bold>\\<^bold>\\<^bold>\ (Inv v \<^bold>\ Inv y \<^bold>\ Inv z)" proof - fix y z assume 2: "Nml (y \<^bold>\ z)" assume yz: "w = y \<^bold>\ z" show "Inv ((v \<^bold>\ y \<^bold>\ z) \<^bold>\\<^bold>\\<^bold>\ u) = Inv u \<^bold>\\<^bold>\\<^bold>\ (Inv v \<^bold>\ Inv y \<^bold>\ Inv z)" using u vw yz I1 I2 1 2 VcompNml_Nml_Ide apply (cases u) apply simp_all by (metis Nml_HcompD(3-4)) qed qed qed thus ?thesis using assms by blast qed lemma Can_and_Nml_implies_Ide: assumes "Can t" and "Nml t" shows "Ide t" proof - have "\ Can t; Nml t \ \ Ide t" apply (induct t) by (simp_all add: Nml_HcompD) thus ?thesis using assms by blast qed lemma VcompNml_Can_Inv [simp]: assumes "Can t" and "Nml t" shows "t \<^bold>\\<^bold>\\<^bold>\ Inv t = Cod t" using assms Can_and_Nml_implies_Ide Ide_in_Hom by simp lemma VcompNml_Inv_Can [simp]: assumes "Can t" and "Nml t" shows "Inv t \<^bold>\\<^bold>\\<^bold>\ t = Dom t" using assms Can_and_Nml_implies_Ide Ide_in_Hom by simp text \ The next fact is a syntactic version of the interchange law, for normal terms. \ lemma VcompNml_HcompNml: assumes "Nml t" and "Nml u" and "Nml v" and "Nml w" and "VSeq t v" and "VSeq u w" and "Src v = Trg w" shows "(t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = (t \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" proof - have "\u v w. \ Nml t; Nml u; Nml v; Nml w; VSeq t v; VSeq u w; Src t = Trg u; Src v = Trg w \ \ (t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = (t \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" proof (induct t, simp_all) fix u v w x assume u: "Nml u" assume v: "Nml v" assume w: "Nml w" assume uw: "VSeq u w" show "\x. Arr v \ \<^bold>\x\<^bold>\\<^sub>0 = Cod v \ (Cod v \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = v \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\\<^bold>\\<^bold>\ w" using u v w uw by (cases v) simp_all show "\x. \ arr x; Arr v \ \<^bold>\dom x\<^bold>\ = Cod v; \<^bold>\src x\<^bold>\\<^sub>0 = Trg u; Src v = Trg w \ \ (\<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = \<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ v \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\\<^bold>\\<^bold>\ w" proof - fix x assume x: "arr x" assume 1: "Arr v \ \<^bold>\dom x\<^bold>\ = Cod v" assume tu: "\<^bold>\src x\<^bold>\\<^sub>0 = Trg u" assume vw: "Src v = Trg w" show "(\<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = \<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ v \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\\<^bold>\\<^bold>\ w" proof - have 2: "v = \<^bold>\un_Prim v\<^bold>\ \ arr (un_Prim v)" using v 1 by (cases v) simp_all have "is_Prim\<^sub>0 u \ ?thesis" using u v w x tu uw vw 1 2 Cod.simps(3) VcompNml_Cod_Nml Dom.simps(2) HcompNml_Prim HcompNml_term_Prim\<^sub>0 Nml_HcompNml(3) HcompNml_Trg_Nml apply (cases u) apply simp_all by (cases w, simp_all add: Src_VcompNml) moreover have "\ is_Prim\<^sub>0 u \ ?thesis" proof - assume 3: "\ is_Prim\<^sub>0 u" hence 4: "\ is_Prim\<^sub>0 w" using u w uw by (cases u, simp_all; cases w, simp_all) have "(\<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = (\<^bold>\x\<^bold>\ \<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\ w)" proof - have "\<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u = \<^bold>\x\<^bold>\ \<^bold>\ u" using u x 3 HcompNml_Nml by (cases u, simp_all) moreover have "v \<^bold>\\<^bold>\\<^bold>\ w = v \<^bold>\ w" using w 2 4 HcompNml_Nml by (cases v, simp_all; cases w, simp_all) ultimately show ?thesis by simp qed also have 5: "... = (\<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" by simp also have "... = (\<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" using x u w uw vw 1 2 3 4 5 HcompNml_Nml HcompNml_Prim Nml_HcompNml(1) by (metis Cod.simps(3) Nml.simps(3) Dom.simps(2) Dom.simps(3) Nml_VcompNml(1) tu v) finally show ?thesis by blast qed ultimately show ?thesis by blast qed qed fix t1 t2 assume t12: "Nml (t1 \<^bold>\ t2)" assume I1: "\u v w. \ Nml t1; Nml u; Nml v; Nml w; Arr v \ Dom t1 = Cod v; VSeq u w; Trg t2 = Trg u; Src v = Trg w \ \ (t1 \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = t1 \<^bold>\\<^bold>\\<^bold>\ v \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\\<^bold>\\<^bold>\ w" assume I2: "\u' v w. \ Nml t2; Nml u'; Nml v; Nml w; Arr v \ Dom t2 = Cod v; VSeq u' w; Trg u = Trg u'; Src v = Trg w \ \ (t2 \<^bold>\\<^bold>\\<^bold>\ u') \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = (t2 \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (u' \<^bold>\\<^bold>\\<^bold>\ w)" have t1: "t1 = \<^bold>\un_Prim t1\<^bold>\ \ arr (un_Prim t1) \ Nml t1" using t12 by (cases t1, simp_all) have t2: "Nml t2 \ \is_Prim\<^sub>0 t2" using t12 by (cases t1, simp_all) assume vw: "Src v = Trg w" assume tu: "Src t2 = Trg u" assume 1: "Arr t1 \ Arr t2 \ Src t1 = Trg t2 \ Arr v \ (Dom t1 \<^bold>\ Dom t2) = Cod v" show "((t1 \<^bold>\ t2) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = (t1 \<^bold>\ t2) \<^bold>\\<^bold>\\<^bold>\ v \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\\<^bold>\\<^bold>\ w" proof - have "is_Prim\<^sub>0 u \ ?thesis" using u v w uw tu vw t12 I1 I2 1 Obj_Src apply (cases u) apply simp_all by (cases w, simp_all add: Src_VcompNml) moreover have "\ is_Prim\<^sub>0 u \ ?thesis" proof - assume u': "\ is_Prim\<^sub>0 u" hence w': "\ is_Prim\<^sub>0 w" using u w uw by (cases u, simp_all; cases w, simp_all) show ?thesis using 1 v proof (cases v, simp_all) fix v1 v2 assume v12: "v = v1 \<^bold>\ v2" have v1: "v1 = \<^bold>\un_Prim v1\<^bold>\ \ arr (un_Prim v1) \ Nml v1" using v v12 by (cases v1, simp_all) have v2: "Nml v2 \ \ is_Prim\<^sub>0 v2" using v v12 by (cases v1, simp_all) have 2: "v = (\<^bold>\un_Prim v1\<^bold>\ \<^bold>\ v2)" using v1 v12 by simp show "((t1 \<^bold>\ t2) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ ((v1 \<^bold>\ v2) \<^bold>\\<^bold>\\<^bold>\ w) = (t1 \<^bold>\\<^bold>\\<^bold>\ v1 \<^bold>\ t2 \<^bold>\\<^bold>\\<^bold>\ v2) \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\\<^bold>\\<^bold>\ w" proof - have 3: "(t1 \<^bold>\ t2) \<^bold>\\<^bold>\\<^bold>\ u = t1 \<^bold>\\<^bold>\\<^bold>\ (t2 \<^bold>\\<^bold>\\<^bold>\ u)" using u u' by (cases u, simp_all) have 4: "v \<^bold>\\<^bold>\\<^bold>\ w = v1 \<^bold>\\<^bold>\\<^bold>\ v2 \<^bold>\\<^bold>\\<^bold>\ w" proof - have "Src v1 = Trg v2" using v v12 1 Arr.simps(3) Nml_HcompD(7) by blast moreover have "Src v2 = Trg w" using v v12 vw by simp ultimately show ?thesis using v w v1 v2 v12 2 HcompNml_assoc [of v1 v2 w] HcompNml_Nml by metis qed have "((t1 \<^bold>\ t2) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ ((v1 \<^bold>\ v2) \<^bold>\\<^bold>\\<^bold>\ w) = (t1 \<^bold>\\<^bold>\\<^bold>\ (t2 \<^bold>\\<^bold>\\<^bold>\ u)) \<^bold>\\<^bold>\\<^bold>\ (v1 \<^bold>\\<^bold>\\<^bold>\ (v2 \<^bold>\\<^bold>\\<^bold>\ w))" using 3 4 v12 by simp also have "... = (t1 \<^bold>\\<^bold>\\<^bold>\ v1) \<^bold>\\<^bold>\\<^bold>\ ((t2 \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v2 \<^bold>\\<^bold>\\<^bold>\ w))" proof - have "is_Hcomp (t2 \<^bold>\\<^bold>\\<^bold>\ u)" using t2 u u' tu is_Hcomp_HcompNml by auto moreover have "is_Hcomp (v2 \<^bold>\\<^bold>\\<^bold>\ w)" using v2 v12 w w' vw is_Hcomp_HcompNml by auto ultimately show ?thesis using u u' v w t1 v1 t12 v12 HcompNml_Prim by (metis VcompNml.simps(2) VcompNml.simps(3) is_Hcomp_def term.distinct_disc(3)) qed also have "... = (t1 \<^bold>\\<^bold>\\<^bold>\ v1) \<^bold>\\<^bold>\\<^bold>\ (t2 \<^bold>\\<^bold>\\<^bold>\ v2) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" using u w tu uw vw t2 v2 1 2 Nml_implies_Arr I2 by auto also have "... = ((t1 \<^bold>\\<^bold>\\<^bold>\ v1) \<^bold>\ (t2 \<^bold>\\<^bold>\\<^bold>\ v2)) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" proof - have "\is_Prim\<^sub>0 (u \<^bold>\\<^bold>\\<^bold>\ w)" using u w u' w' by (cases u, simp_all; cases w, simp_all) hence "((t1 \<^bold>\\<^bold>\\<^bold>\ v1) \<^bold>\ (t2 \<^bold>\\<^bold>\\<^bold>\ v2)) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w) = (t1 \<^bold>\\<^bold>\\<^bold>\ v1) \<^bold>\\<^bold>\\<^bold>\ ((t2 \<^bold>\\<^bold>\\<^bold>\ v2) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w))" by (cases "u \<^bold>\\<^bold>\\<^bold>\ w") simp_all thus ?thesis by presburger qed finally show ?thesis by blast qed qed qed ultimately show ?thesis by blast qed qed moreover have "Src t = Trg u" using assms Src_Dom Trg_Dom Src_Cod Trg_Cod Nml_implies_Arr by metis ultimately show ?thesis using assms by simp qed text \ The following function reduces a formal arrow to normal form. \ fun Nmlize :: "'a term \ 'a term" ("\<^bold>\_\<^bold>\") where "\<^bold>\\<^bold>\\\<^bold>\\<^sub>0\<^bold>\ = \<^bold>\\\<^bold>\\<^sub>0" | "\<^bold>\\<^bold>\\\<^bold>\\<^bold>\ = \<^bold>\\\<^bold>\" | "\<^bold>\t \<^bold>\ u\<^bold>\ = \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\" | "\<^bold>\t \<^bold>\ u\<^bold>\ = \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\" | "\<^bold>\\<^bold>\\<^bold>[t\<^bold>]\<^bold>\ = \<^bold>\t\<^bold>\" | "\<^bold>\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<^bold>\ = \<^bold>\t\<^bold>\" | "\<^bold>\\<^bold>\\<^bold>[t\<^bold>]\<^bold>\ = \<^bold>\t\<^bold>\" | "\<^bold>\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<^bold>\ = \<^bold>\t\<^bold>\" | "\<^bold>\\<^bold>\\<^bold>[t, u, v\<^bold>]\<^bold>\ = (\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\" | "\<^bold>\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<^bold>\ = \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\)" lemma Nml_Nmlize: assumes "Arr t" shows "Nml \<^bold>\t\<^bold>\" and "Src \<^bold>\t\<^bold>\ = Src t" and "Trg \<^bold>\t\<^bold>\ = Trg t" and "Dom \<^bold>\t\<^bold>\ = \<^bold>\Dom t\<^bold>\" and "Cod \<^bold>\t\<^bold>\ = \<^bold>\Cod t\<^bold>\" proof - have 0: "Arr t \ Nml \<^bold>\t\<^bold>\ \ Src \<^bold>\t\<^bold>\ = Src t \ Trg \<^bold>\t\<^bold>\ = Trg t \ Dom \<^bold>\t\<^bold>\ = \<^bold>\Dom t\<^bold>\ \ Cod \<^bold>\t\<^bold>\ = \<^bold>\Cod t\<^bold>\" using Nml_HcompNml Nml_VcompNml HcompNml_assoc Src_VcompNml Trg_VcompNml VSeq_implies_HPar apply (induct t) apply auto proof - fix t assume 1: "Arr t" assume 2: "Nml \<^bold>\t\<^bold>\" assume 3: "Src \<^bold>\t\<^bold>\ = Src t" assume 4: "Trg \<^bold>\t\<^bold>\ = Trg t" assume 5: "Dom \<^bold>\t\<^bold>\ = \<^bold>\Dom t\<^bold>\" assume 6: "Cod \<^bold>\t\<^bold>\ = \<^bold>\Cod t\<^bold>\" have 7: "Nml \<^bold>\Dom t\<^bold>\" using 2 5 Nml_Dom by fastforce have 8: "Trg \<^bold>\t\<^bold>\ = \<^bold>\Trg t\<^bold>\" using 1 2 4 Nml_Trg Obj_Trg by (metis Nml.elims(2) Nmlize.simps(1) Nmlize.simps(2) Obj.simps(3)) have 9: "Nml \<^bold>\Cod t\<^bold>\" using 2 6 Nml_Cod by fastforce have 10: "Src \<^bold>\t\<^bold>\ = \<^bold>\Src t\<^bold>\" using 1 2 3 Nml_Src Obj_Src by (metis Nml.elims(2) Nmlize.simps(1) Nmlize.simps(2) Obj.simps(3)) show "\<^bold>\Dom t\<^bold>\ = \<^bold>\Trg t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t\<^bold>\" using 2 5 7 8 Nml_implies_Arr Trg_Dom HcompNml_Trg_Nml by metis show "\<^bold>\Cod t\<^bold>\ = \<^bold>\Trg t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t\<^bold>\" using 2 6 8 9 Nml_implies_Arr Trg_Cod HcompNml_Trg_Nml by metis show "\<^bold>\Dom t\<^bold>\ = \<^bold>\Dom t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Src t\<^bold>\" using 2 5 7 10 Nml_implies_Arr Src_Dom HcompNml_Nml_Src by metis show "\<^bold>\Cod t\<^bold>\ = \<^bold>\Cod t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Src t\<^bold>\" using 2 6 9 10 Nml_implies_Arr Src_Cod HcompNml_Nml_Src by metis next fix t1 t2 t3 assume "Nml \<^bold>\t1\<^bold>\" and "Nml \<^bold>\t2\<^bold>\" and "Nml \<^bold>\t3\<^bold>\" assume "Src t1 = Trg t2" and "Src t2 = Trg t3" assume "Src \<^bold>\t1\<^bold>\ = Trg t2" and "Src \<^bold>\t2\<^bold>\ = Trg t3" assume "Trg \<^bold>\t1\<^bold>\ = Trg t1" and "Trg \<^bold>\t2\<^bold>\ = Trg t2" and "Trg \<^bold>\t3\<^bold>\ = Trg t3" assume "Dom \<^bold>\t1\<^bold>\ = \<^bold>\Dom t1\<^bold>\" and "Cod \<^bold>\t1\<^bold>\ = \<^bold>\Cod t1\<^bold>\" and "Dom \<^bold>\t2\<^bold>\ = \<^bold>\Dom t2\<^bold>\" and "Cod \<^bold>\t2\<^bold>\ = \<^bold>\Cod t2\<^bold>\" and "Dom \<^bold>\t3\<^bold>\ = \<^bold>\Dom t3\<^bold>\" and "Cod \<^bold>\t3\<^bold>\ = \<^bold>\Cod t3\<^bold>\" show "\<^bold>\Dom t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t3\<^bold>\ = (\<^bold>\Dom t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t3\<^bold>\" using Nml_Dom Nml_implies_Arr Src_Dom Trg_Dom HcompNml_assoc [of "\<^bold>\Dom t1\<^bold>\" "\<^bold>\Dom t2\<^bold>\" "\<^bold>\Dom t3\<^bold>\"] \Nml \<^bold>\t1\<^bold>\\ \Nml \<^bold>\t2\<^bold>\\ \Nml \<^bold>\t3\<^bold>\\ \Dom \<^bold>\t1\<^bold>\ = \<^bold>\Dom t1\<^bold>\\ \Dom \<^bold>\t2\<^bold>\ = \<^bold>\Dom t2\<^bold>\\ \Dom \<^bold>\t3\<^bold>\ = \<^bold>\Dom t3\<^bold>\\ \Src \<^bold>\t1\<^bold>\ = Trg t2\ \Trg \<^bold>\t2\<^bold>\ = Trg t2\ \Src \<^bold>\t2\<^bold>\ = Trg t3\ \Trg \<^bold>\t3\<^bold>\ = Trg t3\ by metis show "\<^bold>\Cod t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t3\<^bold>\ = (\<^bold>\Cod t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t3\<^bold>\" using Nml_Cod Nml_implies_Arr Src_Cod Trg_Cod HcompNml_assoc [of "\<^bold>\Cod t1\<^bold>\" "\<^bold>\Cod t2\<^bold>\" "\<^bold>\Cod t3\<^bold>\"] \Nml \<^bold>\t1\<^bold>\\ \Nml \<^bold>\t2\<^bold>\\ \Nml \<^bold>\t3\<^bold>\\ \Cod \<^bold>\t1\<^bold>\ = \<^bold>\Cod t1\<^bold>\\ \Cod \<^bold>\t2\<^bold>\ = \<^bold>\Cod t2\<^bold>\\ \Cod \<^bold>\t3\<^bold>\ = \<^bold>\Cod t3\<^bold>\\ \Src \<^bold>\t1\<^bold>\ = Trg t2\ \Trg \<^bold>\t2\<^bold>\ = Trg t2\ \Src \<^bold>\t2\<^bold>\ = Trg t3\ \Trg \<^bold>\t3\<^bold>\ = Trg t3\ by metis qed show "Nml \<^bold>\t\<^bold>\" using assms 0 by blast show "Src \<^bold>\t\<^bold>\ = Src t" using assms 0 by blast show "Trg \<^bold>\t\<^bold>\ = Trg t" using assms 0 by blast show "Dom \<^bold>\t\<^bold>\ = \<^bold>\Dom t\<^bold>\" using assms 0 by blast show "Cod \<^bold>\t\<^bold>\ = \<^bold>\Cod t\<^bold>\" using assms 0 by blast qed lemma Nmlize_in_Hom [intro]: assumes "Arr t" shows "\<^bold>\t\<^bold>\ \ HHom (Src t) (Trg t)" and "\<^bold>\t\<^bold>\ \ VHom \<^bold>\Dom t\<^bold>\ \<^bold>\Cod t\<^bold>\" using assms Nml_Nmlize Nml_implies_Arr by auto lemma Nmlize_Src: assumes "Arr t" shows "\<^bold>\Src t\<^bold>\ = Src \<^bold>\t\<^bold>\" and "\<^bold>\Src t\<^bold>\ = Src t" proof - have 1: "Obj (Src t)" using assms by simp obtain a where a: "obj a \ Src t = \<^bold>\a\<^bold>\\<^sub>0" using 1 by (cases "Src t", simp_all) show "\<^bold>\Src t\<^bold>\ = Src t" using a by simp thus "\<^bold>\Src t\<^bold>\ = Src \<^bold>\t\<^bold>\" using assms Nmlize_in_Hom by simp qed lemma Nmlize_Trg: assumes "Arr t" shows "\<^bold>\Trg t\<^bold>\ = Trg \<^bold>\t\<^bold>\" and "\<^bold>\Trg t\<^bold>\ = Trg t" proof - have 1: "Obj (Trg t)" using assms by simp obtain a where a: "obj a \ Trg t = \<^bold>\a\<^bold>\\<^sub>0" using 1 by (cases "Trg t", simp_all) show "\<^bold>\Trg t\<^bold>\ = Trg t" using a by simp thus "\<^bold>\Trg t\<^bold>\ = Trg \<^bold>\t\<^bold>\" using assms Nmlize_in_Hom by simp qed lemma Nmlize_Dom: assumes "Arr t" shows "\<^bold>\Dom t\<^bold>\ = Dom \<^bold>\t\<^bold>\" using assms Nmlize_in_Hom by simp lemma Nmlize_Cod: assumes "Arr t" shows "\<^bold>\Cod t\<^bold>\ = Cod \<^bold>\t\<^bold>\" using assms Nmlize_in_Hom by simp lemma Ide_Nmlize_Ide: assumes "Ide t" shows "Ide \<^bold>\t\<^bold>\" proof - have "Ide t \ Ide \<^bold>\t\<^bold>\" using Ide_HcompNml Nml_Nmlize by (induct t, simp_all) thus ?thesis using assms by blast qed lemma Ide_Nmlize_Can: assumes "Can t" shows "Ide \<^bold>\t\<^bold>\" proof - have "Can t \ Ide \<^bold>\t\<^bold>\" using Can_implies_Arr Ide_HcompNml Nml_Nmlize Ide_VcompNml Nml_HcompNml apply (induct t) apply (auto simp add: Dom_Ide Cod_Ide) by (metis Ide_VcompNml) thus ?thesis using assms by blast qed lemma Can_Nmlize_Can: assumes "Can t" shows "Can \<^bold>\t\<^bold>\" using assms Ide_Nmlize_Can Ide_implies_Can by auto lemma Nmlize_Nml [simp]: assumes "Nml t" shows "\<^bold>\t\<^bold>\ = t" proof - have "Nml t \ \<^bold>\t\<^bold>\ = t" apply (induct t, simp_all) using HcompNml_Prim Nml_HcompD by metis thus ?thesis using assms by blast qed lemma Nmlize_Nmlize [simp]: assumes "Arr t" shows "\<^bold>\\<^bold>\t\<^bold>\\<^bold>\ = \<^bold>\t\<^bold>\" using assms Nml_Nmlize Nmlize_Nml by blast lemma Nmlize_Hcomp: assumes "Arr t" and "Arr u" shows "\<^bold>\t \<^bold>\ u\<^bold>\ = \<^bold>\\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\\<^bold>\" using assms Nmlize_Nmlize by simp lemma Nmlize_Hcomp_Obj_Arr [simp]: assumes "Arr u" shows "\<^bold>\\<^bold>\b\<^bold>\\<^sub>0 \<^bold>\ u\<^bold>\ = \<^bold>\u\<^bold>\" using assms by simp lemma Nmlize_Hcomp_Arr_Obj [simp]: assumes "Arr t" and "Src t = \<^bold>\a\<^bold>\\<^sub>0" shows "\<^bold>\t \<^bold>\ \<^bold>\a\<^bold>\\<^sub>0\<^bold>\ = \<^bold>\t\<^bold>\" using assms HcompNml_Nml_Src Nmlize_in_Hom by simp lemma Nmlize_Hcomp_Prim_Arr [simp]: assumes "Arr u" and "\ is_Prim\<^sub>0 \<^bold>\u\<^bold>\" shows "\<^bold>\\<^bold>\\\<^bold>\ \<^bold>\ u\<^bold>\ = \<^bold>\\\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\" using assms by simp lemma Nmlize_Hcomp_Hcomp: assumes "Arr t" and "Arr u" and "Arr v" and "Src t = Trg u" and "Src u = Trg v" shows "\<^bold>\(t \<^bold>\ u) \<^bold>\ v\<^bold>\ = \<^bold>\\<^bold>\t\<^bold>\ \<^bold>\ (\<^bold>\u\<^bold>\ \<^bold>\ \<^bold>\v\<^bold>\)\<^bold>\" using assms Nml_Nmlize Nmlize_Nmlize by (simp add: HcompNml_assoc) lemma Nmlize_Hcomp_Hcomp': assumes "Arr t" and "Arr u" and "Arr v" and "Src t = Trg u" and "Src u = Trg v" shows "\<^bold>\t \<^bold>\ u \<^bold>\ v\<^bold>\ = \<^bold>\\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\ \<^bold>\ \<^bold>\v\<^bold>\\<^bold>\" using assms Nml_Nmlize Nmlize_Nmlize by (simp add: HcompNml_assoc) lemma Nmlize_Vcomp_Cod_Arr: assumes "Arr t" shows "\<^bold>\Cod t \<^bold>\ t\<^bold>\ = \<^bold>\t\<^bold>\" proof - have "Arr t \ \<^bold>\Cod t \<^bold>\ t\<^bold>\ = \<^bold>\t\<^bold>\" proof (induct t, simp_all) show "\x. arr x \ cod x \ x = x" using comp_cod_arr by blast fix t1 t2 show "\t1 t2. \<^bold>\Cod t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t1\<^bold>\ = \<^bold>\t1\<^bold>\ \ \<^bold>\Cod t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\ = \<^bold>\t2\<^bold>\ \ HSeq t1 t2 \ (\<^bold>\Cod t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\) = \<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\" using VcompNml_HcompNml Ide_Cod Nml_Nmlize Nmlize_in_Hom by simp show "\<^bold>\Cod t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t1\<^bold>\ = \<^bold>\t1\<^bold>\ \ \<^bold>\Cod t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\ = \<^bold>\t2\<^bold>\ \ VSeq t1 t2 \ \<^bold>\Cod t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\ = \<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\" using VcompNml_assoc [of "\<^bold>\Cod t1\<^bold>\" "\<^bold>\t1\<^bold>\" "\<^bold>\t2\<^bold>\"] Ide_Cod Nml_Nmlize by simp next show "\t. \<^bold>\Cod t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t\<^bold>\ = \<^bold>\t\<^bold>\ \ Arr t \ (\<^bold>\Trg t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t\<^bold>\ = \<^bold>\t\<^bold>\" by (metis Arr.simps(6) Cod.simps(6) Nmlize.simps(3) Nmlize.simps(6) Nmlize_Cod) show "\t. \<^bold>\Cod t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t\<^bold>\ = \<^bold>\t\<^bold>\ \ Arr t \ (\<^bold>\Cod t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Src t\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t\<^bold>\ = \<^bold>\t\<^bold>\" by (simp add: Nml_Nmlize(1) Nml_Nmlize(2) Nmlize_Src(2) HcompNml_Nml_Obj) show "\t1 t2 t3. \<^bold>\Cod t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t1\<^bold>\ = \<^bold>\t1\<^bold>\ \ \<^bold>\Cod t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\ = \<^bold>\t2\<^bold>\ \ \<^bold>\Cod t3\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t3\<^bold>\ = \<^bold>\t3\<^bold>\ \ Arr t1 \ Arr t2 \ Arr t3 \ Src t1 = Trg t2 \ Src t2 = Trg t3 \ (\<^bold>\Cod t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t3\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ ((\<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t3\<^bold>\) = (\<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t3\<^bold>\" using VcompNml_HcompNml Ide_Cod HcompNml_in_Hom Nml_HcompNml Nml_Nmlize Nmlize_in_Hom HcompNml_assoc by simp show "\t1 t2 t3. \<^bold>\Cod t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t1\<^bold>\ = \<^bold>\t1\<^bold>\ \ \<^bold>\Cod t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\ = \<^bold>\t2\<^bold>\ \ \<^bold>\Cod t3\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t3\<^bold>\ = \<^bold>\t3\<^bold>\ \ Arr t1 \ Arr t2 \ Arr t3 \ Src t1 = Trg t2 \ Src t2 = Trg t3 \ ((\<^bold>\Cod t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t3\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t3\<^bold>\) = \<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t3\<^bold>\" using VcompNml_HcompNml Ide_Cod HcompNml_in_Hom Nml_HcompNml Nml_Nmlize Nmlize_in_Hom HcompNml_assoc by simp qed thus ?thesis using assms by blast qed lemma Nmlize_Vcomp_Arr_Dom: assumes "Arr t" shows "\<^bold>\t \<^bold>\ Dom t\<^bold>\ = \<^bold>\t\<^bold>\" proof - have "Arr t \ \<^bold>\t \<^bold>\ Dom t\<^bold>\ = \<^bold>\t\<^bold>\" proof (induct t, simp_all) show "\x. arr x \ x \ local.dom x = x" using comp_arr_dom by blast fix t1 t2 show "\<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t1\<^bold>\ = \<^bold>\t1\<^bold>\ \ \<^bold>\t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t2\<^bold>\ = \<^bold>\t2\<^bold>\ \ HSeq t1 t2 \ (\<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\Dom t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t2\<^bold>\) = \<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\" using VcompNml_HcompNml Ide_Dom HcompNml_in_Hom Nml_HcompNml Nml_Nmlize Nmlize_in_Hom HcompNml_assoc by simp show "\<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod t2\<^bold>\ = \<^bold>\t1\<^bold>\ \ \<^bold>\t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t2\<^bold>\ = \<^bold>\t2\<^bold>\ \ VSeq t1 t2 \ (\<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t2\<^bold>\ = \<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\" using VcompNml_assoc [of "\<^bold>\t1\<^bold>\" "\<^bold>\t2\<^bold>\" "\<^bold>\Dom t2\<^bold>\"] Ide_Dom Nml_Nmlize by simp next show "\t. \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t\<^bold>\ = \<^bold>\t\<^bold>\ \ Arr t \ \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\Trg t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t\<^bold>\) = \<^bold>\t\<^bold>\" by (simp add: Nml_Nmlize(1) Nml_Nmlize(3) Nmlize_Trg(2) HcompNml_Obj_Nml bicategorical_language.Ide_Dom bicategorical_language_axioms) show "\t. \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t\<^bold>\ = \<^bold>\t\<^bold>\ \ Arr t \ \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\Dom t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Src t\<^bold>\) = \<^bold>\t\<^bold>\" by (simp add: Nml_Nmlize(1) Nml_Nmlize(2) Nmlize_Src(2) HcompNml_Nml_Obj) show "\t1 t2 t3. \<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t1\<^bold>\ = \<^bold>\t1\<^bold>\ \ \<^bold>\t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t2\<^bold>\ = \<^bold>\t2\<^bold>\ \ \<^bold>\t3\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t3\<^bold>\ = \<^bold>\t3\<^bold>\ \ Arr t1 \ Arr t2 \ Arr t3 \ Src t1 = Trg t2 \ Src t2 = Trg t3 \ ((\<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t3\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ ((\<^bold>\Dom t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t3\<^bold>\) = (\<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t3\<^bold>\" using VcompNml_HcompNml Ide_Dom HcompNml_in_Hom Nml_HcompNml Nml_Nmlize Nmlize_in_Hom HcompNml_assoc by simp show "\t1 t2 t3. \<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t1\<^bold>\ = \<^bold>\t1\<^bold>\ \ \<^bold>\t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t2\<^bold>\ = \<^bold>\t2\<^bold>\ \ \<^bold>\t3\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t3\<^bold>\ = \<^bold>\t3\<^bold>\ \ Arr t1 \ Arr t2 \ Arr t3 \ Src t1 = Trg t2 \ Src t2 = Trg t3 \ (\<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t3\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\Dom t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t3\<^bold>\) = \<^bold>\t1\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t2\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\t3\<^bold>\" using VcompNml_HcompNml Ide_Dom HcompNml_in_Hom Nml_HcompNml Nml_Nmlize Nmlize_in_Hom HcompNml_assoc by simp qed thus ?thesis using assms by blast qed lemma Nmlize_Inv: assumes "Can t" shows "\<^bold>\Inv t\<^bold>\ = Inv \<^bold>\t\<^bold>\" proof - have "Can t \ \<^bold>\Inv t\<^bold>\ = Inv \<^bold>\t\<^bold>\" proof (induct t, simp_all) fix u v assume I1: "\<^bold>\Inv u\<^bold>\ = Inv \<^bold>\u\<^bold>\" assume I2: "\<^bold>\Inv v\<^bold>\ = Inv \<^bold>\v\<^bold>\" show "Can u \ Can v \ Src u = Trg v \ Inv \<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Inv \<^bold>\v\<^bold>\ = Inv (\<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\)" using Inv_HcompNml Nml_Nmlize Can_implies_Arr Can_Nmlize_Can I1 I2 by simp show "Can u \ Can v \ Dom u = Cod v \ Inv \<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Inv \<^bold>\u\<^bold>\ = Inv (\<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\)" using Inv_VcompNml Nml_Nmlize Can_implies_Arr Nmlize_in_Hom Can_Nmlize_Can I1 I2 by simp fix w assume I3: "\<^bold>\Inv w\<^bold>\ = Inv \<^bold>\w\<^bold>\" assume uvw: "Can u \ Can v \ Can w \ Src u = Trg v \ Src v = Trg w" show "Inv \<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (Inv \<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Inv \<^bold>\w\<^bold>\) = Inv ((\<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\)" using uvw I1 I2 I3 Inv_HcompNml Nml_Nmlize Can_implies_Arr Can_Nmlize_Can Nml_HcompNml Can_HcompNml HcompNml_assoc by simp show "(Inv \<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Inv \<^bold>\v\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ Inv \<^bold>\w\<^bold>\ = Inv (\<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\))" using uvw I1 I2 I3 Inv_HcompNml Nml_Nmlize Can_implies_Arr Can_Nmlize_Can Nml_HcompNml Can_HcompNml HcompNml_assoc Can_Inv by simp qed thus ?thesis using assms by blast qed subsection "Reductions" text \ Function \red\ defined below takes a formal identity @{term t} to a canonical arrow \f\<^bold>\ \ Hom f \<^bold>\f\<^bold>\\. The auxiliary function \red2\ takes a pair @{term "(f, g)"} of normalized formal identities and produces a canonical arrow \f \<^bold>\ g \ Hom (f \<^bold>\ g) \<^bold>\f \<^bold>\ g\<^bold>\\. \ fun red2 (infixr "\<^bold>\" 53) where "\<^bold>\b\<^bold>\\<^sub>0 \<^bold>\ u = \<^bold>\\<^bold>[u\<^bold>]" | "\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\\<^sub>0 = \<^bold>\\<^bold>[\<^bold>\f\<^bold>\\<^bold>]" | "\<^bold>\f\<^bold>\ \<^bold>\ u = \<^bold>\f\<^bold>\ \<^bold>\ u" | "(t \<^bold>\ u) \<^bold>\ \<^bold>\a\<^bold>\\<^sub>0 = \<^bold>\\<^bold>[t \<^bold>\ u\<^bold>]" | "(t \<^bold>\ u) \<^bold>\ v = (t \<^bold>\ \<^bold>\u \<^bold>\ v\<^bold>\) \<^bold>\ (t \<^bold>\ (u \<^bold>\ v)) \<^bold>\ \<^bold>\\<^bold>[t, u, v\<^bold>]" | "t \<^bold>\ u = undefined" fun red ("_\<^bold>\" [56] 56) where "\<^bold>\f\<^bold>\\<^sub>0\<^bold>\ = \<^bold>\f\<^bold>\\<^sub>0" | "\<^bold>\f\<^bold>\\<^bold>\ = \<^bold>\f\<^bold>\" | "(t \<^bold>\ u)\<^bold>\ = (if Nml (t \<^bold>\ u) then t \<^bold>\ u else (\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\) \<^bold>\ (t\<^bold>\ \<^bold>\ u\<^bold>\))" | "t\<^bold>\ = undefined" lemma red_Nml [simp]: assumes "Nml a" shows "a\<^bold>\ = a" using assms by (cases a, simp_all) lemma red2_Nml: assumes "Nml (a \<^bold>\ b)" shows "a \<^bold>\ b = a \<^bold>\ b" proof - have a: "a = \<^bold>\un_Prim a\<^bold>\" using assms Nml_HcompD by metis have b: "Nml b \ \ is_Prim\<^sub>0 b" using assms Nml_HcompD by metis show ?thesis using a b apply (cases b) apply simp_all apply (metis red2.simps(3)) by (metis red2.simps(4)) qed lemma Can_red2: assumes "Ide a" and "Nml a" and "Ide b" and "Nml b" and "Src a = Trg b" shows "Can (a \<^bold>\ b)" and "a \<^bold>\ b \ VHom (a \<^bold>\ b) \<^bold>\a \<^bold>\ b\<^bold>\" proof - have 0: "\b. \ Ide a \ Nml a; Ide b \ Nml b; Src a = Trg b \ \ Can (a \<^bold>\ b) \ a \<^bold>\ b \ VHom (a \<^bold>\ b) \<^bold>\a \<^bold>\ b\<^bold>\" proof (induct a, simp_all add: HcompNml_Nml_Src HcompNml_Trg_Nml) fix x b show "Ide b \ Nml b \ Can (Trg b \<^bold>\ b) \ Arr (Trg b \<^bold>\ b) \ Dom (Trg b \<^bold>\ b) = Trg b \<^bold>\ b \ Cod (Trg b \<^bold>\ b) = b" using Ide_implies_Can Ide_in_Hom Nmlize_Nml apply (cases b, simp_all) proof - fix u v assume uv: "Ide u \ Ide v \ Src u = Trg v \ Nml (u \<^bold>\ v)" show "Can (Trg u \<^bold>\ (u \<^bold>\ v)) \ Arr (Trg u \<^bold>\ (u \<^bold>\ v)) \ Dom (Trg u \<^bold>\ (u \<^bold>\ v)) = Trg u \<^bold>\ u \<^bold>\ v \ Cod (Trg u \<^bold>\ (u \<^bold>\ v)) = u \<^bold>\ v" using uv Ide_implies_Can Can_implies_Arr Ide_in_Hom by (cases u, simp_all) qed show "ide x \ arr x \ Ide b \ Nml b \ \<^bold>\src x\<^bold>\\<^sub>0 = Trg b \ Can (\<^bold>\x\<^bold>\ \<^bold>\ b) \ Arr (\<^bold>\x\<^bold>\ \<^bold>\ b) \ Dom (\<^bold>\x\<^bold>\ \<^bold>\ b) = \<^bold>\x\<^bold>\ \<^bold>\ b \ Cod (\<^bold>\x\<^bold>\ \<^bold>\ b) = \<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ b" using Ide_implies_Can Can_implies_Arr Nmlize_Nml Ide_in_Hom by (cases b, simp_all) next fix u v w assume uv: "Ide u \ Ide v \ Src u = Trg v \ Nml (u \<^bold>\ v)" assume vw: "Src v = Trg w" assume w: "Ide w \ Nml w" assume I1: "\w. \ Nml u; Ide w \ Nml w; Trg v = Trg w \ \ Can (u \<^bold>\ w) \ Arr (u \<^bold>\ w) \ Dom (u \<^bold>\ w) = u \<^bold>\ w \ Cod (u \<^bold>\ w) = u \<^bold>\\<^bold>\\<^bold>\ w" assume I2: "\x. \ Nml v; Ide x \ Nml x; Trg w = Trg x \ \ Can (v \<^bold>\ x) \ Arr (v \<^bold>\ x) \ Dom (v \<^bold>\ x) = v \<^bold>\ x \ Cod (v \<^bold>\ x) = v \<^bold>\\<^bold>\\<^bold>\ x" show "Can ((u \<^bold>\ v) \<^bold>\ w) \ Arr ((u \<^bold>\ v) \<^bold>\ w) \ Dom ((u \<^bold>\ v) \<^bold>\ w) = (u \<^bold>\ v) \<^bold>\ w \ Cod ((u \<^bold>\ v) \<^bold>\ w) = (\<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ w" proof - have u: "Nml u \ Ide u" using uv Nml_HcompD by blast have v: "Nml v \ Ide v" using uv Nml_HcompD by blast have "is_Prim\<^sub>0 w \ ?thesis" proof - assume 1: "is_Prim\<^sub>0 w" have 2: "(u \<^bold>\ v) \<^bold>\ w = \<^bold>\\<^bold>[u \<^bold>\ v\<^bold>]" using 1 by (cases w, simp_all) have 3: "Can (u \<^bold>\ v) \ Arr (u \<^bold>\ v) \ Dom (u \<^bold>\ v) = u \<^bold>\ v \ Cod (u \<^bold>\ v) = u \<^bold>\ v" using u v uv 1 2 I1 Nmlize_Nml Nmlize.simps(3) by metis hence 4: "VSeq (u \<^bold>\ v) \<^bold>\\<^bold>[u \<^bold>\ v\<^bold>]" using uv by (metis (mono_tags, lifting) Arr.simps(7) Cod.simps(3) Cod.simps(7) Nml_implies_Arr Ide_in_Hom(2) mem_Collect_eq) have "Can ((u \<^bold>\ v) \<^bold>\ w)" using 1 2 3 4 uv by (simp add: Ide_implies_Can) moreover have "Dom ((u \<^bold>\ v) \<^bold>\ w) = (u \<^bold>\ v) \<^bold>\ w" using 1 2 3 4 u v w uv vw I1 Ide_in_Hom Nml_HcompNml Ide_in_Hom apply (cases w) by simp_all moreover have "Cod ((u \<^bold>\ v) \<^bold>\ w) = \<^bold>\(u \<^bold>\ v) \<^bold>\ w\<^bold>\" using 1 2 3 4 uv using Nmlize_Nml apply (cases w, simp_all) by (metis Nmlize.simps(3) Nmlize_Nml HcompNml.simps(3)) ultimately show ?thesis using w Can_implies_Arr by (simp add: 1 uv) qed moreover have "\ is_Prim\<^sub>0 w \ ?thesis" proof - assume 1: "\ is_Prim\<^sub>0 w" have 2: "(u \<^bold>\ v) \<^bold>\ w = (u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\) \<^bold>\ (u \<^bold>\ v \<^bold>\ w) \<^bold>\ \<^bold>\\<^bold>[u, v, w\<^bold>]" using 1 u v uv w by (cases w; simp) have 3: "Can (u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\) \ Dom (u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\) = u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\ \ Cod (u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\) = \<^bold>\u \<^bold>\ (v \<^bold>\ w)\<^bold>\" proof - have "Can (u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\) \ Dom (u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\) = u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\ \ Cod (u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\) = u \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\" using w uv Ide_HcompNml Nml_HcompNml(1) apply (cases u, simp_all) using u v vw I1 Nmlize_in_Hom(1) [of "v \<^bold>\ w"] Nml_Nmlize Ide_Nmlize_Ide by simp moreover have "u \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\ = \<^bold>\u \<^bold>\ (v \<^bold>\ w)\<^bold>\" using uv u w Nmlize_Hcomp Nmlize_Nmlize Nml_implies_Arr by simp ultimately show ?thesis by presburger qed have 4: "Can (v \<^bold>\ w) \ Dom (v \<^bold>\ w) = v \<^bold>\ w \ Cod (v \<^bold>\ w) = \<^bold>\v \<^bold>\ w\<^bold>\" using v w vw 1 2 I2 by simp hence 5: "Src (v \<^bold>\ w) = Src w \ Trg (v \<^bold>\ w) = Trg v" using Src_Dom Trg_Dom Can_implies_Arr by fastforce have "Can (u \<^bold>\ (v \<^bold>\ w)) \ Dom (u \<^bold>\ (v \<^bold>\ w)) = u \<^bold>\ (v \<^bold>\ w) \ Cod (u \<^bold>\ (v \<^bold>\ w)) = u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\" using u uv vw 4 5 Ide_implies_Can Ide_in_Hom by simp moreover have "\<^bold>\u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\\<^bold>\ = \<^bold>\u \<^bold>\ v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\" proof - have "\<^bold>\u \<^bold>\ \<^bold>\v \<^bold>\ w\<^bold>\\<^bold>\ = u \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w)" using u v w 4 by (metis Ide_Dom Can_implies_Arr Ide_implies_Arr Nml_Nmlize(1) Nmlize.simps(3) Nmlize_Nml) also have "... = (u \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ w" using u v w uv vw HcompNml_assoc by metis also have "... = \<^bold>\u \<^bold>\ v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\" using u v w by (metis Nmlize.simps(3) Nmlize_Nml) finally show ?thesis by blast qed moreover have "Can \<^bold>\\<^bold>[u, v, w\<^bold>] \ Dom \<^bold>\\<^bold>[u, v, w\<^bold>] = (u \<^bold>\ v) \<^bold>\ w \ Cod \<^bold>\\<^bold>[u, v, w\<^bold>] = u \<^bold>\ (v \<^bold>\ w)" using uv vw w Ide_implies_Can Ide_in_Hom by auto ultimately show ?thesis using uv w 2 3 4 Nml_implies_Arr Nmlize_Nmlize Ide_implies_Can Nmlize_Nml Ide_Dom Can_implies_Arr by (metis Can.simps(4) Cod.simps(4) Dom.simps(4) Nmlize.simps(3)) qed ultimately show ?thesis by blast qed qed show "Can (a \<^bold>\ b)" using assms 0 by blast show "a \<^bold>\ b \ VHom (a \<^bold>\ b) \<^bold>\a \<^bold>\ b\<^bold>\" using 0 assms by blast qed lemma red2_in_Hom [intro]: assumes "Ide u" and "Nml u" and "Ide v" and "Nml v" and "Src u = Trg v" shows "u \<^bold>\ v \ HHom (Src v) (Trg u)" and "u \<^bold>\ v \ VHom (u \<^bold>\ v) \<^bold>\u \<^bold>\ v\<^bold>\" proof - show 1: "u \<^bold>\ v \ VHom (u \<^bold>\ v) \<^bold>\u \<^bold>\ v\<^bold>\" using assms Can_red2 Can_implies_Arr by simp show "u \<^bold>\ v \ HHom (Src v) (Trg u)" using assms 1 Src_Dom [of "u \<^bold>\ v"] Trg_Dom [of "u \<^bold>\ v"] Can_red2 Can_implies_Arr by simp qed lemma red2_simps [simp]: assumes "Ide u" and "Nml u" and "Ide v" and "Nml v" and "Src u = Trg v" shows "Src (u \<^bold>\ v) = Src v" and "Trg (u \<^bold>\ v) = Trg u" and "Dom (u \<^bold>\ v) = u \<^bold>\ v" and "Cod (u \<^bold>\ v) = \<^bold>\u \<^bold>\ v\<^bold>\" using assms red2_in_Hom by auto lemma Can_red: assumes "Ide u" shows "Can (u\<^bold>\)" and "u\<^bold>\ \ VHom u \<^bold>\u\<^bold>\" proof - have 0: "Ide u \ Can (u\<^bold>\) \ u\<^bold>\ \ VHom u \<^bold>\u\<^bold>\" proof (induct u, simp_all add: Dom_Ide Cod_Ide) fix v w assume v: "Can (v\<^bold>\) \ Arr (v\<^bold>\) \ Dom (v\<^bold>\) = v \ Cod (v\<^bold>\) = \<^bold>\v\<^bold>\" assume w: "Can (w\<^bold>\) \ Arr (w\<^bold>\) \ Dom (w\<^bold>\) = w \ Cod (w\<^bold>\) = \<^bold>\w\<^bold>\" assume vw: "Ide v \ Ide w \ Src v = Trg w" show "(Nml (v \<^bold>\ w) \ Can v \ Can w \ v \<^bold>\ w = \<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\) \ (\ Nml (v \<^bold>\ w) \ Can (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) \ Src (v\<^bold>\) = Trg (w\<^bold>\) \ Dom (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) = \<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\ \ Arr (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) \ Src (v\<^bold>\) = Trg (w\<^bold>\) \ Dom (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) = \<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\ \ Cod (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) = \<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\)" proof show "Nml (v \<^bold>\ w) \ Can v \ Can w \ v \<^bold>\ w = \<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\" using vw Nml_HcompD Ide_implies_Can Dom_Inv VcompNml_Ide_Nml Inv_Ide Nmlize.simps(3) Nmlize_Nml by metis show "\ Nml (v \<^bold>\ w) \ Can (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) \ Src (v\<^bold>\) = Trg (w\<^bold>\) \ Dom (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) = \<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\ \ Arr (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) \ Src (v\<^bold>\) = Trg (w\<^bold>\) \ Dom (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) = \<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\ \ Cod (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) = \<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\" proof assume 1: "\ Nml (v \<^bold>\ w)" have "Can (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\)" using v w vw Can_red2 Nml_Nmlize Ide_Nmlize_Ide Nml_HcompNml Ide_HcompNml by simp moreover have "Src (v\<^bold>\) = Trg (w\<^bold>\)" using v w vw Src_Dom Trg_Dom by metis moreover have "Dom (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) = \<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\ \ Cod (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) = \<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\" using v w vw Can_red2 Nml_Nmlize Ide_Nmlize_Ide by simp ultimately show "Can (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) \ Src (v\<^bold>\) = Trg (w\<^bold>\) \ Dom (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) = \<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\ \ Arr (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) \ Src (v\<^bold>\) = Trg (w\<^bold>\) \ Dom (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) = \<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\ \ Cod (\<^bold>\v\<^bold>\ \<^bold>\ \<^bold>\w\<^bold>\) = \<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\" using Can_implies_Arr by blast qed qed qed show "Can (u\<^bold>\)" using assms 0 by blast show "u\<^bold>\ \ VHom u \<^bold>\u\<^bold>\" using assms 0 by blast qed lemma red_in_Hom [intro]: assumes "Ide t" shows "t\<^bold>\ \ HHom (Src t) (Trg t)" and "t\<^bold>\ \ VHom t \<^bold>\t\<^bold>\" proof - show 1: "t\<^bold>\ \ VHom t \<^bold>\t\<^bold>\" using assms Can_red Can_implies_Arr by simp show "t\<^bold>\ \ HHom (Src t) (Trg t)" using assms 1 Src_Dom [of "t\<^bold>\"] Trg_Dom [of "t\<^bold>\"] Can_red Can_implies_Arr by simp qed lemma red_simps [simp]: assumes "Ide t" shows "Src (t\<^bold>\) = Src t" and "Trg (t\<^bold>\) = Trg t" and "Dom (t\<^bold>\) = t" and "Cod (t\<^bold>\) = \<^bold>\t\<^bold>\" using assms red_in_Hom by auto lemma red_Src: assumes "Ide t" shows "Src t\<^bold>\ = Src t" using assms is_Prim0_Src [of t] by (cases "Src t", simp_all) lemma red_Trg: assumes "Ide t" shows "Trg t\<^bold>\ = Trg t" using assms is_Prim0_Trg [of t] by (cases "Trg t", simp_all) lemma Nmlize_red [simp]: assumes "Ide t" shows "\<^bold>\t\<^bold>\\<^bold>\ = \<^bold>\t\<^bold>\" using assms Can_red Ide_Nmlize_Can Nmlize_in_Hom Ide_in_Hom by fastforce lemma Nmlize_red2 [simp]: assumes "Ide t" and "Ide u" and "Nml t" and "Nml u" and "Src t = Trg u" shows "\<^bold>\t \<^bold>\ u\<^bold>\ = \<^bold>\t \<^bold>\ u\<^bold>\" using assms Can_red2 Ide_Nmlize_Can Nmlize_in_Hom [of "t \<^bold>\ u"] red2_in_Hom Ide_in_Hom by simp end subsection "Evaluation" text \ The following locale is concerned with the evaluation of terms of the bicategorical language determined by \C\, \src\<^sub>C\, and \trg\<^sub>C\ in a bicategory \(V, H, \, \, src, trg)\, given a source and target-preserving functor from \C\ to \V\. \ locale evaluation_map = C: horizontal_homs C src\<^sub>C trg\<^sub>C + bicategorical_language C src\<^sub>C trg\<^sub>C + bicategory V H \ \ src trg + E: "functor" C V E for C :: "'c comp" (infixr "\\<^sub>C" 55) and src\<^sub>C :: "'c \ 'c" and trg\<^sub>C :: "'c \ 'c" and V :: "'b comp" (infixr "\" 55) and H :: "'b comp" (infixr "\" 53) and \ :: "'b \ 'b \ 'b \ 'b" ("\[_, _, _]") and \ :: "'b \ 'b" ("\[_]") and src :: "'b \ 'b" and trg :: "'b \ 'b" and E :: "'c \ 'b" + assumes preserves_src: "E (src\<^sub>C x) = src (E x)" and preserves_trg: "E (trg\<^sub>C x) = trg (E x)" begin (* TODO: Figure out why this notation has to be reinstated. *) notation Nmlize ("\<^bold>\_\<^bold>\") notation HcompNml (infixr "\<^bold>\\<^bold>\\<^bold>\" 53) notation VcompNml (infixr "\<^bold>\\<^bold>\\<^bold>\" 55) notation red ("_\<^bold>\" [56] 56) notation red2 (infixr "\<^bold>\" 53) primrec eval :: "'c term \ 'b" ("\_\") where "\\<^bold>\f\<^bold>\\<^sub>0\ = E f" | "\\<^bold>\f\<^bold>\\ = E f" | "\t \<^bold>\ u\ = \t\ \ \u\" | "\t \<^bold>\ u\ = \t\ \ \u\" | "\\<^bold>\\<^bold>[t\<^bold>]\ = \ \t\" | "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \'.map \t\" | "\\<^bold>\\<^bold>[t\<^bold>]\ = \ \t\" | "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \'.map \t\" | "\\<^bold>\\<^bold>[t, u, v\<^bold>]\ = \ (\t\, \u\, \v\)" | "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = \'.map (\t\, \u\, \v\)" lemma preserves_obj: assumes "C.obj a" shows "obj (E a)" proof (unfold obj_def) show "arr (E a) \ src (E a) = E a" proof show "arr (E a)" using assms C.obj_simps by simp have "src (E a) = E (src\<^sub>C a)" using assms preserves_src by metis also have "... = E a" using assms C.obj_simps by simp finally show "src (E a) = E a" by simp qed qed lemma eval_in_hom': shows "Arr t \ \\t\ : \Src t\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\" proof (induct t) show "\x. Arr \<^bold>\x\<^bold>\\<^sub>0 \ \\\<^bold>\x\<^bold>\\<^sub>0\ : \Src \<^bold>\x\<^bold>\\<^sub>0\ \ \Trg \<^bold>\x\<^bold>\\<^sub>0\\ \ \\\<^bold>\x\<^bold>\\<^sub>0\ : \Dom \<^bold>\x\<^bold>\\<^sub>0\ \ \Cod \<^bold>\x\<^bold>\\<^sub>0\\" apply (simp add: preserves_src preserves_trg) using preserves_src preserves_trg C.objE by (metis (full_types) C.obj_def' E.preserves_arr E.preserves_ide in_hhom_def ide_in_hom(2)) show "\x. Arr \<^bold>\x\<^bold>\ \ \\\<^bold>\x\<^bold>\\ : \Src \<^bold>\x\<^bold>\\ \ \Trg \<^bold>\x\<^bold>\\\ \ \\\<^bold>\x\<^bold>\\ : \Dom \<^bold>\x\<^bold>\\ \ \Cod \<^bold>\x\<^bold>\\\" by (auto simp add: preserves_src preserves_trg) show "\t1 t2. (Arr t1 \ \\t1\ : \Src t1\ \ \Trg t1\\ \ \\t1\ : \Dom t1\ \ \Cod t1\\) \ (Arr t2 \ \\t2\ : \Src t2\ \ \Trg t2\\ \ \\t2\ : \Dom t2\ \ \Cod t2\\) \ Arr (t1 \<^bold>\ t2) \ \\t1 \<^bold>\ t2\ : \Src (t1 \<^bold>\ t2)\ \ \Trg (t1 \<^bold>\ t2)\\ \ \\t1 \<^bold>\ t2\ : \Dom (t1 \<^bold>\ t2)\ \ \Cod (t1 \<^bold>\ t2)\\" using hcomp_in_hhom in_hhom_def vconn_implies_hpar(1) vconn_implies_hpar(2) by auto show "\t1 t2. (Arr t1 \ \\t1\ : \Src t1\ \ \Trg t1\\ \ \\t1\ : \Dom t1\ \ \Cod t1\\) \ (Arr t2 \ \\t2\ : \Src t2\ \ \Trg t2\\ \ \\t2\ : \Dom t2\ \ \Cod t2\\) \ Arr (t1 \<^bold>\ t2) \ \\t1 \<^bold>\ t2\ : \Src (t1 \<^bold>\ t2)\ \ \Trg (t1 \<^bold>\ t2)\\ \ \\t1 \<^bold>\ t2\ : \Dom (t1 \<^bold>\ t2)\ \ \Cod (t1 \<^bold>\ t2)\\" using VSeq_implies_HPar seqI' by auto show "\t. (Arr t \ \\t\ : \Src t\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\) \ Arr \<^bold>\\<^bold>[t\<^bold>] \ \\\<^bold>\\<^bold>[t\<^bold>]\ : \Src \<^bold>\\<^bold>[t\<^bold>]\ \ \Trg \<^bold>\\<^bold>[t\<^bold>]\\ \ \\\<^bold>\\<^bold>[t\<^bold>]\ : \Dom \<^bold>\\<^bold>[t\<^bold>]\ \ \Cod \<^bold>\\<^bold>[t\<^bold>]\\" proof (simp add: preserves_src preserves_trg) fix t assume t: "\\t\ : \Src t\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\" assume 1: "Arr t" show "\\ \t\ : \Src t\ \ \Trg t\\ \ \\ \t\ : \Trg t\ \ \Dom t\ \ \Cod t\\" proof - have "src (\ \t\) = \Src t\" using t 1 by (metis (no_types, lifting) \.preserves_cod arr_cod in_hhomE map_simp src_cod) moreover have "trg (\ \t\) = \Trg t\" using t 1 by (metis (no_types, lifting) \.preserves_cod arr_cod in_hhomE map_simp trg_cod) moreover have "\\ \t\ : \Trg t\ \ \Dom t\ \ \Cod t\\" using t 1 apply (elim conjE in_hhomE) by (intro in_homI, auto) ultimately show ?thesis by auto qed qed show "\t. (Arr t \ \\t\ : \Src t\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\) \ Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ \\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ : \Src \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ \ \Trg \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\\ \ \\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ : \Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ \ \Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\\" proof (simp add: preserves_src preserves_trg) fix t assume t: "\\t\ : \Src t\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\" assume 1: "Arr t" show "\\'.map \t\ : \Src t\ \ \Trg t\\ \ \\'.map \t\ : \Dom t\ \ \Trg t\ \ \Cod t\\" proof - have "src (\'.map \t\) = \Src t\" using t 1 \'.preserves_dom arr_dom map_simp \'.preserves_reflects_arr src_dom by (metis (no_types, lifting) in_hhomE) moreover have "trg (\'.map \t\) = \Trg t\" using t 1 \'.preserves_dom arr_dom map_simp \'.preserves_reflects_arr trg_dom by (metis (no_types, lifting) in_hhomE) moreover have "\\'.map \t\ : \Dom t\ \ \Trg t\ \ \Cod t\\" using t 1 \'.preserves_hom apply (intro in_homI) apply auto[1] apply fastforce by fastforce ultimately show ?thesis by blast qed qed show "\t. (Arr t \ \\t\ : \Src t\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\) \ Arr \<^bold>\\<^bold>[t\<^bold>] \ \\\<^bold>\\<^bold>[t\<^bold>]\ : \Src \<^bold>\\<^bold>[t\<^bold>]\ \ \Trg \<^bold>\\<^bold>[t\<^bold>]\\ \ \\\<^bold>\\<^bold>[t\<^bold>]\ : \Dom \<^bold>\\<^bold>[t\<^bold>]\ \ \Cod \<^bold>\\<^bold>[t\<^bold>]\\" proof (simp add: preserves_src preserves_trg) fix t assume t: "\\t\ : \Src t\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\" assume 1: "Arr t" show "\\ \t\ : \Src t\ \ \Trg t\\ \ \\ \t\ : \Dom t\ \ \Src t\ \ \Cod t\\" proof - have "src (\ \t\) = \Src t\" using t 1 \.preserves_cod arr_cod map_simp \.preserves_reflects_arr src_cod by (metis (no_types, lifting) in_hhomE) moreover have "trg (\ \t\) = \Trg t\" using t 1 \.preserves_cod arr_cod map_simp \.preserves_reflects_arr trg_cod by (metis (no_types, lifting) in_hhomE) moreover have "\\ \t\ : \Dom t\ \ \Src t\ \ \Cod t\\" using t 1 by force ultimately show ?thesis by blast qed qed show "\t. (Arr t \ \\t\ : \Src t\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\) \ Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ \\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ : \Src \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ \ \Trg \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\\ \ \\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ : \Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ \ \Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\\" proof (simp add: preserves_src preserves_trg) fix t assume t: "\\t\ : \Src t\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\" assume 1: "Arr t" show "\\'.map \t\ : \Src t\ \ \Trg t\\ \ \\'.map \t\ : \Dom t\ \ \Cod t\ \ \Src t\\" proof - have "src (\'.map \t\) = \Src t\" using t 1 \'.preserves_dom arr_dom map_simp \'.preserves_reflects_arr src_dom by (metis (no_types, lifting) in_hhomE) moreover have "trg (\'.map \t\) = \Trg t\" using t 1 \'.preserves_dom arr_dom map_simp \'.preserves_reflects_arr trg_dom by (metis (no_types, lifting) in_hhomE) moreover have "\\'.map \t\ : \Dom t\ \ \Cod t\ \ \Src t\\" using t 1 src_cod arr_cod \'.preserves_hom [of "\t\" "\Dom t\" "\Cod t\"] apply (elim conjE in_hhomE) by (intro in_homI, auto) ultimately show ?thesis by blast qed qed show "\t u v. (Arr t \ \\t\ : \Src t\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\) \ (Arr u \ \\u\ : \Src u\ \ \Trg u\\ \ \\u\ : \Dom u\ \ \Cod u\\) \ (Arr v \ \\v\ : \Src v\ \ \Trg v\\ \ \\v\ : \Dom v\ \ \Cod v\\) \ Arr \<^bold>\\<^bold>[t, u, v\<^bold>] \ \\\<^bold>\\<^bold>[t, u, v\<^bold>]\ : \Src \<^bold>\\<^bold>[t, u, v\<^bold>]\ \ \Trg \<^bold>\\<^bold>[t, u, v\<^bold>]\\ \ \\\<^bold>\\<^bold>[t, u, v\<^bold>]\ : \Dom \<^bold>\\<^bold>[t, u, v\<^bold>]\ \ \Cod \<^bold>\\<^bold>[t, u, v\<^bold>]\\" proof (simp add: preserves_src preserves_trg) fix t u v assume t: "\\t\ : \Trg u\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\" assume u: "\\u\ : \Trg v\ \ \Trg u\\ \ \\u\ : \Dom u\ \ \Cod u\\" assume v: "\\v\ : \Src v\ \ \Trg v\\ \ \\v\ : \Dom v\ \ \Cod v\\" assume tuv: "Arr t \ Arr u \ Arr v \ Src t = Trg u \ Src u = Trg v" show "\\ (\t\, \u\, \v\) : \Src v\ \ \Trg t\\ \ \\ (\t\, \u\, \v\) : (\Dom t\ \ \Dom u\) \ \Dom v\ \ \Cod t\ \ \Cod u\ \ \Cod v\\" proof - have 1: "VVV.in_hom (\t\, \u\, \v\) (\Dom t\, \Dom u\, \Dom v\) (\Cod t\, \Cod u\, \Cod v\)" proof - have "(\t\, \u\, \v\) \ VxVxV.hom (\Dom t\, \Dom u\, \Dom v\) (\Cod t\, \Cod u\, \Cod v\)" using t u v tuv by simp moreover have "(\t\, \u\, \v\) \ {\\\. arr (fst \\\) \ VV.arr (snd \\\) \ src (fst \\\) = trg (fst (snd \\\))}" using t u v tuv by fastforce ultimately show ?thesis using VVV.hom_char [of "(\Dom t\, \Dom u\, \Dom v\)" "(\Cod t\, \Cod u\, \Cod v\)"] by blast qed have 4: "VVV.arr (\Dom t\, \Dom u\, \Dom v\)" using 1 VVV.ide_dom VVV.dom_simp by (elim VVV.in_homE) force have 5: "VVV.arr (\Cod t\, \Cod u\, \Cod v\)" using 1 VVV.ide_cod VVV.cod_simp by (elim VVV.in_homE) force have 2: "\\ (\t\, \u\, \v\) : (\Dom t\ \ \Dom u\) \ \Dom v\ \ \Cod t\ \ \Cod u\ \ \Cod v\\" using 1 4 5 HoHV_def HoVH_def \_def \.preserves_hom [of "(\t\, \u\, \v\)" "(\Dom t\, \Dom u\, \Dom v\)" "(\Cod t\, \Cod u\, \Cod v\)"] by simp have 3: "\\ (\t\, \u\, \v\) : \Src v\ \ \Trg t\\" proof show "arr (\ (\t\, \u\, \v\))" using 2 by auto show "src (\ (\t\, \u\, \v\)) = \Src v\" proof - have "src (\ (\t\, \u\, \v\)) = src ((\Dom t\ \ \Dom u\) \ \Dom v\)" using 2 src_dom [of "\ (\t\, \u\, \v\)"] by fastforce also have "... = src \Dom v\" using 4 VVV.arr_char VV.arr_char by simp also have "... = src (dom \v\)" using v by fastforce also have "... = \Src v\" using v by auto finally show ?thesis by auto qed show "trg (\ (\t\, \u\, \v\)) = \Trg t\" proof - have "trg (\ (\t\, \u\, \v\)) = trg ((\Dom t\ \ \Dom u\) \ \Dom v\)" using 2 trg_dom [of "\ (\t\, \u\, \v\)"] by fastforce also have "... = trg \Dom t\" using 4 VVV.arr_char VV.arr_char by simp also have "... = trg (dom \t\)" using t by fastforce also have "... = \Trg t\" using t by auto finally show ?thesis by auto qed qed show ?thesis using 2 3 by simp qed qed show "\t u v. (Arr t \ \\t\ : \Src t\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\) \ (Arr u \ \\u\ : \Src u\ \ \Trg u\\ \ \\u\ : \Dom u\ \ \Cod u\\) \ (Arr v \ \\v\ : \Src v\ \ \Trg v\\ \ \\v\ : \Dom v\ \ \Cod v\\) \ Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] \ \\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ : \Src \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ \ \Trg \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\\ \ \\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ : \Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ \ \Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\\" proof (simp add: preserves_src preserves_trg) fix t u v assume t: "\\t\ : \Trg u\ \ \Trg t\\ \ \\t\ : \Dom t\ \ \Cod t\\" assume u: "\\u\ : \Trg v\ \ \Trg u\\ \ \\u\ : \Dom u\ \ \Cod u\\" assume v: "\\v\ : \Src v\ \ \Trg v\\ \ \\v\ : \Dom v\ \ \Cod v\\" assume tuv: "Arr t \ Arr u \ Arr v \ Src t = Trg u \ Src u = Trg v" show "\\'.map (\t\, \u\, \v\) : \Src v\ \ \Trg t\\ \ \\'.map (\t\, \u\, \v\) : \Dom t\ \ \Dom u\ \ \Dom v\ \ (\Cod t\ \ \Cod u\) \ \Cod v\\" proof - have 1: "VVV.in_hom (\t\, \u\, \v\) (\Dom t\, \Dom u\, \Dom v\) (\Cod t\, \Cod u\, \Cod v\)" using t u v tuv VVV.hom_char VVV.arr_char VV.arr_char VVV.dom_char VVV.cod_char apply (elim conjE in_hhomE in_homE) apply (intro VVV.in_homI) by simp_all have 4: "VVV.arr (\Dom t\, \Dom u\, \Dom v\)" using "1" VVV.in_hom_char by blast have 5: "VVV.arr (\Cod t\, \Cod u\, \Cod v\)" using "1" VVV.in_hom_char by blast have 2: "\\'.map (\t\, \u\, \v\) : \Dom t\ \ \Dom u\ \ \Dom v\ \ (\Cod t\ \ \Cod u\) \ \Cod v\\" using 1 4 5 HoHV_def HoVH_def \'.map_def \'.preserves_hom [of "(\t\, \u\, \v\)" "(\Dom t\, \Dom u\, \Dom v\)" "(\Cod t\, \Cod u\, \Cod v\)"] by simp have 3: "\\'.map (\t\, \u\, \v\) : \Src v\ \ \Trg t\\" proof show "arr (\'.map (\t\, \u\, \v\))" using 2 by auto show "src (\'.map (\t\, \u\, \v\)) = \Src v\" proof - have "src (\'.map (\t\, \u\, \v\)) = src (\Dom t\ \ \Dom u\ \ \Dom v\)" using 2 src_dom [of "\'.map (\t\, \u\, \v\)"] by fastforce also have "... = src \Dom v\" using 4 VVV.arr_char VV.arr_char by simp also have "... = src (dom \v\)" using v by fastforce also have "... = \Src v\" using v by auto finally show ?thesis by auto qed show "trg (\'.map (\t\, \u\, \v\)) = \Trg t\" proof - have "trg (\'.map (\t\, \u\, \v\)) = trg (\Dom t\ \ \Dom u\ \ \Dom v\)" using 2 trg_dom [of "\'.map (\t\, \u\, \v\)"] by fastforce also have "... = trg \Dom t\" using 4 VVV.arr_char VV.arr_char hseqI' by simp also have "... = trg (dom \t\)" using t by fastforce also have "... = \Trg t\" using t by auto finally show ?thesis by auto qed qed show ?thesis using 2 3 by simp qed qed qed lemma eval_in_hom [intro]: assumes "Arr t" shows "\\t\ : \Src t\ \ \Trg t\\" and "\\t\ : \Dom t\ \ \Cod t\\" using assms eval_in_hom' by simp_all (* * TODO: It seems to me that the natural useful orientation of these facts is syntax * to semantics. However, having this as the default makes it impossible to do various * proofs by simp alone. This has to be sorted out. For right now, I am going to include * both versions, which will have to be explicitly invoked where needed. *) lemma eval_simps: assumes "Arr f" shows "arr \f\" and "\Src f\ = src \f\" and "\Trg f\ = trg \f\" and "\Dom f\ = dom \f\" and "\Cod f\ = cod \f\" using assms eval_in_hom [of f] by auto lemma eval_simps': assumes "Arr f" shows "arr \f\" and "src \f\ = \Src f\" and "trg \f\ = \Trg f\" and "dom \f\ = \Dom f\" and "cod \f\ = \Cod f\" using assms eval_in_hom by auto lemma obj_eval_Obj: shows "Obj t \ obj \t\" using preserves_obj by (induct t) auto lemma ide_eval_Ide: shows "Ide t \ ide \t\" by (induct t, auto simp add: eval_simps') lemma arr_eval_Arr [simp]: assumes "Arr t" shows "arr \t\" using assms by (simp add: eval_simps') (* * TODO: The next few results want eval_simps oriented from syntax to semantics. *) lemma eval_Lunit [simp]: assumes "Arr t" shows "\\<^bold>\\<^bold>[t\<^bold>]\ = \[\Cod t\] \ (trg \t\ \ \t\)" using assms \.is_natural_2 \_ide_simp by (simp add: eval_simps) lemma eval_Lunit' [simp]: assumes "Arr t" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \\<^sup>-\<^sup>1[\Cod t\] \ \t\" using assms \'.is_natural_2 \_ide_simp by (simp add: eval_simps) lemma eval_Runit [simp]: assumes "Arr t" shows "\\<^bold>\\<^bold>[t\<^bold>]\ = \[\Cod t\] \ (\t\ \ src \t\)" using assms \.is_natural_2 \_ide_simp by (simp add: eval_simps) lemma eval_Runit' [simp]: assumes "Arr t" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \\<^sup>-\<^sup>1[\Cod t\] \ \t\" using assms \'.is_natural_2 \_ide_simp by (simp add: eval_simps) lemma eval_Assoc [simp]: assumes "Arr t" and "Arr u" and "Arr v" and "Src t = Trg u" and "Src u = Trg v" shows "\\<^bold>\\<^bold>[t, u, v\<^bold>]\ = \ (cod \t\, cod \u\, cod \v\) \ ((\t\ \ \u\) \ \v\)" proof - have "\\<^bold>\\<^bold>[t, u, v\<^bold>]\ = \ (\t\, \u\, \v\)" by simp also have "... = \ (VVV.cod (\t\, \u\, \v\)) \ HoHV (\t\, \u\, \v\)" using assms \.is_natural_2 [of "(\t\, \u\, \v\)"] VVV.arr_char VVV.cod_char \.is_extensional \_def by auto also have "... = \ (cod \t\, cod \u\, cod \v\) \ ((\t\ \ \u\) \ \v\)" unfolding HoHV_def \_def using assms VVV.arr_char VV.arr_char VVV.cod_char \.is_extensional by auto finally show ?thesis by simp qed lemma eval_Assoc' [simp]: assumes "Arr t" and "Arr u" and "Arr v" and "Src t = Trg u" and "Src u = Trg v" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = \\<^sup>-\<^sup>1[cod \t\, cod \u\, cod \v\] \ (\t\ \ \u\ \ \v\)" proof - have "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = \'.map (\t\, \u\, \v\)" by simp also have "... = \'.map (VVV.cod (\t\, \u\, \v\)) \ HoVH (\t\, \u\, \v\)" using assms \'.is_natural_2 [of "(\t\, \u\, \v\)"] VVV.arr_char VVV.cod_char \'.is_extensional by simp also have "... = \'.map (cod \t\, cod \u\, cod \v\) \ (\t\ \ \u\ \ \v\)" unfolding HoVH_def using assms VVV.arr_char VV.arr_char VVV.cod_char \'.is_extensional apply simp using eval_simps'(2) eval_simps'(3) by presburger finally show ?thesis using \'_def by simp qed lemma eval_Lunit_Ide [simp]: assumes "Ide t" shows "\\<^bold>\\<^bold>[t\<^bold>]\ = \[\t\]" using assms \_ide_simp ide_eval_Ide by simp lemma eval_Lunit'_Ide [simp]: assumes "Ide t" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \\<^sup>-\<^sup>1[\t\]" using assms \_ide_simp ide_eval_Ide by simp lemma eval_Runit_Ide [simp]: assumes "Ide t" shows "\\<^bold>\\<^bold>[t\<^bold>]\ = \[\t\]" using assms \_ide_simp ide_eval_Ide by simp lemma eval_Runit'_Ide [simp]: assumes "Ide t" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \\<^sup>-\<^sup>1[\t\]" using assms \_ide_simp ide_eval_Ide by simp lemma eval_Assoc_Ide [simp]: assumes "Ide t" and "Ide u" and "Ide v" and "Src t = Trg u" and "Src u = Trg v" shows "\\<^bold>\\<^bold>[t, u, v\<^bold>]\ = \ (\t\, \u\, \v\)" using assms by simp lemma eval_Assoc'_Ide [simp]: assumes "Ide t" and "Ide u" and "Ide v" and "Src t = Trg u" and "Src u = Trg v" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = \\<^sup>-\<^sup>1[\t\, \u\, \v\]" using assms \'_def by simp lemma iso_eval_Can: shows "Can t \ iso \t\" proof (induct t; simp add: fsts.intros snds.intros) show "\x. C.obj x \ iso (E x)" by auto show "\t1 t2. \ iso \t1\; iso \t2\; Can t1 \ Can t2 \ Src t1 = Trg t2 \ \ iso (\t1\ \ \t2\)" using Can_implies_Arr by (simp add: eval_simps') show "\t1 t2. \ iso \t1\; iso \t2\; Can t1 \ Can t2 \ Dom t1 = Cod t2 \ \ iso (\t1\ \ \t2\)" using Can_implies_Arr isos_compose by (simp add: eval_simps') show "\t. \ iso \t\; Can t \ \ iso (\ \t\)" using \.preserves_iso by auto show "\t. \ iso \t\; Can t \ \ iso (\'.map \t\)" using \'.preserves_iso by simp show "\t. \ iso \t\; Can t \ \ iso (\ \t\)" using \.preserves_iso by auto show "\t. \ iso \t\; Can t \ \ iso (\'.map \t\)" using \'.preserves_iso by simp fix t1 t2 t3 assume t1: "iso \t1\" and t2: "iso \t2\" and t3: "iso \t3\" assume 1: "Can t1 \ Can t2 \ Can t3 \ Src t1 = Trg t2 \ Src t2 = Trg t3" have 2: "VVV.iso (\t1\, \t2\, \t3\)" proof - have 3: "VxVxV.iso (\t1\, \t2\, \t3\)" using t1 t2 t3 Can_implies_Arr VxVxV.iso_char VxV.iso_char by simp moreover have "VVV.arr (VxVxV.inv (\t1\, \t2\, \t3\))" proof - have "VxVxV.inv (\t1\, \t2\, \t3\) = (inv \t1\, inv \t2\, inv \t3\)" using t1 t2 t3 3 by simp thus ?thesis using t1 t2 t3 1 Can_implies_Arr VVV.arr_char VV.arr_char by (simp add: eval_simps') qed ultimately show ?thesis using t1 t2 t3 1 Can_implies_Arr VVV.iso_char VVV.arr_char by (auto simp add: eval_simps') qed show "iso (\ (\t1\, \t2\, \t3\))" using 2 \_def \.preserves_iso by auto show "iso (\'.map (\t1\, \t2\, \t3\))" using 2 \'.preserves_iso by simp qed lemma eval_Inv_Can: shows "Can t \ \Inv t\ = inv \t\" proof (induct t) show "\x. Can \<^bold>\x\<^bold>\\<^sub>0 \ \Inv \<^bold>\x\<^bold>\\<^sub>0\ = inv \\<^bold>\x\<^bold>\\<^sub>0\" by auto show "\x. Can \<^bold>\x\<^bold>\ \ \Inv \<^bold>\x\<^bold>\\ = inv \\<^bold>\x\<^bold>\\" by simp show "\t1 t2. (Can t1 \ \Inv t1\ = inv \t1\) \ (Can t2 \ \Inv t2\ = inv \t2\) \ Can (t1 \<^bold>\ t2) \ \Inv (t1 \<^bold>\ t2)\ = inv \t1 \<^bold>\ t2\" using iso_eval_Can Can_implies_Arr by (simp add: eval_simps') show "\t1 t2. (Can t1 \ \Inv t1\ = inv \t1\) \ (Can t2 \ \Inv t2\ = inv \t2\) \ Can (t1 \<^bold>\ t2) \ \Inv (t1 \<^bold>\ t2)\ = inv \t1 \<^bold>\ t2\" using iso_eval_Can inv_comp Can_implies_Arr by (simp add: eval_simps') fix t assume I: "Can t \ \Inv t\ = inv \t\" show "Can \<^bold>\\<^bold>[t\<^bold>] \ \Inv \<^bold>\\<^bold>[t\<^bold>]\ = inv \\<^bold>\\<^bold>[t\<^bold>]\" proof - assume t: "Can \<^bold>\\<^bold>[t\<^bold>]" have "\Inv \<^bold>\\<^bold>[t\<^bold>]\ = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[Inv t\<^bold>]\" by simp also have "... = \'.map (inv \t\)" using t I by simp also have "... = \'.map (cod (inv \t\)) \ inv \t\" using t \'.is_natural_2 iso_inv_iso iso_eval_Can iso_is_arr by (metis (no_types, lifting) Can.simps(5) map_simp) also have "... = inv (\t\ \ \ (dom \t\))" proof - have 1: "iso \t\" using t iso_eval_Can by simp moreover have "iso (\ (dom \t\))" using t 1 \.components_are_iso ide_dom by blast moreover have "seq \t\ (\ (dom \t\))" using t 1 iso_is_arr by auto ultimately show ?thesis using t 1 inv_comp by auto qed also have "... = inv \\<^bold>\\<^bold>[t\<^bold>]\" using t iso_eval_Can \_ide_simp lunit_naturality Can_implies_Arr eval_Lunit by (auto simp add: eval_simps) finally show ?thesis by blast qed show "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ \Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\" proof - assume t: "Can (Lunit' t)" have "\Inv (Lunit' t)\ = \Lunit (Inv t)\" by simp also have "... = \ (inv \t\)" using t I by simp also have "... = inv \t\ \ \ (dom (inv \t\))" using t \.is_natural_1 iso_inv_iso iso_eval_Can iso_is_arr by (metis (no_types, lifting) Can.simps(6) map_simp) also have "... = inv (\'.map (cod \t\) \ \t\)" proof - have 1: "iso \t\" using t iso_eval_Can by simp moreover have "iso (\'.map (cod \t\))" using t 1 \'.components_are_iso ide_cod by blast moreover have "seq (\'.map (cod \t\)) \t\" using t 1 iso_is_arr by auto ultimately show ?thesis using t 1 inv_comp by auto qed also have "... = inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\" using t \'.is_natural_2 iso_eval_Can iso_is_arr by force finally show ?thesis by auto qed show "Can \<^bold>\\<^bold>[t\<^bold>] \ \Inv \<^bold>\\<^bold>[t\<^bold>]\ = inv \\<^bold>\\<^bold>[t\<^bold>]\" proof - assume t: "Can \<^bold>\\<^bold>[t\<^bold>]" have "\Inv \<^bold>\\<^bold>[t\<^bold>]\ = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[Inv t\<^bold>]\" by simp also have "... = \'.map (inv \t\)" using t I by simp also have "... = \'.map (cod (inv \t\)) \ inv \t\" using t \'.is_natural_2 map_simp iso_inv_iso iso_eval_Can iso_is_arr by (metis (no_types, lifting) Can.simps(7)) also have "... = inv (\t\ \ \ (dom \t\))" proof - have 1: "iso \t\" using t iso_eval_Can by simp moreover have "iso (\ (dom \t\))" using t 1 \.components_are_iso ide_dom by blast moreover have "seq \t\ (\ (dom \t\))" using t 1 iso_is_arr by simp ultimately show ?thesis using t 1 inv_comp by auto qed also have "... = inv \\<^bold>\\<^bold>[t\<^bold>]\" using t \_ide_simp iso_eval_Can runit_naturality Can_implies_Arr eval_Runit by (auto simp add: eval_simps) finally show ?thesis by blast qed show "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ \Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\" proof - assume t: "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]" have "\Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \\<^bold>\\<^bold>[Inv t\<^bold>]\" by simp also have "... = \ (inv \t\)" using t I by simp also have "... = inv \t\ \ \ (dom (inv \t\))" using t \.is_natural_1 map_simp iso_inv_iso iso_eval_Can iso_is_arr by (metis (no_types, lifting) Can.simps(8)) also have "... = inv (\'.map (cod \t\) \ \t\)" proof - have 1: "iso \t\" using t iso_eval_Can by simp moreover have "iso (\'.map (cod \t\))" using t 1 \'.components_are_iso ide_cod by blast moreover have "seq (\'.map (cod \t\)) \t\" using t 1 iso_is_arr by auto ultimately show ?thesis using t 1 inv_comp by auto qed also have "... = inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\" using t \'.is_natural_2 iso_eval_Can iso_is_arr by auto finally show ?thesis by auto qed next fix t u v assume I1: "Can t \ \Inv t\ = inv \t\" assume I2: "Can u \ \Inv u\ = inv \u\" assume I3: "Can v \ \Inv v\ = inv \v\" show "Can \<^bold>\\<^bold>[t, u, v\<^bold>] \ \Inv \<^bold>\\<^bold>[t, u, v\<^bold>]\ = inv \\<^bold>\\<^bold>[t, u, v\<^bold>]\" proof - assume "Can \<^bold>\\<^bold>[t, u, v\<^bold>]" hence tuv: "Can t \ Can u \ Can v \ Src t = Trg u \ Src u = Trg v" by simp have "\Inv \<^bold>\\<^bold>[t, u, v\<^bold>]\ = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[Inv t, Inv u, Inv v\<^bold>]\" by simp also have "... = \\<^sup>-\<^sup>1[dom \t\, dom \u\, dom \v\] \ (inv \t\ \ inv \u\ \ inv \v\)" using tuv I1 I2 I3 eval_in_hom \'.map_ide_simp inv_in_hom iso_eval_Can assoc'_naturality Can_implies_Arr Src_Inv Trg_Inv eval_Assoc' Dom_Inv Can_Inv cod_inv by presburger also have "... = inv ((\t\ \ \u\ \ \v\) \ \ (dom \t\, dom \u\, dom \v\))" using tuv iso_eval_Can Can_implies_Arr eval_simps'(2) eval_simps'(3) \_def by (simp add: inv_comp) also have "... = inv (\ (\t\, \u\, \v\))" using tuv Can_implies_Arr \_def by (metis assoc_is_natural_1 arr_eval_Arr eval_simps'(2) eval_simps'(3) fst_conv snd_conv) also have "... = inv \\<^bold>\\<^bold>[t, u, v\<^bold>]\" by simp finally show ?thesis by blast qed show "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] \ \Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\" proof - assume tuv: "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]" have t: "Can t" using tuv by simp have u: "Can u" using tuv by simp have v: "Can v" using tuv by simp have "\Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = \\<^bold>\\<^bold>[Inv t, Inv u, Inv v\<^bold>]\" by simp also have "... = (inv \t\ \ inv \u\ \ inv \v\) \ \ (cod \t\, cod \u\, cod \v\)" using \_def tuv I1 I2 I3 iso_eval_Can Can_implies_Arr eval_simps'(2) eval_simps'(3) apply simp using assoc_is_natural_1 arr_inv dom_inv src_inv trg_inv by presburger also have "... = inv (\\<^sup>-\<^sup>1[cod \t\, cod \u\, cod \v\] \ (\t\ \ \u\ \ \v\))" using tuv inv_comp [of "\t\ \ \u\ \ \v\" "\\<^sup>-\<^sup>1[cod \t\, cod \u\, cod \v\]"] Can_implies_Arr iso_assoc \_def by (simp add: eval_simps'(2) eval_simps'(3) iso_eval_Can) also have 1: "... = inv (((\t\ \ \u\) \ \v\) \ \\<^sup>-\<^sup>1[dom \t\, dom \u\, dom \v\])" using tuv assoc'_naturality [of "\t\" "\u\" "\v\"] Can_implies_Arr eval_simps'(2) eval_simps'(3) by simp also have "... = inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\" using tuv 1 Can_implies_Arr eval_Assoc' by auto finally show ?thesis by blast qed qed lemma eval_VcompNml: assumes "Nml t" and "Nml u" and "VSeq t u" shows "\t \<^bold>\\<^bold>\\<^bold>\ u\ = \t\ \ \u\" proof - have "\u. \ Nml t; Nml u; VSeq t u \ \ \t \<^bold>\\<^bold>\\<^bold>\ u\ = \t\ \ \u\" proof (induct t, simp_all add: eval_simps) fix u a assume u: "Nml u" assume 1: "Arr u \ \<^bold>\a\<^bold>\\<^sub>0 = Cod u" show "\u\ = cod \u\ \ \u\" using 1 comp_cod_arr by simp next fix u f assume u: "Nml u" assume f: "C.arr f" assume 1: "Arr u \ \<^bold>\C.dom f\<^bold>\ = Cod u" show "\\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u\ = E f \ \u\" using f u 1 as_nat_trans.preserves_comp_2 by (cases u; simp) next fix u v w assume I1: "\u. \ Nml v; Nml u; Arr u \ Dom v = Cod u \ \ \v \<^bold>\\<^bold>\\<^bold>\ u\ = \v\ \ \u\" assume I2: "\u. \ Nml w; Nml u; Arr u \ Dom w = Cod u \ \ \w \<^bold>\\<^bold>\\<^bold>\ u\ = \w\ \ \u\" assume vw: "Nml (v \<^bold>\ w)" have v: "Nml v \ v = Prim (un_Prim v)" using vw by (simp add: Nml_HcompD) have w: "Nml w" using vw by (simp add: Nml_HcompD) assume u: "Nml u" assume 1: "Arr v \ Arr w \ Src v = Trg w \ Arr u \ Dom v \<^bold>\ Dom w = Cod u" show "\(v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u\ = (\v\ \ \w\) \ \u\" using u 1 HcompNml_in_Hom apply (cases u, simp_all) proof - fix x y assume 3: "u = x \<^bold>\ y" have x: "Nml x" using u 1 3 Nml_HcompD by simp have y: "Nml y" using u x 1 3 Nml_HcompD by simp assume 4: "Arr v \ Arr w \ Src v = Trg w \ Dom v = Cod x \ Dom w = Cod y" have "\v \<^bold>\\<^bold>\\<^bold>\ x\ \ \w \<^bold>\\<^bold>\\<^bold>\ y\ = \v \<^bold>\\<^bold>\\<^bold>\ x \<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ y\" using v w x y 4 HcompNml_in_Hom by simp moreover have "... = \v \<^bold>\\<^bold>\\<^bold>\ x\ \ \w \<^bold>\\<^bold>\\<^bold>\ y\" by simp moreover have "... = \v\ \ \x\ \ \w\ \ \y\" using v w x y 4 I1 [of x] I2 [of y] Nml_implies_Arr by simp moreover have "... = (\v\ \ \w\) \ (\x\ \ \y\)" using v w x y 4 Nml_implies_Arr interchange [of "\v\" "\x\" "\w\" "\y\"] by (simp add: eval_simps') ultimately have "\v \<^bold>\\<^bold>\\<^bold>\ x\ \ \w \<^bold>\\<^bold>\\<^bold>\ y\ = (\v\ \ \w\) \ (\x\ \ \y\)" by presburger moreover have "arr \v \<^bold>\\<^bold>\\<^bold>\ x\ \ arr \w \<^bold>\\<^bold>\\<^bold>\ y\" using v w x y 4 VcompNml_in_Hom by simp ultimately show "\v \<^bold>\\<^bold>\\<^bold>\ x\ \ \w \<^bold>\\<^bold>\\<^bold>\ y\ = (\v\ \ \w\) \ (\x\ \ \y\)" by simp qed qed thus ?thesis using assms by blast qed lemma eval_red_Hcomp: assumes "Ide a" and "Ide b" shows "\(a \<^bold>\ b)\<^bold>\\ = \\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\\ \ (\a\<^bold>\\ \ \b\<^bold>\\)" proof - have "Nml (a \<^bold>\ b) \ ?thesis" proof - assume 1: "Nml (a \<^bold>\ b)" hence 2: "Nml a \ Nml b \ Src a = Trg b" using Nml_HcompD(3-4,7) by simp have "\(a \<^bold>\ b)\<^bold>\\ = \a\ \ \b\" using 1 Nml_HcompD by simp also have "... = \\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\\ \ (\a\<^bold>\\ \ \b\<^bold>\\)" using assms 1 2 ide_eval_Ide Nmlize_in_Hom red2_Nml Nmlize_Nml by (simp add: eval_simps') finally show ?thesis by simp qed moreover have "\ Nml (a \<^bold>\ b) \ ?thesis" using assms Can_red2 by (simp add: Can_red(1) iso_eval_Can) ultimately show ?thesis by blast qed (* TODO: Would the following still be useful if \<^bold>\a\<^bold>\\<^sub>0 is replaced by Src t? *) lemma eval_red2_Nml_Prim\<^sub>0: assumes "Ide t" and "Nml t" and "Src t = \<^bold>\a\<^bold>\\<^sub>0" shows "\t \<^bold>\ \<^bold>\a\<^bold>\\<^sub>0\ = \[\t\]" using assms \_ide_simp apply (cases t) apply simp_all proof - show "C.obj a \ t = \<^bold>\a\<^bold>\\<^sub>0 \ \ (E a) = \[E a]" using unitor_coincidence obj_eval_Obj [of t] \_ide_simp by auto show "\b c. Ide b \ Ide c \ Src b = Trg c \ \ (\b\ \ \c\) = \[\b\ \ \c\]" using \_ide_simp by (simp add: eval_simps' ide_eval_Ide) qed end text \ Most of the time when we interpret the @{locale evaluation_map} locale, we are evaluating terms formed from the arrows in a bicategory as arrows of the bicategory itself. The following locale streamlines that use case. \ locale self_evaluation_map = bicategory begin sublocale bicategorical_language V src trg .. sublocale evaluation_map V src trg V H \ \ src trg \\\. if arr \ then \ else null\ using src.is_extensional trg.is_extensional by (unfold_locales, auto) notation eval ("\_\") notation Nmlize ("\<^bold>\_\<^bold>\") end subsection "Coherence" text \ We define an individual term to be \emph{coherent} if it commutes, up to evaluation, with the reductions of its domain and codomain. We then formulate the coherence theorem as the statement ``every formal arrow is coherent''. Because reductions evaluate to isomorphisms, this implies the standard version of coherence, which says that ``parallel canonical terms have equal evaluations''. \ context evaluation_map begin abbreviation coherent where "coherent t \ \Cod t\<^bold>\\ \ \t\ = \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" lemma Nml_implies_coherent: assumes "Nml t" shows "coherent t" using assms Nml_implies_Arr Ide_Dom Ide_Cod Nml_Dom Nml_Cod Nmlize_Nml red_Nml by (metis Dom_Cod VcompNml_Cod_Nml arr_eval_Arr comp_arr_dom eval_VcompNml eval_simps(4)) lemma canonical_factorization: assumes "Arr t" shows "coherent t \ \t\ = inv \Cod t\<^bold>\\ \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" proof assume 1: "coherent t" have "inv \Cod t\<^bold>\\ \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\ = inv \Cod t\<^bold>\\ \ \Cod t\<^bold>\\ \ \t\" using 1 by simp also have "... = (inv \Cod t\<^bold>\\ \ \Cod t\<^bold>\\) \ \t\" using comp_assoc by simp also have "... = \t\" using assms red_in_Hom Ide_Cod Can_red iso_eval_Can comp_cod_arr by (simp add: comp_inv_arr' eval_simps'(4) eval_simps'(5)) finally show "\t\ = inv \Cod t\<^bold>\\ \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" by presburger next assume 1: "\t\ = inv \Cod t\<^bold>\\ \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" hence "\Cod t\<^bold>\\ \ \t\ = \Cod t\<^bold>\\ \ inv \Cod t\<^bold>\\ \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" by simp also have "... = (\Cod t\<^bold>\\ \ inv \Cod t\<^bold>\\) \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" using comp_assoc by simp also have "... = \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" proof - have "\Cod t\<^bold>\\ \ inv \Cod t\<^bold>\\ = cod \\<^bold>\t\<^bold>\\" using assms red_in_Hom Ide_Cod Can_red iso_eval_Can inv_is_inverse Nmlize_in_Hom comp_arr_inv by (simp add: eval_simps') thus ?thesis using assms red_in_Hom Ide_Cod Can_red iso_eval_Can Ide_Dom Nmlize_in_Hom comp_cod_arr by (auto simp add: eval_simps') qed finally show "coherent t" by blast qed lemma coherent_iff_coherent_Inv: assumes "Can t" shows "coherent t \ coherent (Inv t)" proof have 1: "\t. Can t \ coherent t \ coherent (Inv t)" proof - fix t assume "Can t" hence t: "Can t \ Arr t \ Ide (Dom t) \ Ide (Cod t) \ arr \t\ \ iso \t\ \ inverse_arrows \t\ (inv \t\) \ Can \<^bold>\t\<^bold>\ \ Arr \<^bold>\t\<^bold>\ \ arr \\<^bold>\t\<^bold>\\ \ iso \\<^bold>\t\<^bold>\\ \ \<^bold>\t\<^bold>\ \ VHom \<^bold>\Dom t\<^bold>\ \<^bold>\Cod t\<^bold>\ \ inverse_arrows \\<^bold>\t\<^bold>\\ (inv \\<^bold>\t\<^bold>\\) \ Inv t \ VHom (Cod t) (Dom t)" using assms Can_implies_Arr Ide_Dom Ide_Cod iso_eval_Can inv_is_inverse Nmlize_in_Hom Can_Nmlize_Can Inv_in_Hom by simp assume coh: "coherent t" have "\Cod (Inv t)\<^bold>\\ \ \Inv t\ = (inv \\<^bold>\t\<^bold>\\ \ \\<^bold>\t\<^bold>\\) \ \Cod (Inv t)\<^bold>\\ \ \Inv t\" using t comp_inv_arr red_in_Hom comp_cod_arr [of "\Cod (Inv t)\<^bold>\\ \ \Inv t\" "inv \\<^bold>\t\<^bold>\\ \ \\<^bold>\t\<^bold>\\"] by (auto simp add: eval_simps') also have "... = inv \\<^bold>\t\<^bold>\\ \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\ \ inv \t\" using t eval_Inv_Can comp_assoc by auto also have "... = inv \\<^bold>\t\<^bold>\\ \ (\\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\) \ inv \t\" using comp_assoc by simp also have "... = inv \\<^bold>\t\<^bold>\\ \ (\Cod t\<^bold>\\ \ \t\) \ inv \t\" using t coh by simp also have "... = inv \\<^bold>\t\<^bold>\\ \ \Cod t\<^bold>\\ \ \t\ \ inv \t\" using comp_assoc by simp also have "... = \\<^bold>\Inv t\<^bold>\\ \ \Dom (Inv t)\<^bold>\\" proof - have "\Cod t\<^bold>\\ \ \t\ \ inv \t\ = \Dom (Inv t)\<^bold>\\" using t eval_Inv_Can red_in_Hom comp_arr_inv comp_arr_dom by (simp add: eval_simps') thus ?thesis using t Nmlize_Inv eval_Inv_Can by simp qed finally show "coherent (Inv t)" by blast qed show "coherent t \ coherent (Inv t)" using assms 1 by simp show "coherent (Inv t) \ coherent t" proof - assume "coherent (Inv t)" hence "coherent (Inv (Inv t))" using assms 1 Can_Inv by blast thus ?thesis using assms by simp qed qed text \ The next two facts are trivially proved by the simplifier, so formal named facts are not really necessary, but we include them for logical completeness of the following development, which proves coherence by structural induction. \ lemma coherent_Prim\<^sub>0: assumes "C.obj a" shows "coherent \<^bold>\a\<^bold>\\<^sub>0" by simp lemma coherent_Prim: assumes "Arr \<^bold>\f\<^bold>\" shows "coherent \<^bold>\f\<^bold>\" using assms by simp lemma coherent_Lunit_Ide: assumes "Ide t" shows "coherent \<^bold>\\<^bold>[t\<^bold>]" proof - have t: "Ide t \ Arr t \ Dom t = t \ Cod t = t \ ide \t\ \ ide \\<^bold>\t\<^bold>\\ \ \t\<^bold>\\ \ hom \t\ \\<^bold>\t\<^bold>\\" using assms Ide_in_Hom Ide_Nmlize_Ide red_in_Hom eval_in_hom ide_eval_Ide by force have 1: "Obj (Trg t)" using t by auto have "\Cod \<^bold>\\<^bold>[t\<^bold>]\<^bold>\\ \ \\<^bold>\\<^bold>[t\<^bold>]\ = \[\\<^bold>\t\<^bold>\\] \ (\Trg t\ \ \t\<^bold>\\)" using t \_ide_simp lunit_naturality [of "\t\<^bold>\\"] red_in_Hom by (simp add: eval_simps') also have "... = \\<^bold>\t\<^bold>\\ \ \[\\<^bold>\t\<^bold>\\] \ (\Trg t\ \ \t\<^bold>\\)" using t 1 lunit_in_hom Nmlize_in_Hom ide_eval_Ide red_in_Hom comp_cod_arr Obj_implies_Ide by (auto simp add: eval_simps') also have "... = \\<^bold>\\<^bold>\\<^bold>[t\<^bold>]\<^bold>\\ \ \Dom \<^bold>\\<^bold>[t\<^bold>]\<^bold>\\" using 1 t Nml_Trg \_ide_simp by (cases "Trg t"; simp) finally show ?thesis by simp qed text \ Unlike many of the other results, the next one was not quite so straightforward to adapt from @{session \MonoidalCategory\}. \ lemma coherent_Runit_Ide: assumes "Ide t" shows "coherent \<^bold>\\<^bold>[t\<^bold>]" proof - have t: "Ide t \ Arr t \ Dom t = t \ Cod t = t \ ide \t\ \ ide \\<^bold>\t\<^bold>\\ \ \t\<^bold>\\ \ hom \t\ \\<^bold>\t\<^bold>\\" using assms Ide_in_Hom Ide_Nmlize_Ide red_in_Hom eval_in_hom ide_eval_Ide by force have "\Cod \<^bold>\\<^bold>[t\<^bold>]\<^bold>\\ \ \\<^bold>\\<^bold>[t\<^bold>]\ = \[\\<^bold>\t\<^bold>\\] \ (\t\<^bold>\\ \ \Src t\)" using t \_ide_simp red_in_Hom runit_naturality [of "\t\<^bold>\\"] by (simp add: eval_simps') also have "... = \\<^bold>\\<^bold>\\<^bold>[t\<^bold>]\<^bold>\\ \ \Dom \<^bold>\\<^bold>[t\<^bold>]\<^bold>\\" proof - have "\\<^bold>\\<^bold>\\<^bold>[t\<^bold>]\<^bold>\\ \ \Dom \<^bold>\\<^bold>[t\<^bold>]\<^bold>\\ = \\<^bold>\t\<^bold>\\ \ \\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\Src t\<^bold>\\ \ (\t\<^bold>\\ \ \Src t\<^bold>\\)" using t by (cases t; simp; cases "Src t"; simp) also have "... = (\\<^bold>\t\<^bold>\\ \ \\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\Src t\<^bold>\\) \ (\t\<^bold>\\ \ \Src t\<^bold>\\)" proof - have "\\<^bold>\t\<^bold>\\ \ hom \\<^bold>\t\<^bold>\\ \\<^bold>\t\<^bold>\\" using t Nmlize_in_Hom by auto moreover have "\\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\Src t\<^bold>\\ \ hom (\\<^bold>\t\<^bold>\\ \ \Src t\) \\<^bold>\t\<^bold>\\" proof - have "\\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\Src t\<^bold>\\ \ hom \\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\Src t\<^bold>\\ \\<^bold>\t\<^bold>\\" proof - have "Src \<^bold>\t\<^bold>\ = Trg \<^bold>\Src t\<^bold>\ \ \<^bold>\\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\Src t\<^bold>\\<^bold>\ = \<^bold>\t\<^bold>\" using t Nmlize_Src Nml_Nmlize HcompNml_Nml_Src [of "\<^bold>\t\<^bold>\"] by simp thus ?thesis using t Ide_Nmlize_Ide Nml_Nmlize Obj_Src red2_in_Hom(2) Obj_implies_Ide by (auto simp add: eval_simps') qed thus ?thesis using t Nmlize_in_Hom Nmlize_Src by simp qed moreover have "\t\<^bold>\\ \ \Src t\<^bold>\\ \ hom (\t\ \ \Src t\) (\\<^bold>\t\<^bold>\\ \ \Src t\)" using t red_in_Hom red_Src Obj_Src eval_simps' Obj_implies_Ide by auto ultimately show ?thesis using comp_assoc by fastforce qed also have "... = \[\\<^bold>\t\<^bold>\\] \ (\t\<^bold>\\ \ \Src t\<^bold>\\)" proof - have "\\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\Src t\<^bold>\\ = \[\\<^bold>\t\<^bold>\\]" proof - have "Nml \<^bold>\t\<^bold>\" using t Nml_Nmlize by blast moreover have "is_Prim\<^sub>0 \<^bold>\Src t\<^bold>\" using t is_Prim0_Src Nmlize_Src by presburger ultimately show ?thesis apply (cases "\<^bold>\t\<^bold>\"; simp; cases "\<^bold>\Src t\<^bold>\"; simp) using t unitor_coincidence \_ide_simp \_ide_simp Nmlize_in_Hom apply simp_all using t is_Prim0_Src apply (cases "\<^bold>\t\<^bold>\"; simp) using t Nmlize_Src unitor_coincidence preserves_obj by simp qed moreover have "\[\\<^bold>\t\<^bold>\\] \ hom (\\<^bold>\t\<^bold>\\ \ \Src t\) \\<^bold>\t\<^bold>\\" using t Nmlize_in_Hom by (auto simp add: eval_simps'(2)) ultimately show ?thesis using comp_cod_arr [of "\[\\<^bold>\t\<^bold>\\]"] by fastforce qed also have "... = \[\\<^bold>\t\<^bold>\\] \ (\t\<^bold>\\ \ \Src t\)" using t red_Src by auto finally show ?thesis by argo qed finally show ?thesis by blast qed lemma coherent_Lunit'_Ide: assumes "Ide a" shows "coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[a\<^bold>]" using assms Ide_implies_Can coherent_Lunit_Ide coherent_iff_coherent_Inv [of "Lunit a"] by simp lemma coherent_Runit'_Ide: assumes "Ide a" shows "coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[a\<^bold>]" using assms Ide_implies_Can coherent_Runit_Ide coherent_iff_coherent_Inv [of "Runit a"] by simp lemma red2_Nml_Src: assumes "Ide t" and "Nml t" shows "\t \<^bold>\ Src t\ = \[\t\]" using assms eval_red2_Nml_Prim\<^sub>0 is_Prim0_Src [of t] by (cases "Src t"; simp) lemma red2_Trg_Nml: assumes "Ide t" and "Nml t" shows "\Trg t \<^bold>\ t\ = \[\t\]" using assms is_Prim0_Trg [of t] \_ide_simp ide_eval_Ide by (cases "Trg t"; simp) lemma coherence_key_fact: assumes "Ide a \ Nml a" and "Ide b \ Nml b" and "Ide c \ Nml c" and "Src a = Trg b" and "Src b = Trg c" shows "\(a \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\a \<^bold>\ b\ \ \c\) = (\a \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (\a\ \ \b \<^bold>\ c\)) \ \[\a\, \b\, \c\]" proof - have "is_Prim\<^sub>0 b \ ?thesis" proof - assume b: "is_Prim\<^sub>0 b" have "\a \<^bold>\ c\ \ (\[\a\] \ \c\) = (\a \<^bold>\ c\ \ (\a\ \ \[\c\])) \ \[\a\, \Trg c\, \c\]" proof - have "Src b = Trg b" using b by (cases b; simp) thus ?thesis using assms triangle [of "\c\" "\a\"] ide_eval_Ide comp_assoc by (simp add: eval_simps') qed thus ?thesis using assms b HcompNml_Nml_Src [of a] HcompNml_Trg_Nml red2_Nml_Src [of a] red2_Trg_Nml by (cases b, simp_all) qed moreover have "\ \ is_Prim\<^sub>0 b; is_Prim\<^sub>0 c \ \ ?thesis" proof - assume b: "\ is_Prim\<^sub>0 b" assume c: "is_Prim\<^sub>0 c" have "\(a \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\a \<^bold>\ b\ \ \c\) = \(a \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ Src b\ \ (\a \<^bold>\ b\ \ src \b\)" using assms b c by (cases c, simp_all add: eval_simps') also have "... = \[\a \<^bold>\\<^bold>\\<^bold>\ b\] \ (\a \<^bold>\ b\ \ src \b\)" using assms red2_Nml_Src [of "a \<^bold>\\<^bold>\\<^bold>\ b"] Nml_HcompNml(1) Src_HcompNml Ide_HcompNml by simp also have "... = \a \<^bold>\ b\ \ \[\a\ \ \b\]" proof - have "\\a \<^bold>\ b\ : \a\ \ \b\ \ \a \<^bold>\\<^bold>\\<^bold>\ b\\" using assms red2_in_Hom eval_in_hom [of "a \<^bold>\ b"] by simp thus ?thesis using assms runit_naturality [of "\a \<^bold>\ b\"] arr_dom in_homE src_dom src_hcomp hcomp_simps(1) by (elim in_homE, metis) qed also have "... = (\a \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (\a\ \ \b \<^bold>\ c\)) \ \[\a\, \b\, \c\]" proof - have "(\a \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (\a\ \ \b \<^bold>\ c\)) \ \[\a\, \b\, \c\] = (\a \<^bold>\ b\ \ (\a\ \ \[\b\])) \ \[\a\, \b\, src \b\]" using assms c red2_Nml_Src [of b] by (cases c, simp_all add: eval_simps') thus ?thesis using assms runit_hcomp ide_eval_Ide comp_assoc by (simp add: eval_simps') qed finally show ?thesis by blast qed moreover have "\ \ is_Prim\<^sub>0 b; \ is_Prim\<^sub>0 c \ \ ?thesis" proof - assume b': "\ is_Prim\<^sub>0 b" hence b: "Ide b \ Nml b \ Arr b \ \ is_Prim\<^sub>0 b \ ide \b\ \ arr \b\ \ \<^bold>\b\<^bold>\ = b \ b\<^bold>\ = b \ Dom b = b \ Cod b = b" using assms Ide_Nmlize_Ide Ide_in_Hom ide_eval_Ide by simp assume c': "\ is_Prim\<^sub>0 c" hence c: "Ide c \ Nml c \ Arr c \ \ is_Prim\<^sub>0 c \ ide \c\ \ arr \c\ \ \<^bold>\c\<^bold>\ = c \ c\<^bold>\ = c \ Dom c = c \ Cod c = c" using assms Ide_Nmlize_Ide Ide_in_Hom ide_eval_Ide by simp have "\a. Ide a \ Nml a \ Src a = Trg b \ \(a \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\a \<^bold>\ b\ \ \c\) = (\a \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (\a\ \ \b \<^bold>\ c\)) \ \[\a\, \b\, \c\]" proof - fix a :: "'c term" show "Ide a \ Nml a \ Src a = Trg b \ \(a \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\a \<^bold>\ b\ \ \c\) = (\a \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (\a\ \ \b \<^bold>\ c\)) \ \[\a\, \b\, \c\]" apply (induct a) using b c HcompNml_in_Hom apply (simp_all add: HcompNml_Nml_Src HcompNml_Trg_Nml) proof - fix f assume f: "C.ide f \ C.arr f \ \<^bold>\src\<^sub>C f\<^bold>\\<^sub>0 = Trg b" show "\(\<^bold>\f\<^bold>\ \<^bold>\ b) \<^bold>\ c\ \ (\\<^bold>\f\<^bold>\ \<^bold>\ b\ \ \c\) = (\\<^bold>\f\<^bold>\ \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\ \ (E f \ \b \<^bold>\ c\)) \ \[E f, \b\, \c\]" proof - have "\(\<^bold>\f\<^bold>\ \<^bold>\ b) \<^bold>\ c\ \ (\\<^bold>\f\<^bold>\ \<^bold>\ b\ \ \c\) = ((E f \ \b \<^bold>\\<^bold>\\<^bold>\ c\) \ (E f \ \b \<^bold>\ c\) \ \[E f, \b\, \c\]) \ ((E f \ \b\) \ \c\)" proof - have "((E f \ \b \<^bold>\\<^bold>\\<^bold>\ c\) \ (E f \ \b \<^bold>\ c\) \ \[E f, \b\, \c\]) \ ((E f \ \b\) \ \c\) = ((E f \ \b \<^bold>\\<^bold>\\<^bold>\ c\) \ (E f \ \b \<^bold>\ c\) \ \[E f, \b\, \c\]) \ (\\<^bold>\f\<^bold>\ \<^bold>\ b\ \ \c\)" using f b red2_Nml by simp also have "... = (\\<^bold>\f\<^bold>\ \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\ \ (E f \ \b \<^bold>\ c\) \ \[E f, \b\, \c\]) \ (\\<^bold>\f\<^bold>\ \<^bold>\ b\ \ \c\)" proof - have "\\<^bold>\f\<^bold>\ \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\ = E f \ \b \<^bold>\\<^bold>\\<^bold>\ c\" using assms(5) b c is_Hcomp_HcompNml red2_Nml Nml_HcompNml(1) is_Hcomp_def by (metis eval.simps(2) eval.simps(3) red2.simps(4)) thus ?thesis by argo qed also have "... = \(\<^bold>\f\<^bold>\ \<^bold>\ b) \<^bold>\ c\ \ (\\<^bold>\f\<^bold>\ \<^bold>\ b\ \ \c\)" using b c \_def by (cases c, simp_all) finally show ?thesis by argo qed also have "... = ((E f \ \b \<^bold>\\<^bold>\\<^bold>\ c\) \ (E f \ \b \<^bold>\ c\)) \ \[E f, \b\, \c\]" proof - have "src (E f) = trg \b\" using b f preserves_src by (cases "Trg b", auto simp add: eval_simps') thus ?thesis using assms b c f comp_arr_dom comp_assoc by (auto simp add: eval_simps') qed also have "... = (\\<^bold>\f\<^bold>\ \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (E f \ \b \<^bold>\ c\)) \ \[E f, \b\, \c\]" using assms f b c Ide_HcompNml Nml_HcompNml is_Hcomp_HcompNml [of b c] \_def by (cases "b \<^bold>\\<^bold>\\<^bold>\ c") simp_all finally show ?thesis by blast qed next fix x assume x: "C.obj x \ \<^bold>\x\<^bold>\\<^sub>0 = Trg b" show "\b \<^bold>\ c\ \ (\Trg b \<^bold>\ b\ \ \c\) = (\Trg b \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\ \ (\Trg b\ \ \b \<^bold>\ c\)) \ \[\Trg b\, \b\, \c\]" proof - have 1: "Trg (b \<^bold>\\<^bold>\\<^bold>\ c) = Trg b" using assms b c Trg_HcompNml by blast have 2: "\Trg b \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\ = \[\b \<^bold>\\<^bold>\\<^bold>\ c\]" using assms b c 1 Nml_HcompNml red2_Trg_Nml [of "b \<^bold>\\<^bold>\\<^bold>\ c"] Ide_HcompNml by simp have "\b \<^bold>\ c\ \ (\Trg b \<^bold>\ b\ \ \c\) = \b \<^bold>\ c\ \ (\[\b\] \ \c\)" using b c 1 2 HcompNml_Trg_Nml red2_Trg_Nml Trg_HcompNml by simp also have "... = \b \<^bold>\ c\ \ \[\b\ \ \c\] \ \[\Trg b\, \b\, \c\]" using assms b c lunit_hcomp [of "\b\" "\c\"] by (metis (no_types, lifting) eval_simps'(3) eval_simps(2)) also have "... = (\b \<^bold>\ c\ \ \[\b\ \ \c\]) \ \[\Trg b\, \b\, \c\]" using comp_assoc by simp also have "... = (\[\b \<^bold>\\<^bold>\\<^bold>\ c\] \ (\Trg b\ \ \b \<^bold>\ c\)) \ \[\Trg b\, \b\, \c\]" using assms b c lunit_naturality [of "\b \<^bold>\ c\"] red2_in_Hom by (simp add: eval_simps') also have "... = (\Trg b \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\ \ (\Trg b\ \ \b \<^bold>\ c\)) \ \[\Trg b\, \b\, \c\]" using b c 1 2 HcompNml_Trg_Nml red2_Trg_Nml Trg_HcompNml comp_assoc by simp finally show ?thesis by blast qed next fix d e assume I: "Nml e \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\e \<^bold>\ b\ \ \c\) = (\e \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\ \ (\e\ \ \b \<^bold>\ c\)) \ \[\e\, \b\, \c\]" assume de: "Ide d \ Ide e \ Src d = Trg e \ Nml (d \<^bold>\ e) \ Src e = Trg b" show "\((d \<^bold>\ e) \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\(d \<^bold>\ e) \<^bold>\ b\ \ \c\) = (\(d \<^bold>\ e) \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ ((\d\ \ \e\) \ \b \<^bold>\ c\)) \ \[\d\ \ \e\, \b\, \c\]" proof - let ?f = "un_Prim d" have "is_Prim d" using de Nml_HcompD by (metis term.disc(12)) hence "d = \<^bold>\?f\<^bold>\ \ C.ide ?f" using de by (cases d; simp) hence d: "Ide d \ Arr d \ Dom d = d \ Cod d = d \ Nml d \ d = \<^bold>\?f\<^bold>\ \ C.ide ?f \ ide \d\ \ arr \d\" using de ide_eval_Ide Nml_Nmlize(1) Ide_in_Hom Nml_HcompD [of d e] by simp have "Nml e \ \ is_Prim\<^sub>0 e" using de Nml_HcompD by metis hence e: "Ide e \ Arr e \ Dom e = e \ Cod e = e \ Nml e \ \ is_Prim\<^sub>0 e \ ide \e\ \ arr \e\" using de Ide_in_Hom ide_eval_Ide by simp have 1: "is_Hcomp (e \<^bold>\\<^bold>\\<^bold>\ b) \ is_Hcomp (b \<^bold>\\<^bold>\\<^bold>\ c) \ is_Hcomp (e \<^bold>\\<^bold>\\<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c)" using assms b c e de is_Hcomp_HcompNml [of e b] Nml_HcompNml is_Hcomp_HcompNml [of b c] is_Hcomp_HcompNml [of e "b \<^bold>\\<^bold>\\<^bold>\ c"] by auto have eb: "Src e = Trg b" using assms b c e de by argo have bc: "Src b = Trg c" using assms b c by simp have 4: "is_Hcomp ((e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c)" using assms b c e eb de 1 is_Hcomp_HcompNml [of e b] is_Hcomp_HcompNml [of "e \<^bold>\\<^bold>\\<^bold>\ b" c] is_Hcomp_HcompNml [of e b] Nml_HcompNml(1) [of e b] Src_HcompNml by auto have "\((d \<^bold>\ e) \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\(d \<^bold>\ e) \<^bold>\ b\ \ \c\) = ((\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c\) \ (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ \[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\]) \ ((\d\ \ \e \<^bold>\\<^bold>\\<^bold>\ b\) \ (\d\ \ \e \<^bold>\ b\) \ \[\d\, \e\, \b\] \ \c\)" proof - have "\((d \<^bold>\ e) \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ = (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c\) \ (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ \[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\]" proof - have "((d \<^bold>\ e) \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ b)) \<^bold>\ c" using b c d e de 1 HcompNml_Nml Nml_HcompNml HcompNml_assoc HcompNml_Prim by (metis term.distinct_disc(4)) also have "... = (d \<^bold>\ ((e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c)) \<^bold>\ (d \<^bold>\ ((e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c)) \<^bold>\ \<^bold>\\<^bold>[d, e \<^bold>\\<^bold>\\<^bold>\ b, c\<^bold>]" using b c d e de 1 Nml_HcompNml Nmlize_Nml by (cases c) simp_all also have "... = (d \<^bold>\ ((e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c)) \<^bold>\ (d \<^bold>\ ((e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c)) \<^bold>\ \<^bold>\\<^bold>[d, e \<^bold>\\<^bold>\\<^bold>\ b, c\<^bold>]" using d 4 apply (cases d, simp_all) by (cases "(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c") simp_all finally show ?thesis using b c d e HcompNml_in_Hom red2_in_Hom Nml_HcompNml Ide_HcompNml \_def by simp qed moreover have "\(d \<^bold>\ e) \<^bold>\ b\ = (\d\ \ \e \<^bold>\\<^bold>\\<^bold>\ b\) \ (\d\ \ \e \<^bold>\ b\) \ \[\d\, \e\, \b\]" proof - have "(d \<^bold>\ e) \<^bold>\ b = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ b)) \<^bold>\ (d \<^bold>\ (e \<^bold>\ b)) \<^bold>\ \<^bold>\\<^bold>[d, e, b\<^bold>]" using b c d e de 1 HcompNml_Prim Nmlize_Nml by (cases b, simp_all) also have "... = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ b)) \<^bold>\ (d \<^bold>\ (e \<^bold>\ b)) \<^bold>\ \<^bold>\\<^bold>[d, e, b\<^bold>]" using b c d e de 1 HcompNml_Nml Nml_HcompNml apply (cases d, simp_all) by (cases "e \<^bold>\\<^bold>\\<^bold>\ b", simp_all) finally show ?thesis using b d e HcompNml_in_Hom red2_in_Hom \_def by simp qed ultimately show ?thesis by argo qed also have "... = ((\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ \[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\]) \ ((\d\ \ \e \<^bold>\ b\) \ \c\) \ (\[\d\, \e\, \b\] \ \c\)" proof - have "(\d\ \ \e \<^bold>\\<^bold>\\<^bold>\ b\) \ (\d\ \ \e \<^bold>\ b\) \ \[\d\, \e\, \b\] \ \c\ = ((\d\ \ \e \<^bold>\ b\) \ \c\) \ (\[\d\, \e\, \b\] \ \c\)" using assms b c d e de eb HcompNml_in_Hom red2_in_Hom comp_cod_arr Ide_HcompNml Nml_HcompNml comp_assoc interchange [of "\d\ \ \e \<^bold>\ b\" "\[\d\, \e\, \b\]" "\c\" "\c\"] by (auto simp add: eval_simps') moreover have "(\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c\) \ (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ \[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\] = (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ \[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\]" proof - have "(\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c\) \ (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ \[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\] = ((\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c\) \ (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\)) \ \[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\]" using comp_assoc by simp thus ?thesis using assms b c d e de eb HcompNml_in_Hom red2_in_Hom Ide_HcompNml Nml_HcompNml comp_cod_arr by (simp add: eval_simps') qed ultimately show ?thesis by argo qed also have "... = (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ (\d\ \ (\e \<^bold>\ b\ \ \c\)) \ \[\d\, \e\ \ \b\, \c\] \ (\[\d\, \e\, \b\] \ \c\)" using assms b c d e de HcompNml_in_Hom red2_in_Hom Ide_HcompNml Nml_HcompNml ide_eval_Ide assoc_naturality [of "\d\" "\e \<^bold>\ b\" "\c\"] comp_permute [of "\[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\]" "(\d\ \ \e \<^bold>\ b\) \ \c\" "\d\ \ (\e \<^bold>\ b\ \ \c\)" "\[\d\, \e\ \ \b\, \c\]"] comp_assoc by (simp add: eval_simps') also have "... = ((\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ (\d\ \ (\e \<^bold>\ b\ \ \c\))) \ \[\d\, \e\ \ \b\, \c\] \ (\[\d\, \e\, \b\] \ \c\)" using comp_assoc by simp also have "... = (((\d\ \ \e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\) \ (\d\ \ \e\ \ \b \<^bold>\ c\)) \ (\d\ \ \[\e\, \b\, \c\])) \ \[\d\, \e\ \ \b\, \c\] \ (\[\d\, \e\, \b\] \ \c\)" using assms b c d e de eb I HcompNml_in_Hom red2_in_Hom Ide_HcompNml Nml_HcompNml whisker_left [of "\d\"] interchange [of "\d\" "\d\" "\(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\" "\e \<^bold>\ b\ \ \c\"] by (auto simp add: eval_simps') also have "... = ((\d\ \ \e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\) \ (\d\ \ \e\ \ \b \<^bold>\ c\)) \ ((\d\ \ \[\e\, \b\, \c\]) \ \[\d\, \e\ \ \b\, \c\] \ (\[\d\, \e\, \b\] \ \c\))" using comp_assoc by simp also have "... = ((\d\ \ \e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\) \ (\d\ \ (\e\ \ \b \<^bold>\ c\))) \ \[\d\, \e\, \b\ \ \c\] \ \[\d\ \ \e\, \b\, \c\]" using assms b c d e de pentagon by (simp add: eval_simps') also have "... = (\d\ \ \e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\) \ ((\d\ \ (\e\ \ \b \<^bold>\ c\)) \ \[\d\, \e\, \b\ \ \c\]) \ \[\d\ \ \e\, \b\, \c\]" using comp_assoc by simp also have "... = ((\d\ \ \e \<^bold>\\<^bold>\\<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\) \ (\d\ \ \e \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\)) \ (\[\d\, \e\, \b \<^bold>\\<^bold>\\<^bold>\ c\] \ ((\d\ \ \e\) \ \b \<^bold>\ c\)) \ \[\d\ \ \e\, \b\, \c\]" using assms d e de HcompNml_in_Hom red2_in_Hom Ide_HcompNml Nml_HcompNml assoc_naturality [of "\d\" "\e\" "\b \<^bold>\ c\"] comp_cod_arr [of "\d\ \ \e \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\"] by (simp add: eval_simps') also have "... = ((\d\ \ \e \<^bold>\\<^bold>\\<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\) \ (\d\ \ \e \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\) \ \[\d\, \e\, \b \<^bold>\\<^bold>\\<^bold>\ c\]) \ ((\d\ \ \e\) \ \b \<^bold>\ c\) \ \[\d\ \ \e\, \b\, \c\]" using comp_assoc by simp also have "... = \(d \<^bold>\ e) \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ ((\d\ \ \e\) \ \b \<^bold>\ c\) \ \[\d\ \ \e\, \b\, \c\]" proof - have "\(d \<^bold>\ e) \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ = (\d\ \ \e \<^bold>\\<^bold>\\<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\) \ (\d\ \ \e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\) \ \[\d\, \e\, \b \<^bold>\\<^bold>\\<^bold>\ c\]" proof - have "(d \<^bold>\ e) \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c) = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b \<^bold>\\<^bold>\\<^bold>\ c\<^bold>\)) \<^bold>\ (d \<^bold>\ (e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ \<^bold>\\<^bold>[d, e, b \<^bold>\\<^bold>\\<^bold>\ c\<^bold>]" using e 1 by (cases "b \<^bold>\\<^bold>\\<^bold>\ c") auto also have "... = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ (d \<^bold>\ (e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ \<^bold>\\<^bold>[d, e, b \<^bold>\\<^bold>\\<^bold>\ c\<^bold>]" using assms Nml_HcompNml Nmlize_Nml by simp also have "... = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ (d \<^bold>\ (e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ \<^bold>\\<^bold>[d, e, b \<^bold>\\<^bold>\\<^bold>\ c\<^bold>]" using d 1 apply (cases "e \<^bold>\\<^bold>\\<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c", simp_all) by (cases d, simp_all) finally show ?thesis using \_def by simp qed thus ?thesis by simp qed also have "... = (\(d \<^bold>\ e) \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ ((\d\ \ \e\) \ \b \<^bold>\ c\)) \ \[\d\ \ \e\, \b\, \c\]" using comp_assoc by simp finally show ?thesis by auto qed qed qed thus ?thesis using assms(1,4) by blast qed ultimately show ?thesis by blast qed lemma coherent_Assoc_Ide: assumes "Ide a" and "Ide b" and "Ide c" and "Src a = Trg b" and "Src b = Trg c" shows "coherent \<^bold>\\<^bold>[a, b, c\<^bold>]" proof - have a: "Ide a \ Arr a \ Dom a = a \ Cod a = a \ ide \a\ \ ide \\<^bold>\a\<^bold>\\ \ \a\<^bold>\\ \ hom \a\ \\<^bold>\a\<^bold>\\" using assms Ide_in_Hom Ide_Nmlize_Ide ide_eval_Ide red_in_Hom eval_in_hom(2) by force have b: "Ide b \ Arr b \ Dom b = b \ Cod b = b \ ide \b\ \ ide \\<^bold>\b\<^bold>\\ \ \b\<^bold>\\ \ hom \b\ \\<^bold>\b\<^bold>\\" using assms Ide_in_Hom Ide_Nmlize_Ide ide_eval_Ide red_in_Hom(2) eval_in_hom(2) [of "b\<^bold>\"] by auto have c: "Ide c \ Arr c \ Dom c = c \ Cod c = c \ ide \c\ \ ide \\<^bold>\c\<^bold>\\ \ \c\<^bold>\\ \ hom \c\ \\<^bold>\c\<^bold>\\" using assms Ide_in_Hom Ide_Nmlize_Ide red_in_Hom eval_in_hom(2) [of "c\<^bold>\"] ide_eval_Ide by auto have "\Cod \<^bold>\\<^bold>[a, b, c\<^bold>]\<^bold>\\ \ \\<^bold>\\<^bold>[a, b, c\<^bold>]\ = (\\<^bold>\a\<^bold>\ \<^bold>\ (\<^bold>\b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\)\ \ (\\<^bold>\a\<^bold>\\ \ \\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\\) \ (\a\<^bold>\\ \ \b\<^bold>\\ \ \c\<^bold>\\)) \ \[\a\, \b\, \c\]" using assms a b c red_in_Hom red2_in_Hom Nml_Nmlize Ide_Nmlize_Ide \_def eval_red_Hcomp interchange [of "\\<^bold>\a\<^bold>\\" "\a\<^bold>\\"] comp_cod_arr [of "\a\<^bold>\\"] by (simp add: eval_simps') also have "... = ((\\<^bold>\a\<^bold>\ \<^bold>\ (\<^bold>\b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\)\ \ (\\<^bold>\a\<^bold>\\ \ \\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\\)) \ \[\\<^bold>\a\<^bold>\\, \\<^bold>\b\<^bold>\\, \\<^bold>\c\<^bold>\\]) \ ((\a\<^bold>\\ \ \b\<^bold>\\) \ \c\<^bold>\\)" using assms red_in_Hom Ide_HcompNml assoc_naturality [of "\a\<^bold>\\" "\b\<^bold>\\" "\c\<^bold>\\"] comp_assoc by (simp add: eval_simps') also have "... = (\(\<^bold>\a\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b\<^bold>\) \<^bold>\ \<^bold>\c\<^bold>\\ \ (\\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\\ \ \\<^bold>\c\<^bold>\\)) \ ((\a\<^bold>\\ \ \b\<^bold>\\) \ \c\<^bold>\\)" using assms Nml_Nmlize Ide_Nmlize_Ide coherence_key_fact by simp also have "... = \\<^bold>\\<^bold>\\<^bold>[a, b, c\<^bold>]\<^bold>\\ \ \Dom \<^bold>\\<^bold>[a, b, c\<^bold>]\<^bold>\\" using assms a b c red_in_Hom red2_in_Hom Ide_Nmlize_Ide Nml_Nmlize eval_red_Hcomp HcompNml_assoc interchange [of "\\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\\" "\a\<^bold>\\ \ \b\<^bold>\\" "\\<^bold>\c\<^bold>\\" "\c\<^bold>\\"] comp_cod_arr [of "\c\<^bold>\\" "\\<^bold>\c\<^bold>\\"] apply (simp add: eval_simps') proof - have "seq \(\<^bold>\a\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b\<^bold>\) \<^bold>\ \<^bold>\c\<^bold>\\ ((\\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\\ \ \\<^bold>\c\<^bold>\\) \ ((\a\<^bold>\\ \ \b\<^bold>\\) \ \c\<^bold>\\))" using assms c red_in_Hom red2_in_Hom Nml_HcompNml Ide_Nmlize_Ide Ide_HcompNml Nml_Nmlize by (simp_all add: eval_simps') moreover have "cod (\(\<^bold>\a\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b\<^bold>\) \<^bold>\ \<^bold>\c\<^bold>\\ \ (\\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\\ \ \\<^bold>\c\<^bold>\\) \ ((\a\<^bold>\\ \ \b\<^bold>\\) \ \c\<^bold>\\)) = \\<^bold>\a\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\\" using assms c red_in_Hom red2_in_Hom Nml_HcompNml Ide_Nmlize_Ide Ide_HcompNml Nml_Nmlize HcompNml_assoc by (simp add: eval_simps') ultimately show "(\(\<^bold>\a\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b\<^bold>\) \<^bold>\ \<^bold>\c\<^bold>\\ \ (\\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\\ \ \\<^bold>\c\<^bold>\\)) \ ((\a\<^bold>\\ \ \b\<^bold>\\) \ \c\<^bold>\\) = \\<^bold>\a\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\\ \ \(\<^bold>\a\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b\<^bold>\) \<^bold>\ \<^bold>\c\<^bold>\\ \ (\\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\\ \ \\<^bold>\c\<^bold>\\) \ ((\a\<^bold>\\ \ \b\<^bold>\\) \ \c\<^bold>\\)" using comp_cod_arr comp_assoc by simp qed finally show ?thesis by blast qed lemma coherent_Assoc'_Ide: assumes "Ide a" and "Ide b" and "Ide c" and "Src a = Trg b" and "Src b = Trg c" shows "coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[a, b, c\<^bold>]" using assms Ide_implies_Can coherent_Assoc_Ide Inv_Ide coherent_iff_coherent_Inv Can.simps(10) Inv.simps(10) by presburger lemma eval_red2_naturality: assumes "Nml t" and "Nml u" and "Src t = Trg u" shows "\Cod t \<^bold>\ Cod u\ \ (\t\ \ \u\) = \t \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom t \<^bold>\ Dom u\" proof - have *: "\t u. Nml (t \<^bold>\ u) \ arr \t\ \ arr \u\" using Nml_implies_Arr Nml_HcompD by simp have "is_Prim\<^sub>0 t \ ?thesis" using assms Nml_implies_Arr is_Prim0_Trg \.naturality [of "\u\"] by (cases t) (simp_all add: eval_simps', cases "Trg t", simp_all) moreover have "\ is_Prim\<^sub>0 t \ is_Prim\<^sub>0 u \ ?thesis" using assms Nml_implies_Arr eval_red2_Nml_Prim\<^sub>0 runit_naturality [of "\t\"] by (cases u) (simp_all add: eval_simps') moreover have "\ is_Prim\<^sub>0 t \ \ is_Prim\<^sub>0 u \ ?thesis" using assms * Nml_implies_Arr apply (induct t, simp_all) proof - fix f assume f: "C.arr f" assume "\ is_Prim\<^sub>0 u" hence u: "\ is_Prim\<^sub>0 u \ Nml u \ Nml (Dom u) \ Nml (Cod u) \ Ide (Dom u) \ Ide (Cod u) \ arr \u\ \ arr \Dom u\ \ arr \Cod u\ \ ide \Dom u\ \ ide \Cod u\" using assms(2) Nml_implies_Arr ide_eval_Ide by simp hence 1: "\ is_Prim\<^sub>0 (Dom u) \ \ is_Prim\<^sub>0 (Cod u)" using u by (cases u) simp_all assume "\<^bold>\src\<^sub>C f\<^bold>\\<^sub>0 = Trg u" hence "\\<^bold>\src\<^sub>C f\<^bold>\\<^sub>0\ = \Trg u\" by simp hence fu: "src (E f) = trg \u\" using f u preserves_src Nml_implies_Arr by (simp add: eval_simps') show "\\<^bold>\C.cod f\<^bold>\ \<^bold>\ Cod u\ \ (E f \ \u\) = (E f \ \u\) \ \\<^bold>\C.dom f\<^bold>\ \<^bold>\ Dom u\" proof - have "\\<^bold>\C.cod f\<^bold>\ \<^bold>\ Cod u\ = E (C.cod f) \ cod \u\" using f u 1 Nml_implies_Arr by (cases "Cod u", simp_all add: eval_simps') moreover have "\\<^bold>\C.dom f\<^bold>\ \<^bold>\ Dom u\ = E (C.dom f) \ dom \u\" using f u 1 Nml_implies_Arr by (cases "Dom u", simp_all add: eval_simps') moreover have "\E f \ \u\ : E (C.dom f) \ \Dom u\ \ E (C.cod f) \ \Cod u\\" using f u fu Nml_implies_Arr apply (intro hcomp_in_vhom) apply auto by (metis C.src_dom eval_simps(4) preserves_src trg_dom u) ultimately show ?thesis using f u comp_arr_dom comp_cod_arr by (simp add: fu) qed next fix v w assume I2: "\ \ is_Prim\<^sub>0 w; Nml w \ \ \Cod w \<^bold>\ Cod u\ \ (\w\ \ \u\) = \w \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom w \<^bold>\ Dom u\" assume "\ is_Prim\<^sub>0 u" hence u: "\ is_Prim\<^sub>0 u \ Arr u \ Arr (Dom u) \ Arr (Cod u) \ Nml u \ Nml (Dom u) \ Nml (Cod u) \ Ide (Dom u) \ Ide (Cod u) \ arr \u\ \ arr \Dom u\ \ arr \Cod u\ \ ide \Dom u\ \ ide \Cod u\" using assms(2) Nml_implies_Arr ide_eval_Ide by simp assume vw: "Nml (v \<^bold>\ w)" assume wu: "Src w = Trg u" let ?f = "un_Prim v" have "v = \<^bold>\?f\<^bold>\ \ C.arr ?f" using vw by (metis Nml_HcompD(1) Nml_HcompD(2)) hence "Arr v \ v = \<^bold>\un_Prim v\<^bold>\ \ C.arr ?f \ Nml v" by (cases v; simp) hence v: "v = \<^bold>\?f\<^bold>\ \ C.arr ?f \ Arr v \ Arr (Dom v) \ Arr (Cod v) \ Nml v \ Nml (Dom v) \ Nml (Cod v) \ arr \v\ \ arr \Dom v\ \ arr \Cod v\ \ ide \Dom v\ \ ide \Cod v\" using vw * by (cases v, simp_all) have "Nml w \ \ is_Prim\<^sub>0 w" using vw v by (metis Nml.simps(3)) hence w: "\ is_Prim\<^sub>0 w \ Arr w \ Arr (Dom w) \ Arr (Cod w) \ Nml w \ Nml (Dom w) \ Nml (Cod w) \ Ide (Dom w) \ Ide (Cod w) \ arr \w\ \ arr \Dom w\ \ arr \Cod w\ \ ide \Dom w\ \ ide \Cod w\" using vw * Nml_implies_Arr ide_eval_Ide by simp have u': "\ is_Prim\<^sub>0 (Dom u) \ \ is_Prim\<^sub>0 (Cod u)" using u by (cases u, simp_all) have w': "\ is_Prim\<^sub>0 (Dom w) \ \ is_Prim\<^sub>0 (Cod w)" using w by (cases w, simp_all) have vw': "Src v = Trg w" using vw Nml_HcompD(7) by simp have X: "Nml (Dom v \<^bold>\ (Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u))" using u u' v w w' wu vw is_Hcomp_HcompNml Nml_HcompNml Nml_Dom - by (cases v) auto + by (metis Dom.simps(3) Nml.simps(3) term.distinct_disc(3)) have Y: "Nml (Cod v \<^bold>\ (Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u))" using u u' w w' wu vw is_Hcomp_HcompNml Nml_HcompNml Src_Cod Trg_Cod - by (cases v) auto + by (metis Cod.simps(3) Nml.simps(3) Nml_Cod term.distinct_disc(3) v) show "\(Cod v \<^bold>\ Cod w) \<^bold>\ Cod u\ \ ((\v\ \ \w\) \ \u\) = \(v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u\ \ \(Dom v \<^bold>\ Dom w) \<^bold>\ Dom u\" proof - have "\(Cod v \<^bold>\ Cod w) \<^bold>\ Cod u\ \ ((\v\ \ \w\) \ \u\) = (\Cod v \<^bold>\ (Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u)\ \ (\Cod v\ \ \Cod w \<^bold>\ Cod u\) \ \[\Cod v\, \Cod w\, \Cod u\]) \ ((\v\ \ \w\) \ \u\)" proof - have "(Cod v \<^bold>\ Cod w) \<^bold>\ Cod u = (Cod v \<^bold>\ (Cod w \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod u\<^bold>\)) \<^bold>\ (Cod v \<^bold>\ Cod w \<^bold>\ Cod u) \<^bold>\ \<^bold>\\<^bold>[Cod v, Cod w, Cod u\<^bold>]" using u v w by (cases u) simp_all hence "\(Cod v \<^bold>\ Cod w) \<^bold>\ Cod u\ = \Cod v \<^bold>\ (Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u)\ \ (\Cod v\ \ \Cod w \<^bold>\ Cod u\) \ \[\Cod v\, \Cod w\, \Cod u\]" using u v w \_def by simp thus ?thesis by presburger qed also have "... = ((\Cod v\ \ \Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u\) \ (\Cod v\ \ \Cod w \<^bold>\ Cod u\) \ \[\Cod v\, \Cod w\, \Cod u\]) \ ((\v\ \ \w\) \ \u\)" using u v w Y red2_Nml by simp also have "... = ((\Cod v\ \ \Cod w \<^bold>\ Cod u\) \ \[\Cod v\, \Cod w\, \Cod u\]) \ ((\v\ \ \w\) \ \u\)" using u v w vw' wu comp_cod_arr red2_in_Hom HcompNml_in_Hom comp_reduce by (simp add: eval_simps') also have "... = (\Cod v\ \ \Cod w \<^bold>\ Cod u\) \ \[\Cod v\, \Cod w\, \Cod u\] \ ((\v\ \ \w\) \ \u\)" using comp_assoc by simp also have "... = (\Cod v\ \ \Cod w \<^bold>\ Cod u\) \ (\v\ \ \w\ \ \u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using u v w vw' wu assoc_naturality [of "\v\" "\w\" "\u\"] by (simp add: eval_simps') also have "... = ((\Cod v\ \ \Cod w \<^bold>\ Cod u\) \ (\v\ \ \w\ \ \u\)) \ \[\Dom v\, \Dom w\, \Dom u\]" using comp_assoc by simp also have "... = (\v\ \ \w \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using v w u vw' wu I2 red2_in_Hom HcompNml_in_Hom comp_cod_arr interchange [of "\Cod v\" "\v\" "\Cod w \<^bold>\ Cod u\" "\w\ \ \u\"] by (simp add: eval_simps') also have "... = ((\v\ \ \w \<^bold>\\<^bold>\\<^bold>\ u\) \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\)) \ \[\Dom v\, \Dom w\, \Dom u\]" using v w u vw' wu red2_in_Hom HcompNml_in_Hom comp_arr_dom interchange [of "\v\" "\Dom v\" "\w \<^bold>\\<^bold>\\<^bold>\ u\" "\Dom w \<^bold>\ Dom u\"] by (simp add: eval_simps') also have "... = (\v\ \ \w \<^bold>\\<^bold>\\<^bold>\ u\) \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using comp_assoc by simp also have "... = \v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using u u' v w vw' wu is_Hcomp_HcompNml HcompNml_Prim [of "w \<^bold>\\<^bold>\\<^bold>\ u" ?f] by force also have "... = \v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom v \<^bold>\\<^bold>\\<^bold>\ Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" proof - have "\v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\] = (\v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom v \<^bold>\\<^bold>\\<^bold>\ Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u\) \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using u v w vw' wu comp_arr_dom Nml_HcompNml HcompNml_in_Hom by (simp add: eval_simps') also have "... = \v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom v \<^bold>\\<^bold>\\<^bold>\ Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using comp_assoc by simp finally show ?thesis by blast qed also have "... = \(v \<^bold>\\<^bold>\\<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u\ \ \(Dom v \<^bold>\ Dom w) \<^bold>\ Dom u\" proof - have "(Dom v \<^bold>\ Dom w) \<^bold>\ Dom u = (Dom v \<^bold>\ (Dom w \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom u\<^bold>\)) \<^bold>\ (Dom v \<^bold>\ (Dom w \<^bold>\ Dom u)) \<^bold>\ \<^bold>\\<^bold>[Dom v, Dom w, Dom u\<^bold>]" using u u' v w vw' wu by (cases u, simp_all) hence "\(Dom v \<^bold>\ Dom w) \<^bold>\ Dom u\ = \Dom v \<^bold>\ (Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u)\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using u v w \_def by simp also have "... = \Dom v \<^bold>\\<^bold>\\<^bold>\ Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using X HcompNml_Nml red2_Nml by presburger finally have "\(Dom v \<^bold>\ Dom w) \<^bold>\ Dom u\ = \Dom v \<^bold>\\<^bold>\\<^bold>\ Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" by blast thus ?thesis using assms v w vw' wu HcompNml_assoc by presburger qed finally show ?thesis using vw HcompNml_Nml by simp qed qed ultimately show ?thesis by blast qed lemma coherent_Hcomp: assumes "Arr t" and "Arr u" and "Src t = Trg u" and "coherent t" and "coherent u" shows "coherent (t \<^bold>\ u)" proof - have t: "Arr t \ Ide (Dom t) \ Ide (Cod t) \ Ide \<^bold>\Dom t\<^bold>\ \ Ide \<^bold>\Cod t\<^bold>\ \ arr \t\ \ arr \Dom t\ \ ide \Dom t\ \ arr \Cod t\ \ ide \Cod t\" using assms Ide_Nmlize_Ide ide_eval_Ide by auto have u: "Arr u \ Ide (Dom u) \ Ide (Cod u) \ Ide \<^bold>\Dom u\<^bold>\ \ Ide \<^bold>\Cod u\<^bold>\ \ arr \u\ \ arr \Dom u\ \ ide \Dom u\ \ arr \Cod u\ \ ide \Cod u\" using assms Ide_Nmlize_Ide ide_eval_Ide by auto have "\Cod (t \<^bold>\ u)\<^bold>\\ \ (\t\ \ \u\) = (\\<^bold>\Cod t\<^bold>\ \<^bold>\ \<^bold>\Cod u\<^bold>\\ \ (\Cod t\<^bold>\\ \ \Cod u\<^bold>\\)) \ (\t\ \ \u\)" using t u eval_red_Hcomp by simp also have "... = \\<^bold>\Cod t\<^bold>\ \<^bold>\ \<^bold>\Cod u\<^bold>\\ \ (\Cod t\<^bold>\\ \ \Cod u\<^bold>\\) \ (\t\ \ \u\)" using comp_assoc by simp also have "... = \\<^bold>\Cod t\<^bold>\ \<^bold>\ \<^bold>\Cod u\<^bold>\\ \ (\\<^bold>\t\<^bold>\\ \ \\<^bold>\u\<^bold>\\) \ (\Dom t\<^bold>\\ \ \Dom u\<^bold>\\)" using assms t u Nmlize_in_Hom red_in_Hom interchange [of "\Cod t\<^bold>\\" "\t\" "\Cod u\<^bold>\\" "\u\"] interchange [of "\\<^bold>\t\<^bold>\\" "\Dom t\<^bold>\\" "\\<^bold>\u\<^bold>\\" "\Dom u\<^bold>\\"] by (simp add: eval_simps') also have "... = (\\<^bold>\Cod t\<^bold>\ \<^bold>\ \<^bold>\Cod u\<^bold>\\ \ (\\<^bold>\t\<^bold>\\ \ \\<^bold>\u\<^bold>\\)) \ (\Dom t\<^bold>\\ \ \Dom u\<^bold>\\)" using comp_assoc by simp also have "... = (\\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\\ \ \\<^bold>\Dom t\<^bold>\ \<^bold>\ \<^bold>\Dom u\<^bold>\\) \ (\Dom t\<^bold>\\ \ \Dom u\<^bold>\\)" using assms t u Nml_Nmlize Nmlize_in_Hom eval_red2_naturality [of "Nmlize t" "Nmlize u"] by simp also have "... = \\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\\ \ \\<^bold>\Dom t\<^bold>\ \<^bold>\ \<^bold>\Dom u\<^bold>\\ \ (\Dom t\<^bold>\\ \ \Dom u\<^bold>\\)" using comp_assoc by simp also have "... = \\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\\ \ \(Dom t \<^bold>\ Dom u)\<^bold>\\" using t u eval_red_Hcomp by simp finally have "\Cod (t \<^bold>\ u)\<^bold>\\ \ (\t\ \ \u\) = \\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\\ \ \(Dom t \<^bold>\ Dom u)\<^bold>\\" by blast thus ?thesis using t u by simp qed lemma coherent_Vcomp: assumes "Arr t" and "Arr u" and "Dom t = Cod u" and "coherent t" and "coherent u" shows "coherent (t \<^bold>\ u)" proof - have t: "Arr t \ Ide (Dom t) \ Ide (Cod t) \ Ide \<^bold>\Dom t\<^bold>\ \ Ide \<^bold>\Cod t\<^bold>\ \ arr \t\ \ arr \Dom t\ \ ide \Dom t\ \ arr \Cod t\ \ ide \Cod t\" using assms Ide_Nmlize_Ide ide_eval_Ide by auto have u: "Arr u \ Ide (Dom u) \ Ide (Cod u) \ Ide \<^bold>\Dom u\<^bold>\ \ Ide \<^bold>\Cod u\<^bold>\ \ arr \u\ \ arr \Dom u\ \ ide \Dom u\ \ arr \Cod u\ \ ide \Cod u\" using assms Ide_Nmlize_Ide ide_eval_Ide by auto have "\Cod (t \<^bold>\ u)\<^bold>\\ \ \t \<^bold>\ u\ = \Cod t\<^bold>\\ \ \t\ \ \u\" using t u by simp also have "... = (\Cod t\<^bold>\\ \ \t\) \ \u\" proof - have "seq \Cod t\<^bold>\\ \t\" using assms t red_in_Hom by (intro seqI, auto simp add: eval_simps') moreover have "seq \t\ \u\" using assms t u by (auto simp add: eval_simps') ultimately show ?thesis using comp_assoc by auto qed also have "... = \\<^bold>\t \<^bold>\ u\<^bold>\\ \ \Dom (t \<^bold>\ u)\<^bold>\\" using t u assms red_in_Hom Nml_Nmlize comp_assoc by (simp add: eval_simps' Nml_implies_Arr eval_VcompNml) finally show "coherent (t \<^bold>\ u)" by blast qed text \ The main result: ``Every formal arrow is coherent.'' \ theorem coherence: assumes "Arr t" shows "coherent t" proof - have "Arr t \ coherent t" proof (induct t) show "\a. Arr \<^bold>\a\<^bold>\\<^sub>0 \ coherent \<^bold>\a\<^bold>\\<^sub>0" by simp show "\\. Arr \<^bold>\\\<^bold>\ \ coherent \<^bold>\\\<^bold>\" by simp fix u v show "\ Arr u \ coherent u; Arr v \ coherent v \ \ Arr (u \<^bold>\ v) \ coherent (u \<^bold>\ v)" using coherent_Hcomp by simp show "\ Arr u \ coherent u; Arr v \ coherent v \ \ Arr (u \<^bold>\ v) \ coherent (u \<^bold>\ v)" using coherent_Vcomp by simp next fix t assume I: "Arr t \ coherent t" show Lunit: "Arr \<^bold>\\<^bold>[t\<^bold>] \ coherent \<^bold>\\<^bold>[t\<^bold>]" using I Ide_Dom coherent_Lunit_Ide Ide_in_Hom coherent_Vcomp [of t "\<^bold>\\<^bold>[Dom t\<^bold>]"] Nmlize_Vcomp_Arr_Dom eval_in_hom \.is_natural_1 [of "\t\"] by force show Runit: "Arr \<^bold>\\<^bold>[t\<^bold>] \ coherent \<^bold>\\<^bold>[t\<^bold>]" using I Ide_Dom coherent_Runit_Ide Ide_in_Hom ide_eval_Ide coherent_Vcomp [of t "\<^bold>\\<^bold>[Dom t\<^bold>]"] Nmlize_Vcomp_Arr_Dom \_ide_simp eval_in_hom \.is_natural_1 [of "\t\"] by force show "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]" proof - assume "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]" hence t: "Arr t" by simp have "coherent (\<^bold>\\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>] \<^bold>\ t)" using t I Ide_Cod coherent_Lunit'_Ide Ide_in_Hom coherent_Vcomp [of "\<^bold>\\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>]" t] Arr.simps(6) Dom.simps(6) Dom_Cod Ide_implies_Arr by presburger moreover have "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>] \<^bold>\ t\" using t \'.is_natural_2 [of "\t\"] by (simp add: eval_simps(5)) ultimately show ?thesis using t Nmlize_Vcomp_Cod_Arr by simp qed show "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]" proof - assume "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]" hence t: "Arr t" by simp have "coherent (\<^bold>\\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>] \<^bold>\ t)" using t I Ide_Cod coherent_Runit'_Ide Ide_in_Hom coherent_Vcomp [of "\<^bold>\\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>]" t] Arr.simps(8) Dom.simps(8) Dom_Cod Ide_implies_Arr by presburger moreover have "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>] \<^bold>\ t\" using t \'.is_natural_2 [of "\t\"] by (simp add: eval_simps(5)) ultimately show ?thesis using t Nmlize_Vcomp_Cod_Arr by simp qed next fix t u v assume I1: "Arr t \ coherent t" assume I2: "Arr u \ coherent u" assume I3: "Arr v \ coherent v" show "Arr \<^bold>\\<^bold>[t, u, v\<^bold>] \ coherent \<^bold>\\<^bold>[t, u, v\<^bold>]" proof - assume tuv: "Arr \<^bold>\\<^bold>[t, u, v\<^bold>]" have t: "Arr t" using tuv by simp have u: "Arr u" using tuv by simp have v: "Arr v" using tuv by simp have tu: "Src t = Trg u" using tuv by simp have uv: "Src u = Trg v" using tuv by simp have "coherent ((t \<^bold>\ u \<^bold>\ v) \<^bold>\ \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>])" proof - have "Arr (t \<^bold>\ u \<^bold>\ v) \ coherent (t \<^bold>\ u \<^bold>\ v)" using t u v tu uv tuv I1 I2 I3 coherent_Hcomp Arr.simps(3) Trg.simps(3) by presburger moreover have "Arr \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v tu uv Ide_Dom by simp moreover have "coherent \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v tu uv Src_Dom Trg_Dom Ide_Dom coherent_Assoc_Ide by metis moreover have "Dom (t \<^bold>\ u \<^bold>\ v) = Cod \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v by simp ultimately show ?thesis using t u v coherent_Vcomp by blast qed moreover have "VPar \<^bold>\\<^bold>[t, u, v\<^bold>] ((t \<^bold>\ u \<^bold>\ v) \<^bold>\ \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>])" using t u v tu uv Ide_Dom by simp moreover have "\<^bold>\\<^bold>\\<^bold>[t, u, v\<^bold>]\<^bold>\ = \<^bold>\(t \<^bold>\ u \<^bold>\ v) \<^bold>\ \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>]\<^bold>\" proof - have "(\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\ = (\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ ((\<^bold>\Dom t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom u\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom v\<^bold>\)" proof - have 1: "Nml \<^bold>\t\<^bold>\ \ Nml \<^bold>\u\<^bold>\ \ Nml \<^bold>\v\<^bold>\ \ Dom \<^bold>\t\<^bold>\ = \<^bold>\Dom t\<^bold>\ \ Dom \<^bold>\u\<^bold>\ = \<^bold>\Dom u\<^bold>\ \ Dom \<^bold>\v\<^bold>\ = \<^bold>\Dom v\<^bold>\" using t u v Nml_Nmlize by blast moreover have "Nml (\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\)" using 1 t u tu Nmlize_Src Nmlize_Trg Nml_HcompNml(1) by presburger moreover have "\t. Arr t \ \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t\<^bold>\ = \<^bold>\t\<^bold>\" using t Nmlize_Vcomp_Arr_Dom by simp moreover have "Dom \<^bold>\\<^bold>\\<^bold>[t, u, v\<^bold>]\<^bold>\ = \<^bold>\Dom \<^bold>\\<^bold>[t, u, v\<^bold>]\<^bold>\" using Nml_Nmlize tuv by blast ultimately show ?thesis using t u v tu uv tuv 1 HcompNml_assoc Nml_HcompNml Nml_Nmlize VcompNml_Nml_Dom [of "(\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\"] by auto qed thus ?thesis using t u v Nmlize_Vcomp_Arr_Dom VcompNml_HcompNml Nml_Nmlize by simp qed moreover have "\\<^bold>\\<^bold>[t, u, v\<^bold>]\ = \(t \<^bold>\ u \<^bold>\ v) \<^bold>\ \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>]\" using t u v tu uv Ide_Dom comp_cod_arr ide_eval_Ide \_def apply (simp add: eval_simps') using assoc_is_natural_1 arr_eval_Arr eval_simps'(2-4) by presburger ultimately show "coherent \<^bold>\\<^bold>[t, u, v\<^bold>]" by argo qed show "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] \ coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]" proof - assume tuv: "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]" have t: "Arr t" using tuv by simp have u: "Arr u" using tuv by simp have v: "Arr v" using tuv by simp have tu: "Src t = Trg u" using tuv by simp have uv: "Src u = Trg v" using tuv by simp have "coherent (((t \<^bold>\ u) \<^bold>\ v) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>])" proof - have "Arr ((t \<^bold>\ u) \<^bold>\ v) \ coherent ((t \<^bold>\ u) \<^bold>\ v)" using t u v tu uv tuv I1 I2 I3 coherent_Hcomp Arr.simps(3) Src.simps(3) by presburger moreover have "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v tu uv Ide_Dom by simp moreover have "coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v tu uv Src_Dom Trg_Dom Ide_Dom coherent_Assoc'_Ide by metis moreover have "Dom ((t \<^bold>\ u) \<^bold>\ v) = Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v by simp ultimately show ?thesis using t u v coherent_Vcomp by metis qed moreover have "VPar \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] (((t \<^bold>\ u) \<^bold>\ v) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>])" using t u v tu uv Ide_Dom by simp moreover have "\<^bold>\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<^bold>\ = \<^bold>\((t \<^bold>\ u) \<^bold>\ v) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]\<^bold>\" using t u v tu uv Nmlize_Vcomp_Arr_Dom VcompNml_HcompNml Nml_Nmlize HcompNml_assoc Nml_HcompNml HcompNml_in_Hom VcompNml_Nml_Dom [of "(\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\"] by simp moreover have "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = \((t \<^bold>\ u) \<^bold>\ v) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]\" proof - have 1: "VVV.arr (\t\, \u\, \v\)" using tuv \'.preserves_reflects_arr arr_eval_Arr eval.simps(10) by (metis (no_types, lifting)) have "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = ((\t\ \ \u\) \ \v\) \ \\<^sup>-\<^sup>1[dom \t\, dom \u\, dom \v\]" proof - have "VVV.arr (\t\, \u\, \v\)" using tuv \'.preserves_reflects_arr arr_eval_Arr eval.simps(10) by (metis (no_types, lifting)) thus ?thesis using t u v \'.is_natural_1 [of "(\t\, \u\, \v\)"] HoHV_def \'_def VVV.dom_simp by simp qed also have "... = \((t \<^bold>\ u) \<^bold>\ v) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]\" by (simp add: eval_simps'(4) t u v \'_def) finally show ?thesis by blast qed ultimately show "coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]" by argo qed qed thus ?thesis using assms by blast qed corollary eval_eqI: assumes "VPar t u" and "\<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" shows "\t\ = \u\" using assms coherence canonical_factorization by simp text \ The following allows us to prove that two 1-cells in a bicategory are isomorphic simply by expressing them as the evaluations of terms having the same normalization. The benefits are: (1) we do not have to explicitly exhibit the isomorphism, which is canonical and is obtained by evaluating the reductions of the terms to their normalizations, and (2) the normalizations can be computed automatically by the simplifier. \ lemma canonically_isomorphicI: assumes "f = \t\" and "g = \u\" and "Ide t" and "Ide u" and "\<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" shows "f \ g" proof - have "f \ \t\" using assms isomorphic_reflexive ide_eval_Ide by blast also have "... \ \\<^bold>\t\<^bold>\\" proof - have "\\t\<^bold>\\ : \t\ \ \\<^bold>\t\<^bold>\\\ \ iso \t\<^bold>\\" using assms(1,3) Can_red iso_eval_Can red_in_Hom(2) eval_in_hom(2) by fastforce thus ?thesis using isomorphic_def by blast qed also have "... \ \\<^bold>\u\<^bold>\\" using assms isomorphic_reflexive by (simp add: Ide_Nmlize_Ide ide_eval_Ide) also have "... \ \u\" proof - have "\\u\<^bold>\\ : \u\ \ \\<^bold>\u\<^bold>\\\ \ iso \u\<^bold>\\" using assms(2,4) Can_red iso_eval_Can red_in_Hom(2) eval_in_hom(2) by fastforce thus ?thesis using isomorphic_def isomorphic_symmetric by blast qed also have "... \ g" using assms isomorphic_reflexive ide_eval_Ide by blast finally show ?thesis by simp qed end end diff --git a/thys/Bicategory/InternalEquivalence.thy b/thys/Bicategory/InternalEquivalence.thy --- a/thys/Bicategory/InternalEquivalence.thy +++ b/thys/Bicategory/InternalEquivalence.thy @@ -1,1444 +1,1445 @@ (* Title: InternalEquivalence Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) section "Internal Equivalences" theory InternalEquivalence imports Bicategory begin text \ An \emph{internal equivalence} in a bicategory consists of antiparallel 1-cells \f\ and \g\ together with invertible 2-cells \\\ : src f \ g \ f\\ and \\\ : f \ g \ src g\\. Objects in a bicategory are said to be \emph{equivalent} if they are connected by an internal equivalence. In this section we formalize the definition of internal equivalence and the related notions ``equivalence map'' and ``equivalent objects'', and we establish some basic facts about these notions. \ subsection "Definition of Equivalence" text \ The following locale is defined to prove some basic facts about an equivalence (or an adjunction) in a bicategory that are ``syntactic'' in the sense that they depend only on the configuration (source, target, domain, codomain) of the arrows involved and not on further properties such as the triangle identities (for adjunctions) or assumptions about invertibility (for equivalences). Proofs about adjunctions and equivalences become more automatic once we have introduction and simplification rules in place about this syntax. \ locale adjunction_data_in_bicategory = bicategory + fixes f :: 'a and g :: 'a and \ :: 'a and \ :: 'a assumes ide_left [simp]: "ide f" and ide_right [simp]: "ide g" and unit_in_vhom: "\\ : src f \ g \ f\" and counit_in_vhom: "\\ : f \ g \ src g\" begin (* * TODO: It is difficult to orient the following equations to make them useful as * default simplifications, so they have to be cited explicitly where they are used. *) lemma antipar (*[simp]*): shows "trg g = src f" and "src g = trg f" apply (metis counit_in_vhom hseqE ideD(1) ide_right src.preserves_reflects_arr vconn_implies_hpar(3)) by (metis arrI not_arr_null seq_if_composable src.preserves_reflects_arr unit_in_vhom vconn_implies_hpar(1) vconn_implies_hpar(3)) lemma counit_in_hom [intro]: shows "\\ : trg f \ trg f\" and "\\ : f \ g \ trg f\" using counit_in_vhom vconn_implies_hpar antipar by auto lemma unit_in_hom [intro]: shows "\\ : src f \ src f\" and "\\ : src f \ g \ f\" using unit_in_vhom vconn_implies_hpar antipar by auto lemma unit_simps [simp]: shows "arr \" and "dom \ = src f" and "cod \ = g \ f" and "src \ = src f" and "trg \ = src f" using unit_in_hom antipar by auto lemma counit_simps [simp]: shows "arr \" and "dom \ = f \ g" and "cod \ = trg f" and "src \ = trg f" and "trg \ = trg f" using counit_in_hom antipar by auto text \ The expressions found in the triangle identities for an adjunction come up relatively frequently, so it is useful to have established some basic facts about them, even if the triangle identities themselves have not actually been introduced as assumptions in the current context. \ lemma triangle_in_hom: shows "\(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) : f \ src f \ trg f \ f\" and "\(g \ \) \ \[g, f, g] \ (\ \ g) : trg g \ g \ g \ src g\" and "\\[f] \ (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ \\<^sup>-\<^sup>1[f] : f \ f\" and "\\[g] \ (g \ \) \ \[g, f, g] \ (\ \ g) \ \\<^sup>-\<^sup>1[g] : g \ g\" using antipar by auto lemma triangle_equiv_form: shows "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f] \ \[f] \ (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ \\<^sup>-\<^sup>1[f] = f" and "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g] \ \[g] \ (g \ \) \ \[g, f, g] \ (\ \ g) \ \\<^sup>-\<^sup>1[g] = g" proof - show "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f] \ \[f] \ (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ \\<^sup>-\<^sup>1[f] = f" proof assume 1: "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" have "\[f] \ (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ \\<^sup>-\<^sup>1[f] = \[f] \ ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) \ \\<^sup>-\<^sup>1[f]" using comp_assoc by simp also have "... = \[f] \ (\\<^sup>-\<^sup>1[f] \ \[f]) \ \\<^sup>-\<^sup>1[f]" using 1 by simp also have "... = f" using comp_assoc comp_arr_inv' comp_inv_arr' iso_lunit iso_runit comp_arr_dom comp_cod_arr by simp finally show "\[f] \ (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ \\<^sup>-\<^sup>1[f] = f" by simp next assume 2: "\[f] \ (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) \ \\<^sup>-\<^sup>1[f] = f" have "\\<^sup>-\<^sup>1[f] \ \[f] = \\<^sup>-\<^sup>1[f] \ f \ \[f]" using comp_cod_arr by simp also have "... = (\\<^sup>-\<^sup>1[f] \ \[f]) \ ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) \ (\\<^sup>-\<^sup>1[f] \ \[f])" using 2 comp_assoc by (metis (no_types, lifting)) also have "... = (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)" using comp_arr_inv' comp_inv_arr' iso_lunit iso_runit comp_arr_dom comp_cod_arr hseqI' antipar by (metis ide_left in_homE lunit_simps(4) runit_simps(4) triangle_in_hom(1)) finally show "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" by simp qed show "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g] \ \[g] \ (g \ \) \ \[g, f, g] \ (\ \ g) \ \\<^sup>-\<^sup>1[g] = g" proof assume 1: "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" have "\[g] \ (g \ \) \ \[g, f, g] \ (\ \ g) \ \\<^sup>-\<^sup>1[g] = \[g] \ ((g \ \) \ \[g, f, g] \ (\ \ g)) \ \\<^sup>-\<^sup>1[g]" using comp_assoc by simp also have "... = \[g] \ (\\<^sup>-\<^sup>1[g] \ \[g]) \ \\<^sup>-\<^sup>1[g]" using 1 by simp also have "... = g" using comp_assoc comp_arr_inv' comp_inv_arr' iso_lunit iso_runit comp_arr_dom comp_cod_arr by simp finally show "\[g] \ (g \ \) \ \[g, f, g] \ (\ \ g) \ \\<^sup>-\<^sup>1[g] = g" by simp next assume 2: "\[g] \ (g \ \) \ \[g, f, g] \ (\ \ g) \ \\<^sup>-\<^sup>1[g] = g" have "\\<^sup>-\<^sup>1[g] \ \[g] = \\<^sup>-\<^sup>1[g] \ g \ \[g]" using comp_cod_arr by simp also have "... = \\<^sup>-\<^sup>1[g] \ (\[g] \ (g \ \) \ \[g, f, g] \ (\ \ g) \ \\<^sup>-\<^sup>1[g]) \ \[g]" using 2 by simp also have "... = (\\<^sup>-\<^sup>1[g] \ \[g]) \ ((g \ \) \ \[g, f, g] \ (\ \ g)) \ (\\<^sup>-\<^sup>1[g] \ \[g])" using comp_assoc by simp also have "... = (g \ \) \ \[g, f, g] \ (\ \ g)" using comp_arr_inv' comp_inv_arr' iso_lunit iso_runit comp_arr_dom comp_cod_arr hseqI' antipar by (metis ide_right in_homE lunit_simps(4) runit_simps(4) triangle_in_hom(2)) finally show "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" by simp qed qed end locale equivalence_in_bicategory = adjunction_data_in_bicategory + assumes unit_is_iso [simp]: "iso \" and counit_is_iso [simp]: "iso \" begin lemma dual_equivalence: shows "equivalence_in_bicategory V H \ \ src trg g f (inv \) (inv \)" using antipar by unfold_locales auto end abbreviation (in bicategory) internal_equivalence where "internal_equivalence f g \ \ \ equivalence_in_bicategory V H \ \ src trg f g \ \" subsection "Quasi-Inverses and Equivalence Maps" text \ Antiparallel 1-cells \f\ and \g\ are \emph{quasi-inverses} if they can be extended to an internal equivalence. We will use the term \emph{equivalence map} to refer to a 1-cell that has a quasi-inverse. \ context bicategory begin definition quasi_inverses where "quasi_inverses f g \ \\ \. internal_equivalence f g \ \" lemma quasi_inversesI: assumes "ide f" and "ide g" and "src f \ g \ f" and "f \ g \ trg f" shows "quasi_inverses f g" proof (unfold quasi_inverses_def) have 1: "src g = trg f" using assms ideD(1) isomorphic_implies_ide(2) by blast obtain \ where \: "\\ : src f \ g \ f\ \ iso \" using assms isomorphic_def by auto obtain \ where \: "\\ : f \ g \ trg f\ \ iso \" using assms isomorphic_def by auto have "equivalence_in_bicategory V H \ \ src trg f g \ \" using assms 1 \ \ by unfold_locales auto thus "\\ \. internal_equivalence f g \ \" by auto qed lemma quasi_inversesE: assumes "quasi_inverses f g" and "\ide f; ide g; src f \ g \ f; f \ g \ trg f\ \ T" shows T proof - obtain \ \ where \\: "internal_equivalence f g \ \" using assms quasi_inverses_def by auto interpret \\: equivalence_in_bicategory V H \ \ src trg f g \ \ using \\ by simp have "ide f \ ide g" by simp moreover have "src f \ g \ f" using isomorphic_def \\.unit_in_hom by auto moreover have "f \ g \ trg f" using isomorphic_def \\.counit_in_hom by auto ultimately show T using assms by blast qed lemma quasi_inverse_unique: assumes "quasi_inverses f g" and "quasi_inverses f g'" shows "isomorphic g g'" proof - obtain \ \ where \\: "internal_equivalence f g \ \" using assms quasi_inverses_def by auto interpret \\: equivalence_in_bicategory V H \ \ src trg f g \ \ using \\ by simp obtain \' \' where \'\': "internal_equivalence f g' \' \'" using assms quasi_inverses_def by auto interpret \'\': equivalence_in_bicategory V H \ \ src trg f g' \' \' using \'\' by simp have "\\[g'] \ (g' \ \) \ \[g', f, g] \ (\' \ g) \ \\<^sup>-\<^sup>1[g] : g \ g'\" using \\.unit_in_hom \\.unit_is_iso \\.antipar \'\'.antipar by (intro comp_in_homI' hseqI') auto moreover have "iso (\[g'] \ (g' \ \) \ \[g', f, g] \ (\' \ g) \ \\<^sup>-\<^sup>1[g])" using \\.unit_in_hom \\.unit_is_iso \\.antipar \'\'.antipar by (intro isos_compose) auto ultimately show ?thesis using isomorphic_def by auto qed lemma quasi_inverses_symmetric: assumes "quasi_inverses f g" shows "quasi_inverses g f" using assms quasi_inverses_def equivalence_in_bicategory.dual_equivalence by metis definition equivalence_map where "equivalence_map f \ \g \ \. equivalence_in_bicategory V H \ \ src trg f g \ \" lemma equivalence_mapI: assumes "quasi_inverses f g" shows "equivalence_map f" using assms quasi_inverses_def equivalence_map_def by auto lemma equivalence_mapE: assumes "equivalence_map f" obtains g where "quasi_inverses f g" using assms equivalence_map_def quasi_inverses_def by auto lemma equivalence_map_is_ide: assumes "equivalence_map f" shows "ide f" using assms adjunction_data_in_bicategory.ide_left equivalence_in_bicategory_def equivalence_map_def by fastforce lemma obj_is_equivalence_map: assumes "obj a" shows "equivalence_map a" using assms by (metis equivalence_mapI isomorphic_symmetric objE obj_self_composable(2) quasi_inversesI) lemma equivalence_respects_iso: assumes "equivalence_in_bicategory V H \ \ src trg f g \ \" and "\\ : f \ f'\" and "iso \" and "\\ : g \ g'\" and "iso \" shows "internal_equivalence f' g' ((g' \ \) \ (\ \ f) \ \) (\ \ (inv \ \ g) \ (f' \ inv \))" proof - interpret E: equivalence_in_bicategory V H \ \ src trg f g \ \ using assms by auto show ?thesis proof show f': "ide f'" using assms by auto show g': "ide g'" using assms by auto show 1: "\(g' \ \) \ (\ \ f) \ \ : src f' \ g' \ f'\" using assms f' g' E.unit_in_hom E.antipar(2) vconn_implies_hpar(3) apply (intro comp_in_homI) apply auto by (intro hcomp_in_vhom) auto show "iso ((g' \ \) \ (\ \ f) \ \)" using assms 1 g' vconn_implies_hpar(3) E.antipar(2) iso_hcomp by (intro isos_compose) auto show 1: "\\ \ (inv \ \ g) \ (f' \ inv \) : f' \ g' \ src g'\" using assms f' ide_in_hom(2) vconn_implies_hpar(3-4) E.antipar(1-2) by (intro comp_in_homI) auto show "iso (\ \ (inv \ \ g) \ (f' \ inv \))" using assms 1 isos_compose by (metis E.counit_is_iso E.ide_right arrI f' hseqE ide_is_iso iso_hcomp iso_inv_iso seqE) qed qed lemma equivalence_map_preserved_by_iso: assumes "equivalence_map f" and "f \ f'" shows "equivalence_map f'" proof - obtain g \ \ where E: "equivalence_in_bicategory V H \ \ src trg f g \ \" using assms equivalence_map_def by auto interpret E: equivalence_in_bicategory V H \ \ src trg f g \ \ using E by auto obtain \ where \: "\\ : f \ f'\ \ iso \" using assms isomorphic_def isomorphic_symmetric by blast have "equivalence_in_bicategory V H \ \ src trg f' g ((g \ \) \ (g \ f) \ \) (\ \ (inv \ \ g) \ (f' \ inv g))" using E \ equivalence_respects_iso [of f g \ \ \ f' g g] by fastforce thus ?thesis using equivalence_map_def by auto qed lemma equivalence_preserved_by_iso_right: assumes "equivalence_in_bicategory V H \ \ src trg f g \ \" and "\\ : g \ g'\" and "iso \" shows "equivalence_in_bicategory V H \ \ src trg f g' ((\ \ f) \ \) (\ \ (f \ inv \))" proof interpret E: equivalence_in_bicategory V H \ \ src trg f g \ \ using assms by auto show "ide f" by simp show "ide g'" using assms(2) isomorphic_def by auto show "\(\ \ f) \ \ : src f \ g' \ f\" using assms E.antipar(2) E.ide_left by blast show "\\ \ (f \ inv \) : f \ g' \ src g'\" using assms vconn_implies_hpar(3-4) E.counit_in_vhom E.antipar(1) ide_in_hom(2) by (intro comp_in_homI, auto) show "iso ((\ \ f) \ \)" using assms E.antipar isos_compose by auto show "iso (\ \ (f \ inv \))" using assms E.antipar isos_compose by auto qed lemma equivalence_preserved_by_iso_left: assumes "equivalence_in_bicategory V H \ \ src trg f g \ \" and "\\ : f \ f'\" and "iso \" shows "equivalence_in_bicategory V H \ \ src trg f' g ((g \ \) \ \) (\ \ (inv \ \ g))" proof interpret E: equivalence_in_bicategory V H \ \ src trg f g \ \ using assms by auto show "ide f'" using assms by auto show "ide g" by simp show "\(g \ \) \ \ : src f' \ g \ f'\" using assms E.unit_in_hom E.antipar by auto show "\\ \ (inv \ \ g) : f' \ g \ src g\" using assms E.counit_in_hom E.antipar ide_in_hom(2) vconn_implies_hpar(3) by auto show "iso ((g \ \) \ \)" using assms E.antipar isos_compose by auto show "iso (\ \ (inv \ \ g))" using assms E.antipar isos_compose by auto qed definition some_quasi_inverse where "some_quasi_inverse f = (SOME g. quasi_inverses f g)" notation some_quasi_inverse ("_\<^sup>~" [1000] 1000) lemma quasi_inverses_some_quasi_inverse: assumes "equivalence_map f" shows "quasi_inverses f f\<^sup>~" and "quasi_inverses f\<^sup>~ f" using assms some_quasi_inverse_def quasi_inverses_def equivalence_map_def someI_ex [of "\g. quasi_inverses f g"] quasi_inverses_symmetric by auto lemma quasi_inverse_antipar: assumes "equivalence_map f" shows "src f\<^sup>~ = trg f" and "trg f\<^sup>~ = src f" proof - obtain \ \ where \\: "equivalence_in_bicategory V H \ \ src trg f f\<^sup>~ \ \" using assms equivalence_map_def quasi_inverses_some_quasi_inverse quasi_inverses_def by blast interpret \\: equivalence_in_bicategory V H \ \ src trg f \f\<^sup>~\ \ \ using \\ by simp show "src f\<^sup>~ = trg f" using \\.antipar by simp show "trg f\<^sup>~ = src f" using \\.antipar by simp qed lemma quasi_inverse_in_hom [intro]: assumes "equivalence_map f" shows "\f\<^sup>~ : trg f \ src f\" and "\f\<^sup>~ : f\<^sup>~ \ f\<^sup>~\" using assms equivalence_mapE apply (intro in_homI in_hhomI) apply (metis equivalence_map_is_ide ideD(1) not_arr_null quasi_inverse_antipar(2) src.preserves_ide trg.is_extensional) apply (simp_all add: quasi_inverse_antipar) using assms quasi_inversesE quasi_inverses_some_quasi_inverse(2) by blast lemma quasi_inverse_simps [simp]: assumes "equivalence_map f" shows "equivalence_map f\<^sup>~" and "ide f\<^sup>~" and "src f\<^sup>~ = trg f" and "trg f\<^sup>~ = src f" and "dom f\<^sup>~ = f\<^sup>~" and "cod f\<^sup>~ = f\<^sup>~" - using assms equivalence_mapE quasi_inverse_in_hom quasi_inverses_some_quasi_inverse equivalence_map_is_ide + using assms equivalence_mapE quasi_inverse_in_hom quasi_inverses_some_quasi_inverse + equivalence_map_is_ide apply auto by (meson equivalence_mapI) lemma quasi_inverse_quasi_inverse: assumes "equivalence_map f" shows "(f\<^sup>~)\<^sup>~ \ f" proof - have "quasi_inverses f\<^sup>~ (f\<^sup>~)\<^sup>~" using assms quasi_inverses_some_quasi_inverse by simp moreover have "quasi_inverses f\<^sup>~ f" using assms quasi_inverses_some_quasi_inverse quasi_inverses_symmetric by simp ultimately show ?thesis using quasi_inverse_unique by simp qed lemma comp_quasi_inverse: assumes "equivalence_map f" shows "f\<^sup>~ \ f \ src f" and "f \ f\<^sup>~ \ trg f" proof - obtain \ \ where \\: "equivalence_in_bicategory V H \ \ src trg f f\<^sup>~ \ \" using assms equivalence_map_def quasi_inverses_some_quasi_inverse quasi_inverses_def by blast interpret \\: equivalence_in_bicategory V H \ \ src trg f \f\<^sup>~\ \ \ using \\ by simp show "f\<^sup>~ \ f \ src f" using quasi_inverses_some_quasi_inverse quasi_inverses_def \\.unit_in_hom \\.unit_is_iso isomorphic_def by blast show "f \ f\<^sup>~ \ trg f" using quasi_inverses_some_quasi_inverse quasi_inverses_def \\.counit_in_hom \\.counit_is_iso isomorphic_def by blast qed lemma quasi_inverse_transpose: assumes "ide f" and "ide g" and "ide h" and "f \ g \ h" shows "equivalence_map g \ f \ h \ g\<^sup>~" and "equivalence_map f \ g \ f\<^sup>~ \ h" proof - have 1: "src f = trg g" using assms equivalence_map_is_ide by fastforce show "equivalence_map g \ f \ h \ g\<^sup>~" proof - assume g: "equivalence_map g" have 2: "ide g\<^sup>~" using g by simp have "f \ f \ src f" using assms isomorphic_unit_right isomorphic_symmetric by blast also have "... \ f \ trg g" using assms 1 isomorphic_reflexive by auto also have "... \ f \ g \ g\<^sup>~" using assms g 1 comp_quasi_inverse(2) isomorphic_symmetric hcomp_ide_isomorphic by simp also have "... \ (f \ g) \ g\<^sup>~" using assms g 1 2 assoc'_in_hom [of f g "g\<^sup>~"] iso_assoc' isomorphic_def by auto also have "... \ h \ g\<^sup>~" using assms g 1 2 by (simp add: hcomp_isomorphic_ide) finally show ?thesis by blast qed show "equivalence_map f \ g \ f\<^sup>~ \ h" proof - assume f: "equivalence_map f" have 2: "ide f\<^sup>~" using f by simp have "g \ src f \ g" using assms 1 isomorphic_unit_left isomorphic_symmetric by metis also have "... \ (f\<^sup>~ \ f) \ g" using assms f 1 comp_quasi_inverse(1) [of f] isomorphic_symmetric hcomp_isomorphic_ide by simp also have "... \ f\<^sup>~ \ f \ g" using assms f 1 assoc_in_hom [of "f\<^sup>~" f g] iso_assoc isomorphic_def by auto also have "... \ f\<^sup>~ \ h" using assms f 1 equivalence_map_is_ide quasi_inverses_some_quasi_inverse hcomp_ide_isomorphic by simp finally show ?thesis by blast qed qed end subsection "Composing Equivalences" locale composite_equivalence_in_bicategory = bicategory V H \ \ src trg + fg: equivalence_in_bicategory V H \ \ src trg f g \ \ + hk: equivalence_in_bicategory V H \ \ src trg h k \ \ for V :: "'a \ 'a \ 'a" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and f :: "'a" and g :: "'a" and \ :: "'a" and \ :: "'a" and h :: "'a" and k :: "'a" and \ :: "'a" and \ :: "'a" + assumes composable: "src h = trg f" begin abbreviation \ where "\ \ \\<^sup>-\<^sup>1[g, k, h \ f] \ (g \ \[k, h, f]) \ (g \ \ \ f) \ (g \ \\<^sup>-\<^sup>1[f]) \ \" abbreviation \ where "\ \ \ \ (h \ \[k]) \ (h \ \ \ k) \ (h \ \\<^sup>-\<^sup>1[f, g, k]) \ \[h, f, g \ k]" interpretation adjunction_data_in_bicategory V H \ \ src trg \h \ f\ \g \ k\ \ \ proof show "ide (h \ f)" using composable by simp show "ide (g \ k)" using fg.antipar hk.antipar composable by simp show "\\ : src (h \ f) \ (g \ k) \ h \ f\" using fg.antipar hk.antipar composable by fastforce show "\\ : (h \ f) \ g \ k \ src (g \ k)\" using fg.antipar hk.antipar composable by fastforce qed interpretation equivalence_in_bicategory V H \ \ src trg \h \ f\ \g \ k\ \ \ proof show "iso \" using fg.antipar hk.antipar composable fg.unit_is_iso hk.unit_is_iso iso_hcomp iso_lunit' hseq_char by (intro isos_compose, auto) show "iso \" using fg.antipar hk.antipar composable fg.counit_is_iso hk.counit_is_iso iso_hcomp iso_lunit hseq_char by (intro isos_compose, auto) qed lemma is_equivalence: shows "equivalence_in_bicategory V H \ \ src trg (h \ f) (g \ k) \ \" .. sublocale equivalence_in_bicategory V H \ \ src trg \h \ f\ \g \ k\ \ \ using is_equivalence by simp end context bicategory begin lemma equivalence_maps_compose: assumes "equivalence_map f" and "equivalence_map f'" and "src f' = trg f" shows "equivalence_map (f' \ f)" proof - obtain g \ \ where fg: "equivalence_in_bicategory V H \ \ src trg f g \ \" using assms(1) equivalence_map_def by auto interpret fg: equivalence_in_bicategory V H \ \ src trg f g \ \ using fg by auto obtain g' \' \' where f'g': "equivalence_in_bicategory V H \ \ src trg f' g' \' \'" using assms(2) equivalence_map_def by auto interpret f'g': equivalence_in_bicategory V H \ \ src trg f' g' \' \' using f'g' by auto interpret composite_equivalence_in_bicategory V H \ \ src trg f g \ \ f' g' \' \' using assms(3) by (unfold_locales, auto) show ?thesis using equivalence_map_def equivalence_in_bicategory_axioms by auto qed lemma quasi_inverse_hcomp': assumes "equivalence_map f" and "equivalence_map f'" and "equivalence_map (f \ f')" and "quasi_inverses f g" and "quasi_inverses f' g'" shows "quasi_inverses (f \ f') (g' \ g)" proof - obtain \ \ where g: "internal_equivalence f g \ \" using assms quasi_inverses_def by auto interpret g: equivalence_in_bicategory V H \ \ src trg f g \ \ using g by simp obtain \' \' where g': "internal_equivalence f' g' \' \'" using assms quasi_inverses_def by auto interpret g': equivalence_in_bicategory V H \ \ src trg f' g' \' \' using g' by simp interpret gg': composite_equivalence_in_bicategory V H \ \ src trg f' g' \' \' f g \ \ using assms equivalence_map_is_ide [of "f \ f'"] apply unfold_locales using ideD(1) by fastforce show ?thesis unfolding quasi_inverses_def using gg'.equivalence_in_bicategory_axioms by auto qed lemma quasi_inverse_hcomp: assumes "equivalence_map f" and "equivalence_map f'" and "equivalence_map (f \ f')" shows "(f \ f')\<^sup>~ \ f'\<^sup>~ \ f\<^sup>~" using assms quasi_inverses_some_quasi_inverse quasi_inverse_hcomp' quasi_inverse_unique by metis lemma quasi_inverse_respects_isomorphic: assumes "equivalence_map f" and "equivalence_map f'" and "f \ f'" shows "f\<^sup>~ \ f'\<^sup>~" proof - have hpar: "src f = src f' \ trg f = trg f'" using assms isomorphic_implies_hpar by simp have "f\<^sup>~ \ f\<^sup>~ \ trg f" using isomorphic_unit_right by (metis assms(1) isomorphic_symmetric quasi_inverse_simps(2-3)) also have "... \ f\<^sup>~ \ f' \ f'\<^sup>~" proof - have "trg f \ f' \ f'\<^sup>~" using assms quasi_inverse_hcomp by (simp add: comp_quasi_inverse(2) hpar isomorphic_symmetric) thus ?thesis using assms hpar hcomp_ide_isomorphic isomorphic_implies_hpar(2) by auto qed also have "... \ (f\<^sup>~ \ f') \ f'\<^sup>~" using assms hcomp_assoc_isomorphic hpar isomorphic_implies_ide(2) isomorphic_symmetric by auto also have "... \ (f\<^sup>~ \ f) \ f'\<^sup>~" proof - have "f\<^sup>~ \ f' \ f\<^sup>~ \ f" using assms isomorphic_symmetric hcomp_ide_isomorphic isomorphic_implies_hpar(1) by auto thus ?thesis using assms hcomp_isomorphic_ide isomorphic_implies_hpar(1) by auto qed also have "... \ src f \ f'\<^sup>~" proof - have "f\<^sup>~ \ f \ src f" using assms comp_quasi_inverse by simp thus ?thesis using assms hcomp_isomorphic_ide isomorphic_implies_hpar by simp qed also have "... \ f'\<^sup>~" using assms isomorphic_unit_left by (metis hpar quasi_inverse_simps(2,4)) finally show ?thesis by blast qed end subsection "Equivalent Objects" context bicategory begin definition equivalent_objects where "equivalent_objects a b \ \f. \f : a \ b\ \ equivalence_map f" lemma equivalent_objects_reflexive: assumes "obj a" shows "equivalent_objects a a" using assms by (metis equivalent_objects_def ide_in_hom(1) objE obj_is_equivalence_map) lemma equivalent_objects_symmetric: assumes "equivalent_objects a b" shows "equivalent_objects b a" using assms by (metis equivalent_objects_def in_hhomE quasi_inverse_in_hom(1) quasi_inverse_simps(1)) lemma equivalent_objects_transitive [trans]: assumes "equivalent_objects a b" and "equivalent_objects b c" shows "equivalent_objects a c" proof - obtain f where f: "\f : a \ b\ \ equivalence_map f" using assms equivalent_objects_def by auto obtain g where g: "\g : b \ c\ \ equivalence_map g" using assms equivalent_objects_def by auto have "\g \ f : a \ c\ \ equivalence_map (g \ f)" using f g equivalence_maps_compose by auto thus ?thesis using equivalent_objects_def by auto qed end subsection "Transporting Arrows along Equivalences" text \ We show in this section that transporting the arrows of one hom-category to another along connecting equivalence maps yields an equivalence of categories. This is useful, because it seems otherwise hard to establish that the transporting functor is full. \ locale two_equivalences_in_bicategory = bicategory V H \ \ src trg + e\<^sub>0: equivalence_in_bicategory V H \ \ src trg e\<^sub>0 d\<^sub>0 \\<^sub>0 \\<^sub>0 + e\<^sub>1: equivalence_in_bicategory V H \ \ src trg e\<^sub>1 d\<^sub>1 \\<^sub>1 \\<^sub>1 for V :: "'a \ 'a \ 'a" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src :: "'a \ 'a" and trg :: "'a \ 'a" and e\<^sub>0 :: "'a" and d\<^sub>0 :: "'a" and \\<^sub>0 :: "'a" and \\<^sub>0 :: "'a" and e\<^sub>1 :: "'a" and d\<^sub>1 :: "'a" and \\<^sub>1 :: "'a" and \\<^sub>1 :: "'a" begin interpretation hom: subcategory V \\\. \\ : src e\<^sub>0 \ src e\<^sub>1\\ using hhom_is_subcategory by simp (* TODO: The preceding interpretation somehow brings in unwanted notation. *) no_notation in_hom ("\_ : _ \ _\") interpretation hom': subcategory V \\\. \\ : trg e\<^sub>0 \ trg e\<^sub>1\\ using hhom_is_subcategory by simp no_notation in_hom ("\_ : _ \ _\") abbreviation (input) F where "F \ \\. e\<^sub>1 \ \ \ d\<^sub>0" interpretation F: "functor" hom.comp hom'.comp F proof show 1: "\f. hom.arr f \ hom'.arr (e\<^sub>1 \ f \ d\<^sub>0)" using hom.arr_char hom'.arr_char in_hhom_def e\<^sub>0.antipar(1-2) by simp show "\f. \ hom.arr f \ e\<^sub>1 \ f \ d\<^sub>0 = hom'.null" using 1 hom.arr_char hom'.null_char in_hhom_def by (metis e\<^sub>0.antipar(1) hseqE hseq_char' hcomp_simps(2)) show "\f. hom.arr f \ hom'.dom (e\<^sub>1 \ f \ d\<^sub>0) = e\<^sub>1 \ hom.dom f \ d\<^sub>0" using hom.arr_char hom.dom_char hom'.arr_char hom'.dom_char by (metis 1 hcomp_simps(3) e\<^sub>0.ide_right e\<^sub>1.ide_left hom'.inclusion hseq_char ide_char) show "\f. hom.arr f \ hom'.cod (e\<^sub>1 \ f \ d\<^sub>0) = e\<^sub>1 \ hom.cod f \ d\<^sub>0" using hom.arr_char hom.cod_char hom'.arr_char hom'.cod_char by (metis 1 hcomp_simps(4) e\<^sub>0.ide_right e\<^sub>1.ide_left hom'.inclusion hseq_char ide_char) show "\g f. hom.seq g f \ e\<^sub>1 \ hom.comp g f \ d\<^sub>0 = hom'.comp (e\<^sub>1 \ g \ d\<^sub>0) (e\<^sub>1 \ f \ d\<^sub>0)" using 1 hom.seq_char hom.arr_char hom.comp_char hom'.arr_char hom'.comp_char whisker_left [of e\<^sub>1] whisker_right [of d\<^sub>0] apply auto apply (metis hseq_char seqE src_vcomp) by (metis hseq_char hseq_char') qed abbreviation (input) G where "G \ \\'. d\<^sub>1 \ \' \ e\<^sub>0" interpretation G: "functor" hom'.comp hom.comp G proof show 1: "\f. hom'.arr f \ hom.arr (d\<^sub>1 \ f \ e\<^sub>0)" using hom.arr_char hom'.arr_char in_hhom_def e\<^sub>1.antipar(1) e\<^sub>1.antipar(2) by simp show "\f. \ hom'.arr f \ d\<^sub>1 \ f \ e\<^sub>0 = hom.null" using 1 hom.arr_char hom'.null_char in_hhom_def by (metis e\<^sub>1.antipar(2) hom'.arrI hom.null_char hseqE hseq_char' hcomp_simps(2)) show "\f. hom'.arr f \ hom.dom (d\<^sub>1 \ f \ e\<^sub>0) = d\<^sub>1 \ hom'.dom f \ e\<^sub>0" using 1 hom.arr_char hom.dom_char hom'.arr_char hom'.dom_char by (metis hcomp_simps(3) e\<^sub>0.ide_left e\<^sub>1.ide_right hom.inclusion hseq_char ide_char) show "\f. hom'.arr f \ hom.cod (d\<^sub>1 \ f \ e\<^sub>0) = d\<^sub>1 \ hom'.cod f \ e\<^sub>0" using 1 hom.arr_char hom.cod_char hom'.arr_char hom'.cod_char by (metis hcomp_simps(4) e\<^sub>0.ide_left e\<^sub>1.ide_right hom.inclusion hseq_char ide_char) show "\g f. hom'.seq g f \ d\<^sub>1 \ hom'.comp g f \ e\<^sub>0 = hom.comp (d\<^sub>1 \ g \ e\<^sub>0) (d\<^sub>1 \ f \ e\<^sub>0)" using 1 hom'.seq_char hom'.arr_char hom'.comp_char hom.arr_char hom.comp_char whisker_left [of d\<^sub>1] whisker_right [of e\<^sub>0] apply auto apply (metis hseq_char seqE src_vcomp) by (metis hseq_char hseq_char') qed interpretation GF: composite_functor hom.comp hom'.comp hom.comp F G .. interpretation FG: composite_functor hom'.comp hom.comp hom'.comp G F .. abbreviation (input) \\<^sub>0 where "\\<^sub>0 f \ (d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, f \ d\<^sub>0, e\<^sub>0]) \ \[d\<^sub>1, e\<^sub>1, (f \ d\<^sub>0) \ e\<^sub>0] \ ((d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[f, d\<^sub>0, e\<^sub>0]) \ (\\<^sub>1 \ f \ \\<^sub>0) \ \\<^sup>-\<^sup>1[f \ src e\<^sub>0] \ \\<^sup>-\<^sup>1[f]" lemma \\<^sub>0_in_hom: assumes "\f : src e\<^sub>0 \ src e\<^sub>1\" and "ide f" shows "\\\<^sub>0 f : src e\<^sub>0 \ src e\<^sub>1\" and "\\\<^sub>0 f : f \ d\<^sub>1 \ (e\<^sub>1 \ f \ d\<^sub>0) \ e\<^sub>0\" proof - show "\\\<^sub>0 f : f \ d\<^sub>1 \ (e\<^sub>1 \ f \ d\<^sub>0) \ e\<^sub>0\" using assms e\<^sub>0.antipar e\<^sub>1.antipar by fastforce thus "\\\<^sub>0 f : src e\<^sub>0 \ src e\<^sub>1\" using assms src_dom [of "\\<^sub>0 f"] trg_dom [of "\\<^sub>0 f"] by (metis arrI dom_comp in_hhom_def runit'_simps(4) seqE) qed lemma iso_\\<^sub>0: assumes "\f : src e\<^sub>0 \ src e\<^sub>1\" and "ide f" shows "iso (\\<^sub>0 f)" using assms iso_lunit' iso_runit' e\<^sub>0.antipar e\<^sub>1.antipar by (intro isos_compose) auto interpretation \: transformation_by_components hom.comp hom.comp hom.map \G o F\ \\<^sub>0 proof fix f assume f: "hom.ide f" show "hom.in_hom (\\<^sub>0 f) (hom.map f) (GF.map f)" proof (unfold hom.in_hom_char, intro conjI) show "hom.arr (hom.map f)" using f by simp show "hom.arr (GF.map f)" using f by simp show "hom.arr (\\<^sub>0 f)" using f hom.ide_char hom.arr_char \\<^sub>0_in_hom by simp show "\\\<^sub>0 f : hom.map f \ GF.map f\" using f hom.ide_char hom.arr_char hom'.ide_char hom'.arr_char F.preserves_arr \\<^sub>0_in_hom by auto qed next fix \ assume \: "hom.arr \" show "hom.comp (\\<^sub>0 (hom.cod \)) (hom.map \) = hom.comp (GF.map \) (\\<^sub>0 (hom.dom \))" proof - have "hom.comp (\\<^sub>0 (hom.cod \)) (hom.map \) = \\<^sub>0 (cod \) \ \" proof - have "hom.map \ = \" using \ by simp moreover have "\\<^sub>0 (hom.cod \) = \\<^sub>0 (cod \)" using \ hom.arr_char hom.cod_char by simp moreover have "hom.arr (\\<^sub>0 (cod \))" using \ hom.arr_char \\<^sub>0_in_hom [of "cod \"] by (metis hom.arr_cod hom.cod_simp ide_cod in_hhom_def) moreover have "seq (\\<^sub>0 (cod \)) \" proof show 1: "\\ : dom \ \ cod \\" using \ hom.arr_char by auto show "\\\<^sub>0 (cod \) : cod \ \ d\<^sub>1 \ (e\<^sub>1 \ cod \ \ d\<^sub>0) \ e\<^sub>0\" using \ hom.arr_char \\<^sub>0_in_hom by (metis 1 arrI hom.arr_cod hom.cod_simp ide_cod) qed ultimately show ?thesis using \ hom.comp_char by simp qed also have "... = (d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, cod \ \ d\<^sub>0, e\<^sub>0]) \ \[d\<^sub>1, e\<^sub>1, (cod \ \ d\<^sub>0) \ e\<^sub>0] \ ((d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[cod \, d\<^sub>0, e\<^sub>0]) \ (\\<^sub>1 \ cod \ \ \\<^sub>0) \ \\<^sup>-\<^sup>1[cod \ \ src e\<^sub>0] \ \\<^sup>-\<^sup>1[cod \] \ \" using \ hom.arr_char comp_assoc by auto also have "... = ((d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, cod \ \ d\<^sub>0, e\<^sub>0]) \ \[d\<^sub>1, e\<^sub>1, (cod \ \ d\<^sub>0) \ e\<^sub>0] \ ((d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[cod \, d\<^sub>0, e\<^sub>0]) \ (\\<^sub>1 \ cod \ \ \\<^sub>0) \ \\<^sup>-\<^sup>1[cod \ \ src e\<^sub>0] \ (\ \ src e\<^sub>0)) \ \\<^sup>-\<^sup>1[dom \]" using \ hom.arr_char runit'_naturality [of \] comp_assoc by auto also have "... = ((d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, cod \ \ d\<^sub>0, e\<^sub>0]) \ \[d\<^sub>1, e\<^sub>1, (cod \ \ d\<^sub>0) \ e\<^sub>0] \ ((d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[cod \, d\<^sub>0, e\<^sub>0]) \ (\\<^sub>1 \ cod \ \ \\<^sub>0) \ (src e\<^sub>1 \ \ \ src e\<^sub>0) \ \\<^sup>-\<^sup>1[dom \ \ src e\<^sub>0]) \ \\<^sup>-\<^sup>1[dom \]" using \ hom.arr_char lunit'_naturality [of "\ \ src e\<^sub>0"] by force also have "... = ((d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, cod \ \ d\<^sub>0, e\<^sub>0]) \ \[d\<^sub>1, e\<^sub>1, (cod \ \ d\<^sub>0) \ e\<^sub>0] \ ((d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[cod \, d\<^sub>0, e\<^sub>0]) \ (\\<^sub>1 \ cod \ \ \\<^sub>0) \ (src e\<^sub>1 \ \ \ src e\<^sub>0)) \ \\<^sup>-\<^sup>1[dom \ \ src e\<^sub>0] \ \\<^sup>-\<^sup>1[dom \]" using comp_assoc by simp also have "... = ((d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, cod \ \ d\<^sub>0, e\<^sub>0]) \ \[d\<^sub>1, e\<^sub>1, (cod \ \ d\<^sub>0) \ e\<^sub>0] \ ((d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[cod \, d\<^sub>0, e\<^sub>0]) \ ((d\<^sub>1 \ e\<^sub>1) \ \ \ d\<^sub>0 \ e\<^sub>0)) \ (\\<^sub>1 \ dom \ \ \\<^sub>0) \ \\<^sup>-\<^sup>1[dom \ \ src e\<^sub>0] \ \\<^sup>-\<^sup>1[dom \]" proof - have "(\\<^sub>1 \ cod \ \ \\<^sub>0) \ (src e\<^sub>1 \ \ \ src e\<^sub>0) = \\<^sub>1 \ \ \ \\<^sub>0" using \ hom.arr_char comp_arr_dom comp_cod_arr interchange [of \\<^sub>1 "src e\<^sub>1" "cod \ \ \\<^sub>0" "\ \ src e\<^sub>0"] interchange [of "cod \" \ \\<^sub>0 "src e\<^sub>0"] by (metis e\<^sub>0.unit_in_hom(1) e\<^sub>0.unit_simps(2) e\<^sub>1.unit_simps(1) e\<^sub>1.unit_simps(2) hseqI' in_hhom_def) also have "... = ((d\<^sub>1 \ e\<^sub>1) \ \ \ d\<^sub>0 \ e\<^sub>0) \ (\\<^sub>1 \ dom \ \ \\<^sub>0)" proof - have "\\<^sub>1 \ \ \ \\<^sub>0 = (d\<^sub>1 \ e\<^sub>1) \ \\<^sub>1 \ \ \ dom \ \ (d\<^sub>0 \ e\<^sub>0) \ \\<^sub>0" using \ hom.arr_char comp_arr_dom comp_cod_arr by auto also have "... = (d\<^sub>1 \ e\<^sub>1) \ \\<^sub>1 \ (\ \ d\<^sub>0 \ e\<^sub>0) \ (dom \ \ \\<^sub>0)" using \ hom.arr_char comp_cod_arr interchange [of \ "dom \" "d\<^sub>0 \ e\<^sub>0" \\<^sub>0] by auto also have "... = ((d\<^sub>1 \ e\<^sub>1) \ \ \ d\<^sub>0 \ e\<^sub>0) \ (\\<^sub>1 \ dom \ \ \\<^sub>0)" using \ hom.arr_char comp_arr_dom comp_cod_arr interchange [of "d\<^sub>1 \ e\<^sub>1" \\<^sub>1 "\ \ d\<^sub>0 \ e\<^sub>0" "dom \ \ \\<^sub>0"] interchange [of \ "dom \" "d\<^sub>0 \ e\<^sub>0" \\<^sub>0] by (metis e\<^sub>0.unit_in_hom(1) e\<^sub>0.unit_simps(1) e\<^sub>0.unit_simps(3) e\<^sub>1.unit_simps(1) e\<^sub>1.unit_simps(3) hom.inclusion hseqI) finally show ?thesis by simp qed finally have "(\\<^sub>1 \ cod \ \ \\<^sub>0) \ (src e\<^sub>1 \ \ \ src e\<^sub>0) = ((d\<^sub>1 \ e\<^sub>1) \ \ \ d\<^sub>0 \ e\<^sub>0) \ (\\<^sub>1 \ dom \ \ \\<^sub>0)" by simp thus ?thesis using comp_assoc by simp qed also have "... = ((d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, cod \ \ d\<^sub>0, e\<^sub>0]) \ \[d\<^sub>1, e\<^sub>1, (cod \ \ d\<^sub>0) \ e\<^sub>0] \ ((d\<^sub>1 \ e\<^sub>1) \ (\ \ d\<^sub>0) \ e\<^sub>0) \ ((d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[dom \, d\<^sub>0, e\<^sub>0])) \ (\\<^sub>1 \ dom \ \ \\<^sub>0) \ \\<^sup>-\<^sup>1[dom \ \ src e\<^sub>0] \ \\<^sup>-\<^sup>1[dom \]" using \ hom.arr_char e\<^sub>0.antipar e\<^sub>1.antipar assoc'_naturality [of \ d\<^sub>0 e\<^sub>0] whisker_left [of "d\<^sub>1 \ e\<^sub>1" "\\<^sup>-\<^sup>1[cod \, d\<^sub>0, e\<^sub>0]" "\ \ d\<^sub>0 \ e\<^sub>0"] whisker_left [of "d\<^sub>1 \ e\<^sub>1" "(\ \ d\<^sub>0) \ e\<^sub>0" "\\<^sup>-\<^sup>1[dom \, d\<^sub>0, e\<^sub>0]"] by auto also have "... = ((d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, cod \ \ d\<^sub>0, e\<^sub>0]) \ \[d\<^sub>1, e\<^sub>1, (cod \ \ d\<^sub>0) \ e\<^sub>0] \ ((d\<^sub>1 \ e\<^sub>1) \ (\ \ d\<^sub>0) \ e\<^sub>0)) \ ((d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[dom \, d\<^sub>0, e\<^sub>0]) \ (\\<^sub>1 \ dom \ \ \\<^sub>0) \ \\<^sup>-\<^sup>1[dom \ \ src e\<^sub>0] \ \\<^sup>-\<^sup>1[dom \]" using comp_assoc by simp also have "... = ((d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, cod \ \ d\<^sub>0, e\<^sub>0]) \ (d\<^sub>1 \ e\<^sub>1 \ (\ \ d\<^sub>0) \ e\<^sub>0) \ \[d\<^sub>1, e\<^sub>1, (dom \ \ d\<^sub>0) \ e\<^sub>0]) \ ((d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[dom \, d\<^sub>0, e\<^sub>0]) \ (\\<^sub>1 \ dom \ \ \\<^sub>0) \ \\<^sup>-\<^sup>1[dom \ \ src e\<^sub>0] \ \\<^sup>-\<^sup>1[dom \]" using \ hom.arr_char e\<^sub>0.antipar e\<^sub>1.antipar assoc_naturality [of d\<^sub>1 e\<^sub>1 "(\ \ d\<^sub>0) \ e\<^sub>0"] hseqI by auto also have "... = ((d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, cod \ \ d\<^sub>0, e\<^sub>0]) \ (d\<^sub>1 \ e\<^sub>1 \ (\ \ d\<^sub>0) \ e\<^sub>0)) \ \[d\<^sub>1, e\<^sub>1, (dom \ \ d\<^sub>0) \ e\<^sub>0] \ ((d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[dom \, d\<^sub>0, e\<^sub>0]) \ (\\<^sub>1 \ dom \ \ \\<^sub>0) \ \\<^sup>-\<^sup>1[dom \ \ src e\<^sub>0] \ \\<^sup>-\<^sup>1[dom \]" using comp_assoc by simp also have "... = ((d\<^sub>1 \ (e\<^sub>1 \ \ \ d\<^sub>0) \ e\<^sub>0) \ (d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, dom \ \ d\<^sub>0, e\<^sub>0])) \ \[d\<^sub>1, e\<^sub>1, (dom \ \ d\<^sub>0) \ e\<^sub>0] \ ((d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[dom \, d\<^sub>0, e\<^sub>0]) \ (\\<^sub>1 \ dom \ \ \\<^sub>0) \ \\<^sup>-\<^sup>1[dom \ \ src e\<^sub>0] \ \\<^sup>-\<^sup>1[dom \]" using \ hom.arr_char e\<^sub>0.antipar e\<^sub>1.antipar assoc'_naturality [of e\<^sub>1 "\ \ d\<^sub>0" e\<^sub>0] whisker_left [of d\<^sub>1 "\\<^sup>-\<^sup>1[e\<^sub>1, cod \ \ d\<^sub>0, e\<^sub>0]" "e\<^sub>1 \ (\ \ d\<^sub>0) \ e\<^sub>0"] whisker_left [of d\<^sub>1 "(e\<^sub>1 \ \ \ d\<^sub>0) \ e\<^sub>0" "\\<^sup>-\<^sup>1[e\<^sub>1, dom \ \ d\<^sub>0, e\<^sub>0]"] by auto also have "... = hom.comp (GF.map \) (\\<^sub>0 (hom.dom \))" proof - have "hom.arr (GF.map \)" using \ by blast moreover have "hom.arr (\\<^sub>0 (hom.dom \))" using \ hom.arr_char hom.in_hom_char \\<^sub>0_in_hom(1) hom.dom_closed hom.dom_simp hom.inclusion ide_dom by metis moreover have "seq (GF.map \) (\\<^sub>0 (hom.dom \))" proof show "\\\<^sub>0 (hom.dom \) : dom \ \ d\<^sub>1 \ (e\<^sub>1 \ dom \ \ d\<^sub>0) \ e\<^sub>0\" using \ hom.arr_char hom.dom_char hom.in_hom_char \\<^sub>0_in_hom hom.dom_closed hom.dom_simp hom.inclusion ide_dom by metis show "\GF.map \ : d\<^sub>1 \ (e\<^sub>1 \ dom \ \ d\<^sub>0) \ e\<^sub>0 \ d\<^sub>1 \ (e\<^sub>1 \ cod \ \ d\<^sub>0) \ e\<^sub>0\" using \ hom.arr_char hom'.arr_char F.preserves_arr apply simp apply (intro hcomp_in_vhom) by (auto simp add: e\<^sub>0.antipar e\<^sub>1.antipar in_hhom_def) qed ultimately show ?thesis using \ hom.comp_char comp_assoc hom.dom_simp by auto qed finally show ?thesis by blast qed qed lemma transformation_by_components_\\<^sub>0: shows "transformation_by_components hom.comp hom.comp hom.map (G o F) \\<^sub>0" .. interpretation \: natural_isomorphism hom.comp hom.comp hom.map \G o F\ \.map proof fix f assume "hom.ide f" hence f: "ide f \ \f : src e\<^sub>0 \ src e\<^sub>1\" using hom.ide_char hom.arr_char by simp show "hom.iso (\.map f)" proof (unfold hom.iso_char hom.arr_char, intro conjI) show "iso (\.map f)" using f iso_\\<^sub>0 \.map_simp_ide hom.ide_char hom.arr_char by simp moreover show "\\.map f : src e\<^sub>0 \ src e\<^sub>1\" using f hom.ide_char hom.arr_char by blast ultimately show "\inv (\.map f) : src e\<^sub>0 \ src e\<^sub>1\" by auto qed qed lemma natural_isomorphism_\: shows "natural_isomorphism hom.comp hom.comp hom.map (G o F) \.map" .. definition \ where "\ \ \.map" lemma \_ide_simp: assumes "\f : src e\<^sub>0 \ src e\<^sub>1\" and "ide f" shows "\ f = \\<^sub>0 f" unfolding \_def using assms hom.ide_char hom.arr_char by simp lemma \_components_are_iso: assumes "\f : src e\<^sub>0 \ src e\<^sub>1\" and "ide f" shows "iso (\ f)" using assms \_def \.components_are_iso hom.ide_char hom.arr_char hom.iso_char by simp lemma \_eq: shows "\ = (\\. if \\ : src e\<^sub>0 \ src e\<^sub>1\ then \\<^sub>0 (cod \) \ \ else null)" proof fix \ have "\ \\ : src e\<^sub>0 \ src e\<^sub>1\ \ \.map \ = null" using hom.comp_char hom.null_char hom.arr_char by (simp add: \.is_extensional) moreover have "\\ : src e\<^sub>0 \ src e\<^sub>1\ \ \.map \ = \\<^sub>0 (cod \) \ \" unfolding \.map_def apply auto using hom.comp_char hom.arr_char hom.dom_simp hom.cod_simp apply simp by (metis (no_types, lifting) \\<^sub>0_in_hom(1) hom.cod_closed hom.inclusion ide_cod local.ext) ultimately show "\ \ = (if \\ : src e\<^sub>0 \ src e\<^sub>1\ then \\<^sub>0 (cod \) \ \ else null)" unfolding \_def by auto qed lemma \_in_hom [intro]: assumes "\\ : src e\<^sub>0 \ src e\<^sub>1\" shows "\\ \ : src e\<^sub>0 \ src e\<^sub>1\" and "\\ \ : dom \ \ d\<^sub>1 \ (e\<^sub>1 \ cod \ \ d\<^sub>0) \ e\<^sub>0\" proof - show "\\ \ : src e\<^sub>0 \ src e\<^sub>1\" using assms \_eq \_def hom.arr_char \.preserves_reflects_arr by presburger show "\\ \ : dom \ \ d\<^sub>1 \ (e\<^sub>1 \ cod \ \ d\<^sub>0) \ e\<^sub>0\" unfolding \_eq using assms apply simp apply (intro comp_in_homI) apply auto proof - show "\\\<^sup>-\<^sup>1[cod \] : cod \ \ cod \ \ src e\<^sub>0\" using assms by auto show "\\\<^sup>-\<^sup>1[cod \ \ src e\<^sub>0] : cod \ \ src e\<^sub>0 \ src e\<^sub>1 \ cod \ \ src e\<^sub>0\" using assms by auto show "\\\<^sub>1 \ cod \ \ \\<^sub>0 : src e\<^sub>1 \ cod \ \ src e\<^sub>0 \ (d\<^sub>1 \ e\<^sub>1) \ cod \ \ (d\<^sub>0 \ e\<^sub>0)\" using assms e\<^sub>0.unit_in_hom(2) e\<^sub>1.unit_in_hom(2) apply (intro hcomp_in_vhom) apply auto by fastforce show "\(d\<^sub>1 \ e\<^sub>1) \ \\<^sup>-\<^sup>1[cod \, d\<^sub>0, e\<^sub>0] : (d\<^sub>1 \ e\<^sub>1) \ cod \ \ d\<^sub>0 \ e\<^sub>0 \ (d\<^sub>1 \ e\<^sub>1) \ (cod \ \ d\<^sub>0) \ e\<^sub>0\" using assms assoc'_in_hom e\<^sub>0.antipar(1-2) e\<^sub>1.antipar(2) apply (intro hcomp_in_vhom) by auto show "\\[d\<^sub>1, e\<^sub>1, (cod \ \ d\<^sub>0) \ e\<^sub>0] : (d\<^sub>1 \ e\<^sub>1) \ (cod \ \ d\<^sub>0) \ e\<^sub>0 \ d\<^sub>1 \ e\<^sub>1 \ (cod \ \ d\<^sub>0) \ e\<^sub>0\" using assms assoc_in_hom e\<^sub>0.antipar(1-2) e\<^sub>1.antipar(2) by auto show "\d\<^sub>1 \ \\<^sup>-\<^sup>1[e\<^sub>1, cod \ \ d\<^sub>0, e\<^sub>0] : d\<^sub>1 \ e\<^sub>1 \ (cod \ \ d\<^sub>0) \ e\<^sub>0 \ d\<^sub>1 \ (e\<^sub>1 \ cod \ \ d\<^sub>0) \ e\<^sub>0\" using assms assoc'_in_hom [of "d\<^sub>1" "e\<^sub>1 \ cod \ \ d\<^sub>0" "e\<^sub>0"] e\<^sub>0.antipar(1-2) e\<^sub>1.antipar(1-2) apply (intro hcomp_in_vhom) apply auto by fastforce qed qed lemma \_simps [simp]: assumes "\\ : src e\<^sub>0 \ src e\<^sub>1\" shows "arr (\ \)" and "src (\ \) = src e\<^sub>0" and "trg (\ \) = src e\<^sub>1" and "dom (\ \) = dom \" and "cod (\ \) = d\<^sub>1 \ (e\<^sub>1 \ cod \ \ d\<^sub>0) \ e\<^sub>0" using assms \_in_hom by auto interpretation d\<^sub>0: equivalence_in_bicategory V H \ \ src trg d\<^sub>0 e\<^sub>0 \inv \\<^sub>0\ \inv \\<^sub>0\ using e\<^sub>0.dual_equivalence by simp interpretation d\<^sub>1: equivalence_in_bicategory V H \ \ src trg d\<^sub>1 e\<^sub>1 \inv \\<^sub>1\ \inv \\<^sub>1\ using e\<^sub>1.dual_equivalence by simp interpretation d\<^sub>0e\<^sub>0: two_equivalences_in_bicategory V H \ \ src trg d\<^sub>0 e\<^sub>0 \inv \\<^sub>0\ \inv \\<^sub>0\ d\<^sub>1 e\<^sub>1 \inv \\<^sub>1\ \inv \\<^sub>1\ .. interpretation \: inverse_transformation hom'.comp hom'.comp hom'.map \F o G\ d\<^sub>0e\<^sub>0.\ proof - interpret \': natural_isomorphism hom'.comp hom'.comp hom'.map \F o G\ d\<^sub>0e\<^sub>0.\ using d\<^sub>0e\<^sub>0.natural_isomorphism_\ e\<^sub>0.antipar e\<^sub>1.antipar d\<^sub>0e\<^sub>0.\_eq d\<^sub>0e\<^sub>0.\_def by metis show "inverse_transformation hom'.comp hom'.comp hom'.map (F o G) d\<^sub>0e\<^sub>0.\" .. qed definition \ where "\ \ \.map" lemma \_ide_simp: assumes "\f': trg e\<^sub>0 \ trg e\<^sub>1\" and "ide f'" shows "\ f' = \[f'] \ \[f' \ trg e\<^sub>0] \ (\\<^sub>1 \ f' \ \\<^sub>0) \ ((e\<^sub>1 \ d\<^sub>1) \ \[f', e\<^sub>0, d\<^sub>0]) \ \\<^sup>-\<^sup>1[e\<^sub>1, d\<^sub>1, (f' \ e\<^sub>0) \ d\<^sub>0] \ (e\<^sub>1 \ \[d\<^sub>1, f' \ e\<^sub>0, d\<^sub>0])" proof - have "hom'.ide f'" using assms hom'.ide_char hom'.arr_char by simp hence "\.map f' = hom'.inv (d\<^sub>0e\<^sub>0.\ f')" using assms by simp also have "... = inv (d\<^sub>0e\<^sub>0.\ f')" using hom'.inv_char \hom'.ide f'\ by simp also have "... = inv (d\<^sub>0e\<^sub>0.\\<^sub>0 f')" using assms e\<^sub>0.antipar e\<^sub>1.antipar d\<^sub>0e\<^sub>0.\_eq d\<^sub>0e\<^sub>0.\_ide_simp [of f'] by simp also have "... = ((((inv \\<^sup>-\<^sup>1[f'] \ inv \\<^sup>-\<^sup>1[f' \ trg e\<^sub>0]) \ inv (inv \\<^sub>1 \ f' \ inv \\<^sub>0)) \ inv ((e\<^sub>1 \ d\<^sub>1) \ \\<^sup>-\<^sup>1[f', e\<^sub>0, d\<^sub>0])) \ inv \[e\<^sub>1, d\<^sub>1, (f' \ e\<^sub>0) \ d\<^sub>0]) \ inv (e\<^sub>1 \ \\<^sup>-\<^sup>1[d\<^sub>1, f' \ e\<^sub>0, d\<^sub>0])" proof - text \There has to be a better way to do this.\ have 1: "\A B C D E F. \ iso A; iso B; iso C; iso D; iso E; iso F; arr (((((A \ B) \ C) \ D) \ E) \ F) \ \ inv (((((A \ B) \ C) \ D) \ E) \ F) = inv F \ inv E \ inv D \ inv C \ inv B \ inv A" using inv_comp isos_compose seqE by metis have "arr (d\<^sub>0e\<^sub>0.\\<^sub>0 f')" using assms e\<^sub>0.antipar(2) e\<^sub>1.antipar(2) d\<^sub>0e\<^sub>0.iso_\\<^sub>0 [of f'] iso_is_arr by simp moreover have "iso \\<^sup>-\<^sup>1[f']" using assms iso_runit' by simp moreover have "iso \\<^sup>-\<^sup>1[f' \ trg e\<^sub>0]" using assms iso_lunit' by auto moreover have "iso (inv \\<^sub>1 \ f' \ inv \\<^sub>0)" using assms e\<^sub>0.antipar(2) e\<^sub>1.antipar(2) in_hhom_def by simp moreover have "iso ((e\<^sub>1 \ d\<^sub>1) \ \\<^sup>-\<^sup>1[f', e\<^sub>0, d\<^sub>0])" using assms e\<^sub>0.antipar e\<^sub>1.antipar(1) e\<^sub>1.antipar(2) in_hhom_def iso_hcomp by (metis calculation(1) e\<^sub>0.ide_left e\<^sub>0.ide_right e\<^sub>1.ide_left e\<^sub>1.ide_right hseqE ide_is_iso iso_assoc' seqE) moreover have "iso \[e\<^sub>1, d\<^sub>1, (f' \ e\<^sub>0) \ d\<^sub>0]" using assms e\<^sub>0.antipar e\<^sub>1.antipar by auto moreover have "iso (e\<^sub>1 \ \\<^sup>-\<^sup>1[d\<^sub>1, f' \ e\<^sub>0, d\<^sub>0])" using assms e\<^sub>0.antipar e\<^sub>1.antipar by (metis calculation(1) e\<^sub>0.ide_left e\<^sub>0.ide_right e\<^sub>1.ide_left e\<^sub>1.ide_right iso_hcomp ide_hcomp hseqE ideD(1) ide_is_iso in_hhomE iso_assoc' seqE hcomp_simps(1-2)) ultimately show ?thesis using 1 [of "e\<^sub>1 \ \\<^sup>-\<^sup>1[d\<^sub>1, f' \ e\<^sub>0, d\<^sub>0]" "\[e\<^sub>1, d\<^sub>1, (f' \ e\<^sub>0) \ d\<^sub>0]" "(e\<^sub>1 \ d\<^sub>1) \ \\<^sup>-\<^sup>1[f', e\<^sub>0, d\<^sub>0]" "inv \\<^sub>1 \ f' \ inv \\<^sub>0" "\\<^sup>-\<^sup>1[f' \ trg e\<^sub>0]" "\\<^sup>-\<^sup>1[f']"] comp_assoc by (metis e\<^sub>0.antipar(2)) qed also have "... = inv \\<^sup>-\<^sup>1[f'] \ inv \\<^sup>-\<^sup>1[f' \ trg e\<^sub>0] \ inv (inv \\<^sub>1 \ f' \ inv \\<^sub>0) \ inv ((e\<^sub>1 \ d\<^sub>1) \ \\<^sup>-\<^sup>1[f', e\<^sub>0, d\<^sub>0]) \ inv \[e\<^sub>1, d\<^sub>1, (f' \ e\<^sub>0) \ d\<^sub>0] \ inv (e\<^sub>1 \ \\<^sup>-\<^sup>1[d\<^sub>1, f' \ e\<^sub>0, d\<^sub>0])" using comp_assoc by simp also have "... = \[f'] \ \[f' \ trg e\<^sub>0] \ (\\<^sub>1 \ f' \ \\<^sub>0) \ ((e\<^sub>1 \ d\<^sub>1) \ \[f', e\<^sub>0, d\<^sub>0]) \ \\<^sup>-\<^sup>1[e\<^sub>1, d\<^sub>1, (f' \ e\<^sub>0) \ d\<^sub>0] \ (e\<^sub>1 \ \[d\<^sub>1, f' \ e\<^sub>0, d\<^sub>0])" proof - have "inv \\<^sup>-\<^sup>1[f'] = \[f']" using assms inv_inv iso_runit by blast moreover have "inv \\<^sup>-\<^sup>1[f' \ trg e\<^sub>0] = \[f' \ trg e\<^sub>0]" using assms iso_lunit by auto moreover have "inv (inv \\<^sub>1 \ f' \ inv \\<^sub>0) = \\<^sub>1 \ f' \ \\<^sub>0" proof - have "src (inv \\<^sub>1) = trg f'" using assms(1) e\<^sub>1.antipar(2) by auto moreover have "src f' = trg (inv \\<^sub>0)" using assms(1) e\<^sub>0.antipar(2) by auto ultimately show ?thesis using assms(2) e\<^sub>0.counit_is_iso e\<^sub>1.counit_is_iso by simp qed ultimately show ?thesis using assms e\<^sub>0.antipar e\<^sub>1.antipar by auto qed finally show ?thesis using \_def by simp qed lemma \_components_are_iso: assumes "\f' : trg e\<^sub>0 \ trg e\<^sub>1\" and "ide f'" shows "iso (\ f')" using assms \_def \.components_are_iso hom'.ide_char hom'.arr_char hom'.iso_char by simp lemma \_eq: shows "\ = (\\'. if \\': trg e\<^sub>0 \ trg e\<^sub>1\ then \' \ \[dom \'] \ \[dom \' \ trg e\<^sub>0] \ (\\<^sub>1 \ dom \' \ \\<^sub>0) \ ((e\<^sub>1 \ d\<^sub>1) \ \[dom \', e\<^sub>0, d\<^sub>0]) \ \\<^sup>-\<^sup>1[e\<^sub>1, d\<^sub>1, (dom \' \ e\<^sub>0) \ d\<^sub>0] \ (e\<^sub>1 \ \[d\<^sub>1, dom \' \ e\<^sub>0, d\<^sub>0]) else null)" proof fix \' have "\ \\': trg e\<^sub>0 \ trg e\<^sub>1\ \ \.map \' = null" using \.is_extensional hom'.arr_char hom'.null_char by simp moreover have "\\': trg e\<^sub>0 \ trg e\<^sub>1\ \ \.map \' = \' \ \[dom \'] \ \[dom \' \ trg e\<^sub>0] \ (\\<^sub>1 \ dom \' \ \\<^sub>0) \ ((e\<^sub>1 \ d\<^sub>1) \ \[dom \', e\<^sub>0, d\<^sub>0]) \ \\<^sup>-\<^sup>1[e\<^sub>1, d\<^sub>1, (dom \' \ e\<^sub>0) \ d\<^sub>0] \ (e\<^sub>1 \ \[d\<^sub>1, dom \' \ e\<^sub>0, d\<^sub>0])" proof - assume \': "\\': trg e\<^sub>0 \ trg e\<^sub>1\" have "\\.map (dom \') : trg e\<^sub>0 \ trg e\<^sub>1\" using \' hom'.arr_char hom'.dom_closed by auto moreover have "seq \' (\.map (dom \'))" proof - have "hom'.seq \' (\.map (dom \'))" using \' \.preserves_cod hom'.arrI hom'.dom_simp hom'.cod_simp apply (intro hom'.seqI) by auto thus ?thesis using hom'.seq_char by blast qed ultimately show ?thesis using \' \.is_natural_1 [of \'] hom'.comp_char hom'.arr_char hom'.dom_closed \_ide_simp [of "dom \'"] hom'.dom_simp hom'.cod_simp apply auto by (metis \_def hom'.inclusion ide_dom) qed ultimately show "\ \' = (if \\' : trg e\<^sub>0 \ trg e\<^sub>1\ then \' \ \[dom \'] \ \[dom \' \ trg e\<^sub>0] \ (\\<^sub>1 \ dom \' \ \\<^sub>0) \ ((e\<^sub>1 \ d\<^sub>1) \ \[dom \', e\<^sub>0, d\<^sub>0]) \ \\<^sup>-\<^sup>1[e\<^sub>1, d\<^sub>1, (dom \' \ e\<^sub>0) \ d\<^sub>0] \ (e\<^sub>1 \ \[d\<^sub>1, dom \' \ e\<^sub>0, d\<^sub>0]) else null)" unfolding \_def by argo qed lemma \_in_hom [intro]: assumes "\\' : trg e\<^sub>0 \ trg e\<^sub>1\" shows "\\ \' : trg e\<^sub>0 \ trg e\<^sub>1\" and "\\ \' : e\<^sub>1 \ (d\<^sub>1 \ dom \' \ e\<^sub>0) \ d\<^sub>0 \ cod \'\" proof - have 0: "\ \' = \.map \'" using \_def by auto hence 1: "hom'.arr (\ \')" using assms hom'.arr_char \.preserves_reflects_arr by auto show "\\ \' : trg e\<^sub>0 \ trg e\<^sub>1\" using 1 hom'.arr_char by blast show "\\ \' : e\<^sub>1 \ (d\<^sub>1 \ dom \' \ e\<^sub>0) \ d\<^sub>0 \ cod \'\" using assms 0 1 \.preserves_hom hom'.in_hom_char hom'.arr_char e\<^sub>0.antipar e\<^sub>1.antipar \.preserves_dom \.preserves_cod hom'.dom_char apply (intro in_homI) apply auto[1] proof - show "dom (\ \') = e\<^sub>1 \ (d\<^sub>1 \ dom \' \ e\<^sub>0) \ d\<^sub>0" proof - have "hom'.dom (\ \') = FG.map (hom'.dom \')" using assms 0 \.preserves_dom hom'.arr_char by (metis (no_types, lifting)) thus ?thesis using assms 0 1 hom.arr_char hom'.arr_char hom'.dom_char G.preserves_arr hom'.dom_closed by auto qed show "cod (\ \') = cod \'" proof - have "hom'.cod (\ \') = cod \'" using assms 0 \.preserves_cod hom'.arr_char hom'.cod_simp by auto thus ?thesis using assms 0 1 hom'.arr_char hom'.cod_char G.preserves_arr hom'.cod_closed by auto qed qed qed lemma \_simps [simp]: assumes "\\' : trg e\<^sub>0 \ trg e\<^sub>1\" shows "arr (\ \')" and "src (\ \') = trg e\<^sub>0" and "trg (\ \') = trg e\<^sub>1" and "dom (\ \') = e\<^sub>1 \ (d\<^sub>1 \ dom \' \ e\<^sub>0) \ d\<^sub>0" and "cod (\ \') = cod \'" using assms \_in_hom by auto interpretation equivalence_of_categories hom'.comp hom.comp F G \ \ proof - interpret \: natural_isomorphism hom.comp hom.comp hom.map \G o F\ \ using \.natural_isomorphism_axioms \_def by simp interpret \: natural_isomorphism hom'.comp hom'.comp \F o G\ hom'.map \ using \.natural_isomorphism_axioms \_def by simp show "equivalence_of_categories hom'.comp hom.comp F G \ \" .. qed lemma induces_equivalence_of_hom_categories: shows "equivalence_of_categories hom'.comp hom.comp F G \ \" .. lemma equivalence_functor_F: shows "equivalence_functor hom.comp hom'.comp F" proof - interpret \': inverse_transformation hom.comp hom.comp hom.map \G o F\ \ .. interpret \': inverse_transformation hom'.comp hom'.comp \F o G\ hom'.map \ .. interpret E: equivalence_of_categories hom.comp hom'.comp G F \'.map \'.map .. show ?thesis using E.equivalence_functor_axioms by simp qed lemma equivalence_functor_G: shows "equivalence_functor hom'.comp hom.comp G" using equivalence_functor_axioms by simp end context bicategory begin text \ We now use the just-established equivalence of hom-categories to prove some cancellation laws for equivalence maps. It is relatively straightforward to prove these results directly, without using the just-established equivalence, but the proofs are somewhat longer that way. \ lemma equivalence_cancel_left: assumes "equivalence_map e" and "par \ \'" and "src e = trg \" and "e \ \ = e \ \'" shows "\ = \'" proof - obtain d \ \ where d\\: "equivalence_in_bicategory V H \ \ src trg e d \ \" using assms equivalence_map_def by auto interpret e: equivalence_in_bicategory V H \ \ src trg e d \ \ using d\\ by simp interpret i: equivalence_in_bicategory V H \ \ src trg \src \\ \src \\ \inv \[src \]\ \\[src \]\ using assms iso_unit obj_src by unfold_locales simp_all interpret two_equivalences_in_bicategory V H \ \ src trg \src \\ \src \\ \inv \[src \]\ \\[src \]\ e d \ \ .. interpret hom: subcategory V \\\'. in_hhom \' (src (src \)) (src e)\ using hhom_is_subcategory by blast interpret hom': subcategory V \\\'. in_hhom \' (trg (src \)) (trg e)\ using hhom_is_subcategory by blast interpret F: equivalence_functor hom.comp hom'.comp \\\'. e \ \' \ src \\ using equivalence_functor_F by simp have "F \ = F \'" using assms hom.arr_char apply simp by (metis e.ide_left hcomp_reassoc(2) i.ide_right ideD(1) src_dom trg_dom trg_src) moreover have "hom.par \ \'" using assms hom.arr_char hom.dom_char hom.cod_char by (metis (no_types, lifting) in_hhomI src_dom src_src trg_dom) ultimately show "\ = \'" using F.is_faithful by blast qed lemma equivalence_cancel_right: assumes "equivalence_map e" and "par \ \'" and "src \ = trg e" and "\ \ e = \' \ e" shows "\ = \'" proof - obtain d \ \ where d\\: "equivalence_in_bicategory V H \ \ src trg e d \ \" using assms equivalence_map_def by auto interpret e: equivalence_in_bicategory V H \ \ src trg e d \ \ using d\\ by simp interpret i: equivalence_in_bicategory V H \ \ src trg \trg \\ \trg \\ \inv \[trg \]\ \\[trg \]\ using assms iso_unit obj_src by unfold_locales simp_all interpret two_equivalences_in_bicategory V H \ \ src trg e d \ \ \trg \\ \trg \\ \inv \[trg \]\ \\[trg \]\ .. interpret hom: subcategory V \\\'. in_hhom \' (trg e) (trg (trg \))\ using hhom_is_subcategory by blast interpret hom': subcategory V \\\'. in_hhom \' (src e) (src (trg \))\ using hhom_is_subcategory by blast interpret G: equivalence_functor hom.comp hom'.comp \\\'. trg \ \ \' \ e\ using equivalence_functor_G by simp have "G \ = G \'" using assms hom.arr_char by simp moreover have "hom.par \ \'" using assms hom.arr_char hom.dom_char hom.cod_char by (metis (no_types, lifting) in_hhomI src_dom trg_dom trg_trg) ultimately show "\ = \'" using G.is_faithful by blast qed lemma equivalence_isomorphic_cancel_left: assumes "equivalence_map e" and "ide f" and "ide f'" and "src f = src f'" and "src e = trg f" and "e \ f \ e \ f'" shows "f \ f'" proof - have ef': "src e = trg f'" using assms R.as_nat_iso.components_are_iso iso_is_arr isomorphic_implies_hpar(2) by blast obtain d \ \ where e: "equivalence_in_bicategory V H \ \ src trg e d \ \" using assms equivalence_map_def by auto interpret e: equivalence_in_bicategory V H \ \ src trg e d \ \ using e by simp interpret i: equivalence_in_bicategory V H \ \ src trg \src f\ \src f\ \inv \[src f]\ \\[src f]\ using assms iso_unit obj_src by unfold_locales auto interpret two_equivalences_in_bicategory V H \ \ src trg \src f\ \src f\ \inv \[src f]\ \\[src f]\ e d \ \ .. interpret hom: subcategory V \\\'. in_hhom \' (src (src f)) (src e)\ using hhom_is_subcategory by blast interpret hom': subcategory V \\\'. in_hhom \' (trg (src f)) (trg e)\ using hhom_is_subcategory by blast interpret F: equivalence_functor hom.comp hom'.comp \\\'. e \ \' \ src f\ using equivalence_functor_F by simp have 1: "F f \ F f'" proof - have "F f \ (e \ f) \ src f" using assms hcomp_assoc_isomorphic equivalence_map_is_ide isomorphic_symmetric by auto also have "... \ (e \ f') \ src f'" using assms ef' by (simp add: hcomp_isomorphic_ide) also have "... \ F f'" using assms ef' hcomp_assoc_isomorphic equivalence_map_is_ide by auto finally show ?thesis by blast qed show "f \ f'" proof - obtain \ where \: "\\ : F f \ F f'\ \ iso \" using 1 isomorphic_def by auto have 2: "hom'.iso \" using assms \ hom'.iso_char hom'.arr_char vconn_implies_hpar(1,2) by auto have 3: "hom'.in_hom \ (F f) (F f')" using assms 2 \ ef' hom'.in_hom_char hom'.arr_char by (metis F.preserves_reflects_arr hom'.iso_is_arr hom.arr_char i.antipar(1) ideD(1) ide_in_hom(1) trg_src) obtain \ where \: "hom.in_hom \ f f' \ F \ = \" using assms 3 \ F.is_full F.preserves_reflects_arr hom'.in_hom_char hom.ide_char by fastforce have "hom.iso \" using 2 \ F.reflects_iso by auto thus ?thesis using \ isomorphic_def hom.in_hom_char hom.arr_char hom.iso_char by auto qed qed lemma equivalence_isomorphic_cancel_right: assumes "equivalence_map e" and "ide f" and "ide f'" and "trg f = trg f'" and "src f = trg e" and "f \ e \ f' \ e" shows "f \ f'" proof - have f'e: "src f' = trg e" using assms R.as_nat_iso.components_are_iso iso_is_arr isomorphic_implies_hpar(2) by blast obtain d \ \ where d\\: "equivalence_in_bicategory V H \ \ src trg e d \ \" using assms equivalence_map_def by auto interpret e: equivalence_in_bicategory V H \ \ src trg e d \ \ using d\\ by simp interpret i: equivalence_in_bicategory V H \ \ src trg \trg f\ \trg f\ \inv \[trg f]\ \\[trg f]\ using assms iso_unit obj_src by unfold_locales auto interpret two_equivalences_in_bicategory V H \ \ src trg e d \ \ \trg f\ \trg f\ \inv \[trg f]\ \\[trg f]\ .. interpret hom: subcategory V \\\'. in_hhom \' (trg e) (trg (trg f))\ using hhom_is_subcategory by blast interpret hom': subcategory V \\\'. in_hhom \' (src e) (src (trg f))\ using hhom_is_subcategory by blast interpret G: equivalence_functor hom.comp hom'.comp \\\'. trg f \ \' \ e\ using equivalence_functor_G by simp have 1: "G f \ G f'" using assms hcomp_isomorphic_ide hcomp_ide_isomorphic by simp show "f \ f'" proof - obtain \ where \: "\\ : G f \ G f'\ \ iso \" using 1 isomorphic_def by auto have 2: "hom'.iso \" using assms \ hom'.iso_char hom'.arr_char vconn_implies_hpar(1-2) by auto have 3: "hom'.in_hom \ (G f) (G f')" using assms 2 \ f'e hom'.in_hom_char hom'.arr_char by (metis G.preserves_arr hom'.iso_is_arr hom.ideI hom.ide_char ideD(1) ide_in_hom(1) trg_trg) obtain \ where \: "hom.in_hom \ f f' \ G \ = \" using assms 3 \ G.is_full G.preserves_reflects_arr hom'.in_hom_char hom.ide_char by fastforce have "hom.iso \" using 2 \ G.reflects_iso by auto thus ?thesis using \ isomorphic_def hom.in_hom_char hom.arr_char hom.iso_char by auto qed qed end end diff --git a/thys/Bicategory/Modification.thy b/thys/Bicategory/Modification.thy --- a/thys/Bicategory/Modification.thy +++ b/thys/Bicategory/Modification.thy @@ -1,1299 +1,1299 @@ (* Title: Modification Author: Eugene W. Stark , 2020 Maintainer: Eugene W. Stark *) section "Modifications" text \ Modifications are morphisms of pseudonatural transformations. In this section, we give a definition of ``modification'', and we prove that the mappings \\\ and \\\, which were chosen in the course of showing that a pseudonatural equivalence \\\ has a converse pseudonatural equivalence \\'\, are invertible modifications that relate the composites \\'\\ and \\\'\ to the identity pseudonatural transformations on \F\ and \G\. This means that \\\ and \\'\ are quasi-inverses in a suitable bicategory of pseudofunctors, pseudonatural transformations, and modifications, though we do not go so far as to give a formal construction of such a bicategory. \ theory Modification imports PseudonaturalTransformation begin locale modification = (* 12 sec *) C: bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + D: bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + \: 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 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" and \\<^sub>0' :: "'c \ 'd" and \\<^sub>1' :: "'c \ 'd" and \ :: "'c \ 'd" + assumes \_in_hom: "C.obj a \ \\ a : \\<^sub>0 a \\<^sub>D \\<^sub>0' a\" and naturality: "\ \f : a \\<^sub>C b\; C.ide f \ \ \\<^sub>1' f \\<^sub>D (G f \\<^sub>D \ a) = (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>1 f" locale invertible_modification = (* 13 sec *) modification + assumes components_are_iso: "C.obj a \ D.iso (\ a)" locale identity_modification = (* 12 sec *) C: bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + D: bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + \: 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>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 abbreviation map where "map \ \\<^sub>0" sublocale invertible_modification 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 \\<^sub>0 \\<^sub>1 map using D.comp_arr_dom D.comp_cod_arr by unfold_locales auto end locale composite_modification = (* 13 sec *) C: bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C + D: bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D + \: 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 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 F \\<^sub>F G \\<^sub>G \\<^sub>0 \\<^sub>1 + \: modification 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 \\<^sub>0 \\<^sub>1 \ + \: modification 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 \\<^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" and \\<^sub>0 :: "'c \ 'd" and \\<^sub>1 :: "'c \ 'd" and \\<^sub>0 :: "'c \ 'd" and \\<^sub>1 :: "'c \ 'd" and \ :: "'c \ 'd" and \ :: "'c \ 'd" begin abbreviation map where "map a \ \ a \\<^sub>D \ a" sublocale modification 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 \\<^sub>0 \\<^sub>1 map using \.\_in_hom \.\_in_hom \.naturality \.naturality apply unfold_locales apply auto[1] proof - fix f a b assume f: "\f : a \\<^sub>C b\" and ide_f: "C.ide f" have "\\<^sub>1 f \\<^sub>D (G f \\<^sub>D map a) = (\\<^sub>1 f \\<^sub>D (G f \\<^sub>D \ a)) \\<^sub>D (G f \\<^sub>D \ a)" proof - have "D.seq (\ a) (\ a)" using f \.\_in_hom \.\_in_hom by blast thus ?thesis using f ide_f D.whisker_left [of "G f" "\ a" "\ a"] D.comp_assoc by simp qed also have "... = (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>1 f \\<^sub>D (G f \\<^sub>D \ a)" using f ide_f \.naturality [of f a b] D.comp_assoc by simp also have "... = ((\ b \\<^sub>D F f) \\<^sub>D (\ b \\<^sub>D F f)) \\<^sub>D \\<^sub>1 f" using f ide_f \.naturality [of f a b] D.comp_assoc by simp also have "... = (map b \\<^sub>D F f) \\<^sub>D \\<^sub>1 f" proof - have "D.seq (\ b) (\ b)" using f \.\_in_hom \.\_in_hom by blast thus ?thesis using f ide_f D.whisker_right [of "F f" "\ b" "\ b"] by simp qed finally show "\\<^sub>1 f \\<^sub>D (G f \\<^sub>D map a) = (map b \\<^sub>D F f) \\<^sub>D \\<^sub>1 f" by simp qed end context converse_pseudonatural_equivalence begin interpretation EV: self_evaluation_map V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D .. notation EV.eval ("\_\") interpretation \\<^sub>F: identity_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 .. interpretation \\<^sub>G: identity_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 .. interpretation \'\: composite_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 F \\<^sub>F \\<^sub>0 \\<^sub>1 map\<^sub>0 map\<^sub>1 .. interpretation \\': composite_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 G \\<^sub>G map\<^sub>0 map\<^sub>1 \\<^sub>0 \\<^sub>1 .. interpretation \: invertible_modification 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 \\<^sub>F.map\<^sub>0 \\<^sub>F.map\<^sub>1 \'\.map\<^sub>0 \'\.map\<^sub>1 \ proof show "\a. C.obj a \ \\ a : F.map\<^sub>0 a \\<^sub>D \'\.map\<^sub>0 a\" using unit_in_hom \'\.map\<^sub>0_def by simp show "\a. C.obj a \ D.iso (\ a)" by simp show "\f a b. \ \f : a \\<^sub>C b\; C.ide f\ \ \'\.map\<^sub>1 f \\<^sub>D (F f \\<^sub>D \ a) = (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>D[F f]" proof - fix f a b assume f: "\f : a \\<^sub>C b\" and ide_f: "C.ide f" have a: "C.obj a" and b: "C.obj b" using f by auto have "\'\.map\<^sub>1 f \\<^sub>D (F f \\<^sub>D \ a) = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^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>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, \\<^sub>0' a, \\<^sub>0 a]) \\<^sub>D (F f \\<^sub>D \ a)" unfolding \'\.map\<^sub>1_def map\<^sub>1_def using a b f D.comp_assoc by auto also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \ 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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^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>0 a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D (((\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D (((\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, \\<^sub>0' a, \\<^sub>0 a]) \\<^sub>D (F f \\<^sub>D \ a)" proof - have "(\\<^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>0 a = ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \ 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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^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>0 a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D (((\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a)" proof - have "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 a b f ide_f \.iso_map\<^sub>1_ide - by (intro D.seqI) auto + by (metis (no_types, lifting) C.in_hhomE map\<^sub>1_def map\<^sub>1_simps(1)) thus ?thesis using a b f D.whisker_right [of "\\<^sub>0 a"] by fastforce qed thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \ 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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^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>0 a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D (((\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D (F f \\<^sub>D \ a)" proof - have "((\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, \\<^sub>0' a, \\<^sub>0 a] = \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a)" using a b f ide_f D.assoc'_naturality [of "\\<^sub>D\<^sup>-\<^sup>1[F f]" "\\<^sub>0' a" "\\<^sub>0 a"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \ 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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^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>0 a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((((\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a, \\<^sub>0 a]) \\<^sub>D ((F\<^sub>0 b \\<^sub>D F f) \\<^sub>D \ a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D (F f \\<^sub>D \ a) = \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \ a" using a b f ide_f D.comp_arr_dom D.comp_cod_arr D.interchange [of "\\<^sub>D\<^sup>-\<^sup>1[F f]" "F f" "\\<^sub>0' a \\<^sub>D \\<^sub>0 a" "\ a"] by simp also have "... = ((F\<^sub>0 b \\<^sub>D F f) \\<^sub>D \ a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" using a b f ide_f D.comp_arr_dom D.comp_cod_arr D.interchange [of "F\<^sub>0 b \\<^sub>D F f" "\\<^sub>D\<^sup>-\<^sup>1[F f]" "\ a" "F\<^sub>0 a"] by auto finally have "(\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D (F f \\<^sub>D \ a) = ((F\<^sub>0 b \\<^sub>D F f) \\<^sub>D \ a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \ 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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^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>0 a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f, \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (((\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D ((F\<^sub>0 b \\<^sub>D F f) \\<^sub>D \ a)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "(((\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a, \\<^sub>0 a] = \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f, \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a)" using a b f ide_f D.assoc'_naturality [of "\ b \\<^sub>D F f" "\\<^sub>0' a" "\\<^sub>0 a"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \ 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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^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>0 a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f, \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (((\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f) \\<^sub>D \ a) \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "((\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D ((F\<^sub>0 b \\<^sub>D F f) \\<^sub>D \ a) = (\ b \\<^sub>D F f) \\<^sub>D \ a" using a b f ide_f D.comp_arr_dom D.comp_cod_arr D.interchange [of "\ b \\<^sub>D F f" "F\<^sub>0 b \\<^sub>D F f" "\\<^sub>0' a \\<^sub>D \\<^sub>0 a" "\ a"] by auto also have "... = (((\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f) \\<^sub>D \ a) \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a)" using a b f ide_f D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f" "\ b \\<^sub>D F f" "\ a" "F\<^sub>0 a"] by auto finally have "((\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D ((F\<^sub>0 b \\<^sub>D F f) \\<^sub>D \ a) = (((\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f) \\<^sub>D \ a) \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \ 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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D \\<^sub>0 a) \\<^sub>D ((((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f, \\<^sub>0' a, \\<^sub>0 a]) \\<^sub>D (((\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f) \\<^sub>D \ a) \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "(((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) = ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a" using a b f ide_f D.whisker_right \.iso_map\<^sub>1_ide by auto thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \ 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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D (G f \\<^sub>D \\<^sub>0 a), \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D (((\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f) \\<^sub>D \ a)) \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "(((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f, \\<^sub>0' a, \\<^sub>0 a] = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D (G f \\<^sub>D \\<^sub>0 a), \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a)" using a b f ide_f \.iso_map\<^sub>1_ide D.assoc'_naturality [of "(\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f]" "\\<^sub>0' a" "\\<^sub>0 a"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \ 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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D (G f \\<^sub>D \\<^sub>0 a), \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a) \\<^sub>D \ a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D F\<^sub>0 a) \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D (((\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f) \\<^sub>D \ a) = ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f]) \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f) \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D \ a" using a b f ide_f \.iso_map\<^sub>1_ide D.interchange [of "(\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f]" "(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f" "\\<^sub>0' a \\<^sub>D \\<^sub>0 a" "\ a"] by auto also have "... = (\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \ a" using a b f ide_f \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.comp_assoc by auto also have "... = (\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \ a \\<^sub>D F\<^sub>0 a" using a b f ide_f \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr by auto also have "... = ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a) \\<^sub>D \ a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D F\<^sub>0 a)" proof - have "D.seq (\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a) ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f])" using a b f ide_f \.iso_map\<^sub>1_ide by (intro D.seqI D.hseqI') auto thus ?thesis using a b f ide_f \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr D.interchange [of "(\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a)" "(\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f]" "\ a" "F\<^sub>0 a"] by simp qed finally have "((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D (((\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f) \\<^sub>D \ a) = ((\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a) \\<^sub>D \ a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D F\<^sub>0 a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f, G\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, G\<^sub>0 a, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \ a \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a \\<^sub>D \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a \\<^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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D (G f \\<^sub>D \\<^sub>0 a), \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, \\<^sub>0' a \\<^sub>D \\<^sub>0 a]) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>0 a \\<^sub>D \ a) \\<^sub>D \\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, F\<^sub>0 a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D F\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D F\<^sub>0 a) \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "(\\<^sub>0' b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D \\<^sub>0 a = (\\<^sub>D[\\<^sub>0' b, G f, G\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, G\<^sub>0 a, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \ a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a \\<^sub>D \\<^sub>0' a, \\<^sub>0 a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a \\<^sub>D \\<^sub>0' a] \\<^sub>D \\<^sub>0 a)" proof - have "(\\<^sub>0' b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D \\<^sub>0 a = (\\<^sub>D[\\<^sub>0' b, G f, G\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D (((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \ a) \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a \\<^sub>D \\<^sub>0' a] \\<^sub>D \\<^sub>0 a)" using a b f ide_f D.hcomp_reassoc D.whisker_right [of "\\<^sub>0 a"] by auto also have "... = (\\<^sub>D[\\<^sub>0' b, G f, G\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, G\<^sub>0 a, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \ a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a \\<^sub>D \\<^sub>0' a, \\<^sub>0 a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a \\<^sub>D \\<^sub>0' a] \\<^sub>D \\<^sub>0 a)" using a b f ide_f D.hcomp_reassoc(1) [of "\\<^sub>0' b \\<^sub>D G f" "\ a" "\\<^sub>0 a"] by auto finally show ?thesis by blast qed moreover have "(\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a) \\<^sub>D \ a = (\\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, \\<^sub>0' a \\<^sub>D \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>0 a \\<^sub>D \ a) \\<^sub>D \\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, F\<^sub>0 a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D F\<^sub>0 a)" proof - have "(\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a) \\<^sub>D \ a = \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D (\\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D \ a \\<^sub>D F\<^sub>0 a" using a b f ide_f D.comp_arr_dom D.comp_cod_arr D.hcomp_reassoc(2) [of "\\<^sub>0' b" "G f" "\\<^sub>0 a"] by auto also have "... = (\\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D (((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>0 a) \\<^sub>D \ a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D F\<^sub>0 a)" using a b f ide_f D.inv_inv D.iso_inv_iso D.interchange by auto also have "... = (\\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, \\<^sub>0' a \\<^sub>D \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>0 a \\<^sub>D \ a) \\<^sub>D \\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, F\<^sub>0 a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D F\<^sub>0 a)" using a b f ide_f D.hcomp_reassoc(1) [of "\\<^sub>0' b \\<^sub>D G f" "\\<^sub>0 a" "\ a"] by auto finally show ?thesis by blast qed ultimately show ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f, G\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, G\<^sub>0 a, \\<^sub>0 a] \\<^sub>D (((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \ a \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0' a, \\<^sub>0 a]) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>0 a \\<^sub>D \ a)) \\<^sub>D \\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, F\<^sub>0 a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D F\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D F\<^sub>0 a) \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "\\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a \\<^sub>D \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a \\<^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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D (G f \\<^sub>D \\<^sub>0 a), \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, \\<^sub>0' a \\<^sub>D \\<^sub>0 a] = (\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0' a, \\<^sub>0 a]" proof - have "\\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a \\<^sub>D \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a \\<^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>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D (G f \\<^sub>D \\<^sub>0 a), \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D \\<^sub>0' a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, \\<^sub>0' a \\<^sub>D \\<^sub>0 a] = \\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\ \<^bold>\ \<^bold>\G f\<^bold>\, \<^bold>\\\<^sub>0 a\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' a\<^bold>\, \<^bold>\\\<^sub>0 a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\, \<^bold>\G f\<^bold>\, \<^bold>\\\<^sub>0 a\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0' a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0 a\<^bold>\) \<^bold>\ ((\<^bold>\\\<^sub>0' b\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[ \<^bold>\G f\<^bold>\, \<^bold>\\\<^sub>0 a\<^bold>\, \<^bold>\\\<^sub>0' a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\\<^sub>0 a\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\, \<^bold>\G f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 a\<^bold>\, \<^bold>\\\<^sub>0' a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0 a\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\ \<^bold>\ (\<^bold>\G f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 a\<^bold>\), \<^bold>\\\<^sub>0' a\<^bold>\, \<^bold>\\\<^sub>0 a\<^bold>\\<^bold>] \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\, \<^bold>\G f\<^bold>\, \<^bold>\\\<^sub>0 a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0' a\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 a\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\ \<^bold>\ \<^bold>\G f\<^bold>\, \<^bold>\\\<^sub>0 a\<^bold>\, \<^bold>\\\<^sub>0' a\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 a\<^bold>\\<^bold>]\" using a b f ide_f D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by auto also have "... = \(\<^bold>\\\<^sub>0' b\<^bold>\ \<^bold>\ \<^bold>\G f\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 a\<^bold>\, \<^bold>\\\<^sub>0' a\<^bold>\, \<^bold>\\\<^sub>0 a\<^bold>\\<^bold>]\" using a b f ide_f by (intro EV.eval_eqI, auto) also have "... = (\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0' a, \\<^sub>0 a]" using a b f ide_f D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by auto finally show ?thesis by blast qed thus ?thesis using D.comp_assoc by simp qed also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f, G\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, G\<^sub>0 a, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, F\<^sub>0 a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D F\<^sub>0 a)) \\<^sub>D ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D F\<^sub>0 a) \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^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 a chosen_adjoint_equivalence by simp have "((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \ a \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0' a, \\<^sub>0 a]) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>0 a \\<^sub>D \ a) = (\\<^sub>0' b \\<^sub>D G f) \\<^sub>D (\ a \\<^sub>D \\<^sub>0 a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0' a, \\<^sub>0 a] \\<^sub>D (\\<^sub>0 a \\<^sub>D \ a)" using a b f ide_f D.whisker_left by auto also have "... = (\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a]" using triangle_left by simp finally have "((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \ a \\<^sub>D \\<^sub>0 a) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a, \\<^sub>0' a, \\<^sub>0 a]) \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>0 a \\<^sub>D \ a) = (\\<^sub>0' b \\<^sub>D G f) \\<^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>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D (\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D (\\<^sub>D[\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D F\<^sub>0 a)) \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "\\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f, G\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, G\<^sub>0 a, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, F\<^sub>0 a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D F\<^sub>0 a) = \\<^sub>D[\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a]" proof - have "\\<^sub>D[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D[\\<^sub>0' b, G f, G\<^sub>0 a] \\<^sub>D \\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b \\<^sub>D G f, G\<^sub>0 a, \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 a] \\<^sub>D \\<^sub>D[\\<^sub>0 a]) \\<^sub>D \\<^sub>D[\\<^sub>0' b \\<^sub>D G f, \\<^sub>0 a, F\<^sub>0 a]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, G f, \\<^sub>0 a] \\<^sub>D F\<^sub>0 a) = \\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\, \<^bold>\G f\<^bold>\, \<^bold>\\\<^sub>0 a\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\\\<^sub>0' b\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\G f\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\\<^sub>0 a\<^bold>\) \<^bold>\ (\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\, \<^bold>\G f\<^bold>\, \<^bold>\G\<^sub>0 a\<^bold>\\<^sub>0\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0 a\<^bold>\) \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\ \<^bold>\ \<^bold>\G f\<^bold>\, \<^bold>\G\<^sub>0 a\<^bold>\\<^sub>0, \<^bold>\\\<^sub>0 a\<^bold>\\<^bold>] \<^bold>\ ((\<^bold>\\\<^sub>0' b\<^bold>\ \<^bold>\ \<^bold>\G f\<^bold>\) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0 a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0 a\<^bold>\\<^bold>]) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\ \<^bold>\ \<^bold>\G f\<^bold>\, \<^bold>\\\<^sub>0 a\<^bold>\, \<^bold>\F\<^sub>0 a\<^bold>\\<^sub>0\<^bold>]) \<^bold>\ (\<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\, \<^bold>\G f\<^bold>\, \<^bold>\\\<^sub>0 a\<^bold>\\<^bold>] \<^bold>\ \<^bold>\F\<^sub>0 a\<^bold>\\<^sub>0)\" using a b f ide_f 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 auto also have "... = \\<^bold>\\<^bold>[\<^bold>\\\<^sub>0' b\<^bold>\ \<^bold>\ \<^bold>\G f\<^bold>\ \<^bold>\ \<^bold>\\\<^sub>0 a\<^bold>\\<^bold>]\" using a b f ide_f by (intro EV.eval_eqI, auto) also have "... = \\<^sub>D[\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a]" using a b f ide_f 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 auto finally show ?thesis by blast qed thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D (\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)))) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>D[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f] \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "\\<^sub>D[\\<^sub>0' b \\<^sub>D G f \\<^sub>D \\<^sub>0 a] \\<^sub>D ((\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D F\<^sub>0 a) = (\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>D[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f]" using a b f ide_f D.comp_assoc \.iso_map\<^sub>1_ide D.runit_naturality [of "(\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f]"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = ((\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f]) \\<^sub>D \\<^sub>D[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f]) \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D ((\\<^sub>0' b \\<^sub>D \\<^sub>1 f) \\<^sub>D (\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f))) = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f]" using a b f ide_f D.comp_arr_inv' D.comp_arr_dom \.iso_map\<^sub>1_ide D.whisker_left [of "\\<^sub>0' b" "\\<^sub>1 f" "D.inv (\\<^sub>1 f)"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>D[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f] \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a)) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "(\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0' b, \\<^sub>0 b, F f] \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b, F f]) \\<^sub>D \\<^sub>D[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f] = \\<^sub>D[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f]" using a b f ide_f D.comp_inv_arr' D.comp_cod_arr by auto thus ?thesis using D.comp_assoc by simp qed also have "... = (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>D[F\<^sub>0 b \\<^sub>D F f] \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D F\<^sub>0 a)" proof - have "\\<^sub>D[(\\<^sub>0' b \\<^sub>D \\<^sub>0 b) \\<^sub>D F f] \\<^sub>D ((\ b \\<^sub>D F f) \\<^sub>D F\<^sub>0 a) = (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>D[F\<^sub>0 b \\<^sub>D F f]" using a b f ide_f D.runit_naturality [of "\ b \\<^sub>D F f"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>D[F f]" using a b f ide_f D.runit_naturality [of "\\<^sub>D\<^sup>-\<^sup>1[F f]"] by auto finally show "\'\.map\<^sub>1 f \\<^sub>D (F f \\<^sub>D \ a) = (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>D[F f]" by blast qed qed lemma unit_is_invertible_modification: shows "invertible_modification 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 \\<^sub>F.map\<^sub>0 \\<^sub>F.map\<^sub>1 \'\.map\<^sub>0 \'\.map\<^sub>1 \" .. interpretation \: invertible_modification V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D G \\<^sub>G G \\<^sub>G \\'.map\<^sub>0 \\'.map\<^sub>1 \\<^sub>G.map\<^sub>0 \\<^sub>G.map\<^sub>1 \ proof show "\a. C.obj a \ \\ a : \\'.map\<^sub>0 a \\<^sub>D G\<^sub>0 a\" using counit_in_hom \\'.map\<^sub>0_def by simp show "\a. C.obj a \ D.iso (\ a)" by simp show "\f a b. \\f : a \\<^sub>C b\; C.ide f\ \ (\\<^sub>D\<^sup>-\<^sup>1[G f] \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G f \\<^sub>D \ a) = (\ b \\<^sub>D G f) \\<^sub>D \\'.map\<^sub>1 f" proof - fix f a b assume f: "\f : a \\<^sub>C b\" and ide_f: "C.ide f" have a: "C.obj a" and b: "C.obj b" using f by auto have "(\ b \\<^sub>D G f) \\<^sub>D \\'.map\<^sub>1 f = (\ b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, G f] \\<^sub>D (\\<^sub>0 b \\<^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[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" unfolding \\'.map\<^sub>1_def map\<^sub>1_def using a b f D.comp_assoc by auto also have "... = (\ b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, G f] \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D G f \\<^sub>D \ a) \\<^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 b \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^sub>0' a]) \\<^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 b \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^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) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "\\<^sub>0 b \\<^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>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D G f \\<^sub>D \ a) \\<^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 b \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^sub>0' a]) \\<^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 b \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^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) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a)" proof - have "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 a b f ide_f \.iso_map\<^sub>1_ide by (intro D.seqI) auto thus ?thesis using a b f D.whisker_left [of "\\<^sub>0 b"] by fastforce qed thus ?thesis using D.comp_assoc by simp qed also have "... = (\ b \\<^sub>D G f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, G f] \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D G f \\<^sub>D \ a) \\<^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 b \\<^sub>D \\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^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 b \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^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) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "(\\<^sub>0 b \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D (\\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f)) \\<^sub>D \\<^sub>0' a) = \\<^sub>0 b \\<^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 a b f ide_f D.whisker_left [of "\\<^sub>0 b"] \.iso_map\<^sub>1_ide by auto also have "... = \\<^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' b, \\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a]" using a b f ide_f \.iso_map\<^sub>1_ide D.assoc_naturality [of "\\<^sub>0' b" "D.inv (\\<^sub>1 f)" "\\<^sub>0' a"] by auto also have "... = (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a])" using a b f ide_f D.whisker_left [of "\\<^sub>0 b"] \.iso_map\<^sub>1_ide by auto finally have "(\\<^sub>0 b \\<^sub>D \\<^sub>D[\\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D (\\<^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 D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>0 b \\<^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 also have "... = (\ b \\<^sub>D G f) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, G f] \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D \\<^sub>D[G f])) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^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 b \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^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) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "(\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^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 \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)" using a b f ide_f D.whisker_left \.iso_map\<^sub>1_ide by auto thus ?thesis using D.comp_assoc by simp qed also have "... = (\ b \\<^sub>D G f) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, G f \\<^sub>D G\<^sub>0 a] \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D G f \\<^sub>D \ a)) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^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 b \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^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) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, G f] \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D \\<^sub>D[G f]) = ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, G f \\<^sub>D G\<^sub>0 a]" using a b f ide_f D.assoc'_naturality [of "\\<^sub>0 b" "\\<^sub>0' b" "\\<^sub>D[G f]"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = (\ b \\<^sub>D G f) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D (\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a] \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a))) \\<^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 b \\<^sub>D \\<^sub>D[\\<^sub>0' b, \\<^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) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, G f \\<^sub>D G\<^sub>0 a] \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D G f \\<^sub>D \ a) = ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a]" using a b f ide_f D.assoc'_naturality [of "\\<^sub>0 b" "\\<^sub>0' b" "G f \\<^sub>D \ a"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = ((\ b \\<^sub>D G f) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f])) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 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 \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a]) \\<^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 b \\<^sub>D (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a] \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) = ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 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]" using a b f ide_f \.iso_map\<^sub>1_ide D.assoc'_naturality [of "\\<^sub>0 b" "\\<^sub>0' b" "\\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D ((\ b \\<^sub>D G f \\<^sub>D G\<^sub>0 a) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D G f \\<^sub>D \ a)) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 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 \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a]) \\<^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 b \\<^sub>D (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "(\ b \\<^sub>D G f) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f]) = \ b \\<^sub>D \\<^sub>D[G f]" using a b f ide_f D.comp_arr_dom D.comp_cod_arr D.interchange [of "\ b" "\\<^sub>0 b \\<^sub>D \\<^sub>0' b" "G f" "\\<^sub>D[G f]"] by simp also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\ b \\<^sub>D G f \\<^sub>D G\<^sub>0 a)" using a b f ide_f D.comp_arr_dom D.comp_cod_arr D.interchange [of "G\<^sub>0 b" "\ b" "\\<^sub>D[G f]" "G f \\<^sub>D G\<^sub>0 a"] by auto finally have "(\ b \\<^sub>D G f) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f]) = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (\ b \\<^sub>D G f \\<^sub>D G\<^sub>0 a)" by blast thus ?thesis using D.comp_assoc by simp qed also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D ((\ b \\<^sub>D G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 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 \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a]) \\<^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 b \\<^sub>D (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "(\ b \\<^sub>D G f \\<^sub>D G\<^sub>0 a) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D G f \\<^sub>D \ a) = \ b \\<^sub>D G f \\<^sub>D \ a" using a b f ide_f D.comp_arr_dom D.comp_cod_arr D.interchange [of "\ b" "\\<^sub>0 b \\<^sub>D \\<^sub>0' b" "G f \\<^sub>D G\<^sub>0 a" "G f \\<^sub>D \ a"] by auto also have "... = (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D (\ b \\<^sub>D G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a)" using a b f ide_f D.comp_arr_dom D.comp_cod_arr D.interchange [of "G\<^sub>0 b" "\ b" "G f \\<^sub>D \ a" "G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a"] by auto finally have "(\ b \\<^sub>D G f \\<^sub>D G\<^sub>0 a) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D G f \\<^sub>D \ a) = (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^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 also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^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 \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a]) \\<^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 b \\<^sub>D (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have 1: "D.seq (G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a) (\\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a))" using a b f ide_f \.iso_map\<^sub>1_ide by (intro D.seqI D.hseqI') auto have 2: "D.seq (\\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) ((\\<^sub>0 b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a)" using a b f ide_f \.iso_map\<^sub>1_ide by (intro D.seqI D.hseqI') auto have "(\ b \\<^sub>D G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) = \ b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)" using 1 a b f ide_f D.comp_arr_dom D.comp_cod_arr \.iso_map\<^sub>1_ide D.interchange [of "\ b" "\\<^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] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)"] by auto also have "... = G\<^sub>0 b \\<^sub>D \ b \\<^sub>D (\\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^sub>D ((\\<^sub>0 b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a)" using a b f ide_f \.iso_map\<^sub>1_ide D.comp_arr_dom D.comp_cod_arr by auto also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^sub>D (\ b \\<^sub>D (\\<^sub>0 b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a)" using 2 a b f ide_f \.iso_map\<^sub>1_ide D.comp_assoc D.interchange [of "G\<^sub>0 b" "\ b" "\\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (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 "(\ b \\<^sub>D G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^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 "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^sub>D (G\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F f, \\<^sub>0' a]) \\<^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) \\<^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 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^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 \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a]) \\<^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 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[\\<^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) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D[F\<^sub>0 b, F f, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "\ b \\<^sub>D (\\<^sub>0 b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a = (G\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F f, \\<^sub>0' a]) \\<^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) \\<^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 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a])" proof - have "\ b \\<^sub>D (\\<^sub>0 b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a = (G\<^sub>0 b \\<^sub>D \ b \\<^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 b \\<^sub>D F f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a]" using a b f ide_f D.comp_arr_dom D.comp_cod_arr D.hcomp_reassoc(1) [of "\\<^sub>0 b" "F f" "\\<^sub>0' a"] by auto also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F f, \\<^sub>0' a]) \\<^sub>D (\ b \\<^sub>D \\<^sub>0 b \\<^sub>D F f \\<^sub>D \\<^sub>0' a) \\<^sub>D ((\\<^sub>0 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a])" using a b f ide_f D.comp_arr_dom D.comp_cod_arr D.assoc'_is_natural_1 D.interchange [of "G\<^sub>0 b" "\ b \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>0' 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) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a]"] 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[\\<^sub>0 b, F f, \\<^sub>0' a]"] by fastforce also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F f, \\<^sub>0' a]) \\<^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) \\<^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 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a])" using a b f ide_f D.hcomp_reassoc(2) [of "\ b" "\\<^sub>0 b" "F f \\<^sub>D \\<^sub>0' a"] by auto finally show ?thesis by blast qed moreover have "\\<^sub>0 b \\<^sub>D (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a = (\\<^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[\\<^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) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D[F\<^sub>0 b, F f, \\<^sub>0' a])" proof - have "\\<^sub>0 b \\<^sub>D (\ b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a = (\\<^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 b \\<^sub>D \ b \\<^sub>D F f \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D[F\<^sub>0 b, F f, \\<^sub>0' a])" using a b f ide_f D.hcomp_reassoc(1) [of "\ b" "F f" "\\<^sub>0' a"] D.whisker_left [of "\\<^sub>0 b"] by auto also have "... = (\\<^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[\\<^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) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D[F\<^sub>0 b, F f, \\<^sub>0' a])" using a b f ide_f D.hcomp_reassoc(2) [of "\\<^sub>0 b" "\ b" "F f \\<^sub>D \\<^sub>0' a"] by auto finally show ?thesis by blast qed ultimately show ?thesis using D.comp_assoc by simp qed also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^sub>D (G\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F f, \\<^sub>0' a]) \\<^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) \\<^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)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D[F\<^sub>0 b, F f, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "\\<^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 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^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 \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a]) \\<^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 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[\\<^sub>0 b, \\<^sub>0' b \\<^sub>D \\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a] = \\<^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>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 b \\<^sub>D \\<^sub>0' b) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^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 \\<^sub>D[\\<^sub>0' b, \\<^sub>0 b \\<^sub>D F f, \\<^sub>0' a]) \\<^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 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[\\<^sub>0 b, \\<^sub>0' b \\<^sub>D \\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a] = \\<^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 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>\\<^sup>-\<^sup>1\<^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 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 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>\\\<^sub>0 b\<^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>\\<^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 a b f ide_f D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by auto also have "... = \\<^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 a b f ide_f by (intro EV.eval_eqI, auto) also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, \\<^sub>0' b, \\<^sub>0 b] \\<^sub>D F f \\<^sub>D \\<^sub>0' a" using a b f ide_f D.\_def D.\'.map_ide_simp D.VVV.ide_char D.VVV.arr_char D.VV.ide_char D.VV.arr_char by auto finally show ?thesis by blast qed thus ?thesis using D.comp_assoc by auto qed also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^sub>D ((G\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F f, \\<^sub>0' a]) \\<^sub>D (\\<^sub>D[G\<^sub>0 b, \\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a] \\<^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>D\<^sup>-\<^sup>1[\\<^sub>0 b, F\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D[F\<^sub>0 b, F f, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a]) \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "((\ 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) = \\<^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 b chosen_adjoint_equivalence by simp have "((\ 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) = (\ 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 a b f ide_f D.whisker_right by auto also have "... = \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b] \\<^sub>D \\<^sub>D[\\<^sub>0 b] \\<^sub>D F f \\<^sub>D \\<^sub>0' a" using triangle_left by simp finally show ?thesis by blast qed thus ?thesis using D.comp_assoc by simp qed also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D ((G\<^sub>0 b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0 b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a]) \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "(G\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F f, \\<^sub>0' a]) \\<^sub>D (\\<^sub>D[G\<^sub>0 b, \\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a] \\<^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>D\<^sup>-\<^sup>1[\\<^sub>0 b, F\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D[F\<^sub>0 b, F f, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] = \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0 b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a]" proof - have "(G\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[\\<^sub>0 b, F f, \\<^sub>0' a]) \\<^sub>D (\\<^sub>D[G\<^sub>0 b, \\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a] \\<^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>D\<^sup>-\<^sup>1[\\<^sub>0 b, F\<^sub>0 b, F f \\<^sub>D \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D[F\<^sub>0 b, F f, \\<^sub>0' a]) \\<^sub>D (\\<^sub>0 b \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f] \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D[\\<^sub>0 b, F f, \\<^sub>0' a] = \(\<^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>\\<^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>\\<^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>\\<^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>\\\<^sub>0 b\<^bold>\ \<^bold>\ \<^bold>\\<^bold>[\<^bold>\F\<^sub>0 b\<^bold>\\<^sub>0, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' a\<^bold>\\<^bold>]) \<^bold>\ (\<^bold>\\\<^sub>0 b\<^bold>\ \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[\<^bold>\F f\<^bold>\\<^bold>] \<^bold>\ \<^bold>\\\<^sub>0' a\<^bold>\) \<^bold>\ \<^bold>\\<^bold>[\<^bold>\\\<^sub>0 b\<^bold>\, \<^bold>\F f\<^bold>\, \<^bold>\\\<^sub>0' a\<^bold>\\<^bold>]\" using a b f ide_f 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 auto also have "... = \\<^bold>\\<^sup>-\<^sup>1\<^bold>[(\<^bold>\\\<^sub>0 b\<^bold>\ \<^bold>\ \<^bold>\F f\<^bold>\) \<^bold>\ \<^bold>\\\<^sub>0' a\<^bold>\\<^bold>]\" using a b f ide_f by (intro EV.eval_eqI, auto) also have "... = \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0 b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a]" using a b f ide_f 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 by auto finally show ?thesis by blast qed thus ?thesis using D.comp_assoc by simp qed also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D ((D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "(G\<^sub>0 b \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[(\\<^sub>0 b \\<^sub>D F f) \\<^sub>D \\<^sub>0' a] = \\<^sub>D\<^sup>-\<^sup>1[G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)" using a b f ide_f \.iso_map\<^sub>1_ide D.lunit'_naturality [of "\\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D (D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a)"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" proof - have "((D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>0' a) \\<^sub>D (\\<^sub>1 f \\<^sub>D \\<^sub>0' a)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a] = ((D.inv (\\<^sub>1 f) \\<^sub>D \\<^sub>1 f) \\<^sub>D \\<^sub>0' a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" using a b f ide_f \.iso_map\<^sub>1_ide D.whisker_right [of "\\<^sub>0' a"] by simp also have "... = \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a]" using a b f ide_f \.iso_map\<^sub>1_ide D.comp_inv_arr' D.comp_cod_arr by auto finally show ?thesis using D.comp_assoc by simp qed also have "... = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a]" proof - have "\\<^sub>D\<^sup>-\<^sup>1[G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a] \\<^sub>D \\<^sub>D[G f, \\<^sub>0 a, \\<^sub>0' a] \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f, \\<^sub>0 a, \\<^sub>0' a] = \\<^sub>D\<^sup>-\<^sup>1[G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a]" using a b f ide_f D.comp_arr_inv' D.comp_arr_dom by auto thus ?thesis using D.comp_assoc by simp qed also have "... = ((G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f \\<^sub>D G\<^sub>0 a]) \\<^sub>D (G f \\<^sub>D \ a)" proof - have "(G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G\<^sub>0 b \\<^sub>D G f \\<^sub>D \ a) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f \\<^sub>D \\<^sub>0 a \\<^sub>D \\<^sub>0' a] = (G\<^sub>0 b \\<^sub>D \\<^sub>D[G f]) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[G f \\<^sub>D G\<^sub>0 a] \\<^sub>D (G f \\<^sub>D \ a)" using a b f ide_f D.lunit'_naturality [of "G f \\<^sub>D \ a"] by auto thus ?thesis using D.comp_assoc by simp qed also have "... = (\\<^sub>D\<^sup>-\<^sup>1[G f] \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G f \\<^sub>D \ a)" using a b f ide_f D.lunit'_naturality [of "\\<^sub>D[G f]"] by auto finally show "(\\<^sub>D\<^sup>-\<^sup>1[G f] \\<^sub>D \\<^sub>D[G f]) \\<^sub>D (G f \\<^sub>D \ a) = (\ b \\<^sub>D G f) \\<^sub>D \\'.map\<^sub>1 f" by simp qed qed lemma counit_is_invertible_modification: shows "invertible_modification V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D G \\<^sub>G G \\<^sub>G \\'.map\<^sub>0 \\'.map\<^sub>1 \\<^sub>G.map\<^sub>0 \\<^sub>G.map\<^sub>1 \" .. end end diff --git a/thys/Bicategory/Prebicategory.thy b/thys/Bicategory/Prebicategory.thy --- a/thys/Bicategory/Prebicategory.thy +++ b/thys/Bicategory/Prebicategory.thy @@ -1,2970 +1,2960 @@ (* Title: PreBicategory Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) text \ The objective of this section is to construct a formalization of bicategories that is compatible with our previous formulation of categories \cite{Category3-AFP} and that permits us to carry over unchanged as much of the work done on categories as possible. For these reasons, we conceive of a bicategory in what might be regarded as a somewhat unusual fashion. Rather than a traditional development, which would typically define a bicategory to be essentially ``a `category' whose homs themselves have the structure of categories,'' here we regard a bicategory as ``a (vertical) category that has been equipped with a suitable (horizontal) weak composition.'' Stated another way, we think of a bicategory as a generalization of a monoidal category in which the tensor product is a partial operation, rather than a total one. Our definition of bicategory can thus be summarized as follows: a bicategory is a (vertical) category that has been equipped with idempotent endofunctors \src\ and \trg\ that assign to each arrow its ``source'' and ``target'' subject to certain commutativity constraints, a partial binary operation \\\ of horizontal composition that is suitably functorial on the ``hom-categories'' determined by the assignment of sources and targets, ``associativity'' isomorphisms \\\[f, g, h] : (f \ g) \ h \ f \ (g \ h)\\ for each horizontally composable triple of vertical identities \f\, \g\, \h\, subject to the usual naturality and coherence conditions, and for each ``object'' \a\ (defined to be an arrow that is its own source and target) a ``unit isomorphism'' \\\[a] : a \ a \ a\\. As is the case for monoidal categories, the unit isomorphisms and associator isomorphisms together enable a canonical definition of left and right ``unit'' isomorphisms \\\[f] : a \ f \ f\\ and \\\[f] : f \ a \ f\\ when \f\ is a vertical identity horizontally composable on the left or right by \a\, and it can be shown that these are the components of natural transformations. The definition of bicategory just sketched shares with a more traditional version the requirement that assignments of source and target are given as basic data, and these assignments determine horizontal composability in the sense that arrows \\\ and \\\ are composable if the chosen source of \\\ coincides with the chosen target of \\\. Thus it appears, at least on its face, that composability of arrows depends on an assignment of sources and targets. We are interested in establishing whether this is essential or whether bicategories can be formalized in a completely ``object-free'' fashion. It turns out that we can obtain such an object-free formalization through a rather direct generalization of the approach we used in the formalization of categories. Specifically, we define a \emph{weak composition} to be a partial binary operation \\\ on the arrow type of a ``vertical'' category \V\, such that the domain of definition of this operation is itself a category (of ``horizontally composable pairs of arrows''), the operation is functorial, and it is subject to certain matching conditions which include those satisfied by a category. From the axioms for a weak composition we can prove the existence of ``hom-categories'', which are subcategories of \V\ consisting of arrows horizontally composable on the left or right by a specified vertical identity. A \emph{weak unit} is defined to be a vertical identity \a\ such that \a \ a \ a\ and is such that the mappings \a \ \\ and \\ \ a\ are fully faithful endofunctors of the subcategories of \V\ consisting of the arrows for which they are defined. We define the \emph{sources} of an arrow \\\ to be the weak units that are horizontally composable with \\\ on the right, and the \emph{targets} of \\\ to be the weak units that are horizontally composable with \\\ on the left. An \emph{associative weak composition} is defined to be a weak composition that is equipped with ``associator'' isomorphisms \\\[f, g, h] : (f \ g) \ h \ f \ (g \ h)\\ for horizontally composable vertical identities \f\, \g\, \h\, subject to the usual naturality and coherence conditions. A \emph{prebicategory} is defined to be an associative weak composition for which every arrow has a source and a target. We show that the sets of sources and targets of each arrow in a prebicategory is an isomorphism class of weak units, and that horizontal composability of arrows \\\ and \\\ is characterized by the set of sources of \\\ being equal to the set of targets of \\\. We show that prebicategories are essentially ``bicategories without objects''. Given a prebicategory, we may choose an arbitrary representative of each isomorphism class of weak units and declare these to be ``objects'' (this is analogous to choosing a particular unit object in a monoidal category). For each object we may choose a particular \emph{unit isomorphism} \\\[a] : a \ a \ a\\. This choice, together with the associator isomorphisms, enables a canonical definition of left and right unit isomorphisms \\\[f] : a \ f \ f\\ and \\\[f] : f \ a \ f\\ when \f\ is a vertical identity horizontally composable on the left or right by \a\, and it can be shown that these are the components of natural isomorphisms. We may then define ``the source'' of an arrow to be the chosen representative of the set of its sources and ``the target'' to be the chosen representative of the set of its targets. We show that the resulting structure is a bicategory, in which horizontal composability as given by the weak composition coincides with the ``syntactic'' version determined by the chosen sources and targets. Conversely, a bicategory determines a prebicategory, essentially by forgetting the sources, targets and unit isomorphisms. These results make it clear that the assignment of sources and targets to arrows in a bicategory is basically a convenience and that horizontal composability of arrows is not dependent on a particular choice. \ theory Prebicategory imports Category3.EquivalenceOfCategories Category3.Subcategory IsomorphismClass begin section "Weak Composition" text \ In this section we define a locale \weak_composition\, which formalizes a functorial operation of ``horizontal'' composition defined on an underlying ``vertical'' category. The definition is expressed without the presumption of the existence of any sort of ``objects'' that determine horizontal composability. Rather, just as we did in showing that the @{locale partial_magma} locale supported the definition of ``identity arrow'' as a kind of unit for vertical composition which ultimately served as a basis for the definition of ``domain'' and ``codomain'' of an arrow, here we show that the \weak_composition\ locale supports a definition of \emph{weak unit} for horizontal composition which can ultimately be used to define the \emph{sources} and \emph{targets} of an arrow with respect to horizontal composition. In particular, the definition of weak composition involves axioms that relate horizontal and vertical composability. As a consequence of these axioms, for any fixed arrow \\\, the sets of arrows horizontally composable on the left and on the right with \\\ form subcategories with respect to vertical composition. We define the sources of \\\ to be the weak units that are composable with \\\ on the right, and the targets of \\\ to be the weak units that are composable with \\\ on the left. Weak units are then characterized as arrows that are members of the set of their own sources (or, equivalently, of their own targets). \ subsection "Definition" locale weak_composition = category V + VxV: product_category V V + VoV: subcategory VxV.comp \\\\. fst \\ \ snd \\ \ null\ + "functor" VoV.comp V \\\\. fst \\ \ snd \\\ for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) + assumes left_connected: "seq \ \' \ \ \ \ \ null \ \' \ \ \ null" and right_connected: "seq \ \' \ \ \ \ \ null \ \ \ \' \ null" and match_1: "\ \ \ \ \ null; (\ \ \) \ \ \ null \ \ \ \ \ \ null" and match_2: "\ \ \ (\ \ \) \ null; \ \ \ \ null \ \ \ \ \ \ null" and match_3: "\ \ \ \ \ null; \ \ \ \ null \ \ (\ \ \) \ \ \ null" and match_4: "\ \ \ \ \ null; \ \ \ \ null \ \ \ \ (\ \ \) \ null" begin text \ We think of the arrows of the vertical category as ``2-cells'' and the vertical identities as ``1-cells''. In the formal development, the predicate @{term arr} (``arrow'') will have its normal meaning with respect to the vertical composition, hence @{term "arr \"} will mean, essentially, ``\\\ is a 2-cell''. This is somewhat unfortunate, as it is traditional when discussing bicategories to use the term ``arrow'' to refer to the 1-cells. However, we are trying to carry over all the formalism that we have already developed for categories and apply it to bicategories with as little change and redundancy as possible. It becomes too confusing to try to repurpose the name @{term arr} to mean @{term ide} and to introduce a replacement for the name @{term arr}, so we will simply tolerate the situation. In informal text, we will prefer the terms ``2-cell'' and ``1-cell'' over (vertical) ``arrow'' and ``identity'' when there is a chance for confusion. We do, however, make the following adjustments in notation for @{term in_hom} so that it is distinguished from the notion @{term in_hhom} (``in horizontal hom'') to be introduced subsequently. \ no_notation in_hom ("\_ : _ \ _\") notation in_hom ("\_ : _ \ _\") lemma is_partial_magma: shows "partial_magma H" proof show "\!n. \f. n \ f = n \ f \ n = n" proof show 1: "\f. null \ f = null \ f \ null = null" using is_extensional VoV.inclusion VoV.arr_char by force show "\n. \f. n \ f = n \ f \ n = n \ n = null" using 1 VoV.arr_char is_extensional not_arr_null by metis qed qed interpretation H: partial_magma H using is_partial_magma by auto text \ Either \match_1\ or \match_2\ seems essential for the next result, which states that the nulls for the horizontal and vertical compositions coincide. \ lemma null_agreement [simp]: shows "H.null = null" by (metis VoV.inclusion VxV.not_arr_null match_1 H.comp_null(1)) lemma composable_implies_arr: assumes "\ \ \ \ null" shows "arr \" and "arr \" using assms is_extensional VoV.arr_char VoV.inclusion by auto lemma hcomp_null [simp]: shows "null \ \ = null" and "\ \ null = null" using H.comp_null by auto lemma hcomp_simps\<^sub>W\<^sub>C [simp]: assumes "\ \ \ \ null" shows "arr (\ \ \)" and "dom (\ \ \) = dom \ \ dom \" and "cod (\ \ \) = cod \ \ cod \" using assms preserves_arr preserves_dom preserves_cod VoV.arr_char VoV.inclusion VoV.dom_simp VoV.cod_simp by force+ lemma ide_hcomp\<^sub>W\<^sub>C: assumes "ide f" and "ide g" and "g \ f \ null" shows "ide (g \ f)" using assms preserves_ide VoV.ide_char by force lemma hcomp_in_hom\<^sub>W\<^sub>C [intro]: assumes "\ \ \ \ null" shows "\\ \ \ : dom \ \ dom \ \ cod \ \ cod \\" using assms by auto text \ Horizontal composability of arrows is determined by horizontal composability of their domains and codomains (defined with respect to vertical composition). \ lemma hom_connected: shows "\ \ \ \ null \ dom \ \ \ \ null" and "\ \ \ \ null \ \ \ dom \ \ null" and "\ \ \ \ null \ cod \ \ \ \ null" and "\ \ \ \ null \ \ \ cod \ \ null" proof - show "\ \ \ \ null \ dom \ \ \ \ null" using left_connected [of \ "dom \" \] composable_implies_arr arr_dom_iff_arr by force show "\ \ \ \ null \ cod \ \ \ \ null" using left_connected [of "cod \" \ \] composable_implies_arr arr_cod_iff_arr by force show "\ \ \ \ null \ \ \ dom \ \ null" using right_connected [of \ "dom \" \] composable_implies_arr arr_dom_iff_arr by force show "\ \ \ \ null \ \ \ cod \ \ null" using right_connected [of "cod \" \ \] composable_implies_arr arr_cod_iff_arr by force qed lemma isomorphic_implies_equicomposable: assumes "f \ g" shows "\ \ f \ null \ \ \ g \ null" and "f \ \ \ null \ g \ \ \ null" using assms isomorphic_def hom_connected by auto lemma interchange: assumes "seq \ \" and "seq \ \" shows "(\ \ \) \ (\ \ \) = (\ \ \) \ (\ \ \)" proof - have "\ \ \ = null \ ?thesis" by (metis assms comp_null(2) dom_comp hom_connected(1-2)) moreover have "\ \ \ \ null \ ?thesis" proof - assume \\: "\ \ \ \ null" have 1: "VoV.arr (\, \)" using \\ VoV.arr_char by auto have \\: "(\, \) \ VoV.hom (VoV.cod (\, \)) (VoV.cod (\, \))" proof - have "VoV.arr (\, \)" using assms 1 hom_connected VoV.arr_char by (elim seqE, auto, metis) thus ?thesis using assms \\ VoV.dom_char VoV.cod_char by fastforce qed show ?thesis proof - have "VoV.seq (\, \) (\, \)" using assms 1 \\ \\ VoV.seqI by blast thus ?thesis using assms 1 \\ \\ VoV.comp_char preserves_comp [of "(\, \)" "(\, \)"] VoV.seqI by fastforce qed qed ultimately show ?thesis by blast qed lemma paste_1: shows "\ \ \ = (cod \ \ \) \ (\ \ dom \)" using interchange composable_implies_arr comp_arr_dom comp_cod_arr hom_connected(2-3) by (metis comp_null(2)) lemma paste_2: shows "\ \ \ = (\ \ cod \) \ (dom \ \ \)" using interchange composable_implies_arr comp_arr_dom comp_cod_arr hom_connected(1,4) by (metis comp_null(2)) lemma whisker_left: assumes "seq \ \" and "ide f" shows "f \ (\ \ \) = (f \ \) \ (f \ \)" using assms interchange [of f f \ \] hom_connected by auto lemma whisker_right: assumes "seq \ \" and "ide f" shows "(\ \ \) \ f = (\ \ f) \ (\ \ f)" using assms interchange [of \ \ f f] hom_connected by auto subsection "Hom-Subcategories" definition left where "left \ \ \\. \ \ \ \ null" definition right where "right \ \ \\. \ \ \ \ null" lemma right_iff_left: shows "right \ \ \ left \ \" using right_def left_def by simp lemma left_respects_isomorphic: assumes "f \ g" shows "left f = left g" using assms isomorphic_implies_equicomposable left_def by auto lemma right_respects_isomorphic: assumes "f \ g" shows "right f = right g" using assms isomorphic_implies_equicomposable right_def by auto lemma left_iff_left_inv: assumes "iso \" shows "left \ \ \ left \ (inv \)" using assms left_def inv_in_hom hom_connected(2) hom_connected(4) [of \ "inv \"] by auto lemma right_iff_right_inv: assumes "iso \" shows "right \ \ \ right \ (inv \)" using assms right_def inv_in_hom hom_connected(1) hom_connected(3) [of "inv \" \] by auto lemma left_hom_is_subcategory: assumes "arr \" shows "subcategory V (left \)" using composable_implies_arr hom_connected(2,4) apply (unfold left_def, unfold_locales) apply auto by (metis cod_comp seqI) lemma right_hom_is_subcategory: assumes "arr \" shows "subcategory V (right \)" using composable_implies_arr hom_connected(1,3) apply (unfold right_def, unfold_locales) apply auto by (metis cod_comp seqI) abbreviation Left where "Left a \ subcategory.comp V (left a)" abbreviation Right where "Right a \ subcategory.comp V (right a)" text \ We define operations of composition on the left or right with a fixed 1-cell, and show that such operations are functorial in case that 1-cell is horizontally self-composable. \ definition H\<^sub>L where "H\<^sub>L g \ \\. g \ \" definition H\<^sub>R where "H\<^sub>R f \ \\. \ \ f" (* TODO: Why do the following fail when I use @{thm ...} *) text \ Note that \match_3\ and \match_4\ are required for the next results. \ lemma endofunctor_H\<^sub>L: assumes "ide g" and "g \ g \ null" shows "endofunctor (Left g) (H\<^sub>L g)" proof - interpret L: subcategory V \left g\ using assms left_hom_is_subcategory by simp have *: "\\. L.arr \ \ H\<^sub>L g \ = g \ \" using assms H\<^sub>L_def by simp have preserves_arr: "\\. L.arr \ \ L.arr (H\<^sub>L g \)" using assms * L.arr_char left_def match_4 by force show "endofunctor L.comp (H\<^sub>L g)" using assms * apply unfold_locales using H\<^sub>L_def L.arr_char L.null_char left_def apply force using preserves_arr apply blast apply (metis L.dom_simp L.not_arr_null L.null_char hcomp_simps\<^sub>W\<^sub>C(2) ide_char preserves_arr H\<^sub>L_def) apply (metis H\<^sub>L_def L.arrE L.cod_simp hcomp_simps\<^sub>W\<^sub>C(3) ide_char left_def preserves_arr) by (metis L.comp_def L.comp_simp L.seq_char hcomp_simps\<^sub>W\<^sub>C(1) whisker_left preserves_arr) qed lemma endofunctor_H\<^sub>R: assumes "ide f" and "f \ f \ null" shows "endofunctor (Right f) (H\<^sub>R f)" proof - interpret R: subcategory V \right f\ using assms right_hom_is_subcategory by simp have *: "\\. R.arr \ \ H\<^sub>R f \ = \ \ f" using assms H\<^sub>R_def by simp have preserves_arr: "\\. R.arr \ \ R.arr (H\<^sub>R f \)" using assms * R.arr_char right_def match_3 by force show "endofunctor R.comp (H\<^sub>R f)" using assms * apply unfold_locales using H\<^sub>R_def R.arr_char R.null_char right_def apply force using preserves_arr apply blast apply (metis R.dom_simp R.not_arr_null R.null_char hcomp_simps\<^sub>W\<^sub>C(2) ide_char preserves_arr H\<^sub>R_def) apply (metis H\<^sub>R_def R.arrE R.cod_simp hcomp_simps\<^sub>W\<^sub>C(3) ide_char right_def preserves_arr) by (metis R.comp_def R.comp_simp R.seq_char hcomp_simps\<^sub>W\<^sub>C(1) whisker_right preserves_arr) qed end locale left_hom = weak_composition V H + S: subcategory V \left \\ for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) and \ :: 'a + assumes arr_\: "arr \" begin no_notation in_hom ("\_ : _ \ _\") notation in_hom ("\_ : _ \ _\") notation S.comp (infixr "\\<^sub>S" 55) notation S.in_hom ("\_ : _ \\<^sub>S _\") lemma right_hcomp_closed: assumes "\\ : x \\<^sub>S y\" and "\\ : c \ d\" and "\ \ \ \ null" shows "\\ \ \ : x \ c \\<^sub>S y \ d\" using assms arr_\ S.arr_char S.dom_simp S.cod_simp left_def match_4 by (elim S.in_homE, intro S.in_homI) auto lemma interchange: assumes "S.seq \ \" and "S.seq \ \" and "\ \ \ \ null" shows "(\ \\<^sub>S \) \ (\ \\<^sub>S \) = (\ \ \) \\<^sub>S (\ \ \)" proof - have 1: "\ \ \ \ null" using assms hom_connected(1) [of \ \] hom_connected(2) [of \ \] hom_connected(3-4) S.dom_simp S.cod_simp by force have "(\ \\<^sub>S \) \ (\ \\<^sub>S \) = (\ \ \) \ (\ \ \)" using assms S.comp_char S.seq_char by metis also have "... = (\ \ \) \ (\ \ \)" using assms interchange S.seq_char S.arr_char by simp also have "... = (\ \ \) \\<^sub>S (\ \ \)" using assms 1 by (metis S.arr_char S.comp_char S.seq_char ext match_4 left_def) finally show ?thesis by blast qed lemma inv_char: assumes "S.arr \" and "iso \" shows "S.inverse_arrows \ (inv \)" and "S.inv \ = inv \" proof - have 1: "S.arr (inv \)" using assms S.arr_char left_iff_left_inv by (intro S.arrI) meson show "S.inv \ = inv \" using assms 1 S.inv_char S.iso_char by blast thus "S.inverse_arrows \ (inv \)" using assms 1 S.iso_char S.inv_is_inverse by metis qed lemma iso_char: assumes "S.arr \" shows "S.iso \ \ iso \" using assms S.iso_char inv_char by auto end locale right_hom = weak_composition V H + S: subcategory V \right \\ for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) and \ :: 'a + assumes arr_\: "arr \" begin no_notation in_hom ("\_ : _ \ _\") notation in_hom ("\_ : _ \ _\") notation S.comp (infixr "\\<^sub>S" 55) notation S.in_hom ("\_ : _ \\<^sub>S _\") lemma left_hcomp_closed: assumes "\\ : x \\<^sub>S y\" and "\\ : c \ d\" and "\ \ \ \ null" shows "\\ \ \ : c \ x \\<^sub>S d \ y\" using assms arr_\ S.arr_char S.dom_simp S.cod_simp right_def match_3 by (elim S.in_homE, intro S.in_homI) auto lemma interchange: assumes "S.seq \ \" and "S.seq \ \" and "\ \ \ \ null" shows "(\ \\<^sub>S \) \ (\ \\<^sub>S \) = (\ \ \) \\<^sub>S (\ \ \)" proof - have 1: "\ \ \ \ null" using assms hom_connected(1) [of \ \] hom_connected(2) [of \ \] hom_connected(3-4) S.dom_simp S.cod_simp by fastforce have "(\ \\<^sub>S \) \ (\ \\<^sub>S \) = (\ \ \) \ (\ \ \)" using assms S.comp_char S.seq_char by metis also have "... = (\ \ \) \ (\ \ \)" using assms interchange S.seq_char S.arr_char by simp also have "... = (\ \ \) \\<^sub>S (\ \ \)" using assms 1 by (metis S.arr_char S.comp_char S.seq_char ext match_3 right_def) finally show ?thesis by blast qed lemma inv_char: assumes "S.arr \" and "iso \" shows "S.inverse_arrows \ (inv \)" and "S.inv \ = inv \" proof - have 1: "S.arr (inv \)" using assms S.arr_char right_iff_right_inv by (intro S.arrI) meson show "S.inv \ = inv \" using assms 1 S.inv_char S.iso_char by blast thus "S.inverse_arrows \ (inv \)" using assms 1 S.iso_char S.inv_is_inverse by metis qed lemma iso_char: assumes "S.arr \" shows "S.iso \ \ iso \" using assms S.iso_char inv_char by auto end subsection "Weak Units" text \ We now define a \emph{weak unit} to be an arrow \a\ such that: \begin{enumerate} \item \a \ a\ is isomorphic to \a\ (and hence \a\ is a horizontally self-composable 1-cell). \item Horizontal composition on the left with \a\ is a fully faithful endofunctor of the subcategory of arrows that are composable on the left with \a\. \item Horizontal composition on the right with \a\ is fully faithful endofunctor of the subcategory of arrows that are composable on the right with \a\. \end{enumerate} \ context weak_composition begin definition weak_unit :: "'a \ bool" where "weak_unit a \ a \ a \ a \ fully_faithful_functor (Left a) (Left a) (H\<^sub>L a) \ fully_faithful_functor (Right a) (Right a) (H\<^sub>R a)" lemma weak_unit_self_composable: assumes "weak_unit a" shows "ide a" and "ide (a \ a)" and "a \ a \ null" proof - obtain \ where \: "\\ : a \ a \ a\ \ iso \" using assms weak_unit_def isomorphic_def by blast have 1: "arr \" using \ by blast show "ide a" using \ ide_cod by blast thus "ide (a \ a)" using \ ide_dom by force thus "a \ a \ null" using not_arr_null ideD(1) by metis qed lemma weak_unit_self_right: assumes "weak_unit a" shows "right a a" using assms weak_unit_self_composable right_def by simp lemma weak_unit_self_left: assumes "weak_unit a" shows "left a a" using assms weak_unit_self_composable left_def by simp lemma weak_unit_in_vhom: assumes "weak_unit a" shows "\a : a \ a\" using assms weak_unit_self_composable left_def by auto text \ If \a\ is a weak unit, then there exists a ``unit isomorphism'' \\\ : a \ a \ a\\. It need not be unique, but we may choose one arbitrarily. \ definition some_unit where "some_unit a \ SOME \. iso \ \ \\ : a \ a \ a\" lemma iso_some_unit: assumes "weak_unit a" shows "iso (some_unit a)" and "\some_unit a : a \ a \ a\" proof - let ?P = "\\. iso \ \ \\ : a \ a \ a\" have "\\. ?P \" using assms weak_unit_def by auto hence 1: "?P (some_unit a)" using someI_ex [of ?P] some_unit_def by simp show "iso (some_unit a)" using 1 by blast show "\some_unit a : a \ a \ a\" using 1 by blast qed text \ The \emph{sources} of an arbitrary arrow \\\ are the weak units that are composable with \\\ on the right. Similarly, the \emph{targets} of \\\ are the weak units that are composable with \\\ on the left. \ definition sources where "sources \ \ {a. weak_unit a \ \ \ a \ null}" lemma sourcesI [intro]: assumes "weak_unit a" and "\ \ a \ null" shows "a \ sources \" using assms sources_def by blast lemma sourcesD [dest]: assumes "a \ sources \" shows "ide a" and "weak_unit a" and "\ \ a \ null" using assms sources_def weak_unit_self_composable by auto definition targets where "targets \ \ {b. weak_unit b \ b \ \ \ null}" lemma targetsI [intro]: assumes "weak_unit b" and "b \ \ \ null" shows "b \ targets \" using assms targets_def by blast lemma targetsD [dest]: assumes "b \ targets \" shows "ide b" and "weak_unit b" and "b \ \ \ null" using assms targets_def weak_unit_self_composable by auto lemma sources_dom [simp]: assumes "arr \" shows "sources (dom \) = sources \" using assms hom_connected(1) by blast lemma sources_cod [simp]: assumes "arr \" shows "sources (cod \) = sources \" using assms hom_connected(3) by blast lemma targets_dom [simp]: assumes "arr \" shows "targets (dom \) = targets \" using assms hom_connected(2) by blast lemma targets_cod [simp]: assumes "arr \" shows "targets (cod \) = targets \" using assms hom_connected(4) by blast lemma weak_unit_iff_self_source: shows "weak_unit a \ a \ sources a" using weak_unit_self_composable by auto lemma weak_unit_iff_self_target: shows "weak_unit b \ b \ targets b" using weak_unit_self_composable by auto abbreviation (input) in_hhom\<^sub>W\<^sub>C ("\_ : _ \\<^sub>W\<^sub>C _\") where "in_hhom\<^sub>W\<^sub>C \ f g \ arr \ \ f \ sources \ \ g \ targets \" lemma sources_hcomp: assumes "\ \ \ \ null" shows "sources (\ \ \) = sources \" using assms match_1 match_3 null_agreement by blast lemma targets_hcomp: assumes "\ \ \ \ null" shows "targets (\ \ \) = targets \" using assms match_2 match_4 null_agreement by blast lemma H\<^sub>R_preserved_along_iso: assumes "weak_unit a" and "a \ a'" shows "endofunctor (Right a) (H\<^sub>R a')" proof - have a: "ide a \ weak_unit a" using assms isomorphic_def by auto have a': "ide a'" using assms isomorphic_def by auto (* TODO: The following interpretation re-introduces unwanted notation for "in_hom" *) interpret R: subcategory V \right a\ using a right_hom_is_subcategory by simp have *: "\\. R.arr \ \ H\<^sub>R a' \ = \ \ a'" using assms H\<^sub>R_def by simp have preserves_arr: "\\. R.arr \ \ R.arr (H\<^sub>R a' \)" using assms a' * R.arr_char right_def weak_unit_def weak_unit_self_composable isomorphic_implies_equicomposable R.ide_char match_3 hcomp_simps\<^sub>W\<^sub>C(1) null_agreement by metis show "endofunctor R.comp (H\<^sub>R a')" proof show "\\. \ R.arr \ \ H\<^sub>R a' \ = R.null" using assms R.arr_char R.null_char right_def H\<^sub>R_def null_agreement right_respects_isomorphic by metis fix \ assume "R.arr \" hence \: "R.arr \ \ arr \ \ right a \ \ right a' \ \ \ \ a \ null \ \ \ a' \ null" using assms R.arr_char right_respects_isomorphic composable_implies_arr null_agreement right_def by metis show "R.arr (H\<^sub>R a' \)" using \ preserves_arr by blast show "R.dom (H\<^sub>R a' \) = H\<^sub>R a' (R.dom \)" using a' \ * R.arr_char R.dom_char preserves_arr hom_connected(1) right_def by simp show "R.cod (H\<^sub>R a' \) = H\<^sub>R a' (R.cod \)" using a' \ * R.arr_char R.cod_char preserves_arr hom_connected(3) right_def by simp next fix \ \ assume \\: "R.seq \ \" have \: "R.arr \ \ arr \ \ right a \ \ right a' \ \ \ \ a \ null \ \ \ a' \ null" using assms \\ R.arr_char right_respects_isomorphic composable_implies_arr null_agreement right_def by (elim R.seqE) metis have \: "\\ : R.cod \ \ R.cod \\ \ arr \ \ right a \ \ H \ a \ null \ right a' \ \ H \ a' \ null" - using assms \\ right_def right_respects_isomorphic isomorphic_implies_equicomposable - R.hom_char R.dom_char R.cod_char R.inclusion in_homI R.arrE - by (elim R.seqE R.arrE) auto + by (metis "*" R.cod_simp R.comp_def R.inclusion R.not_arr_null R.null_char R.seqE + \\ in_homI preserves_arr right_def) show "H\<^sub>R a' (R.comp \ \) = R.comp (H\<^sub>R a' \) (H\<^sub>R a' \)" proof - have 1: "R.arr (H\<^sub>R a' \)" using \ preserves_arr by blast have 2: "seq (\ \ a') (\ \ a')" using a' \ \ R.arr_char R.inclusion R.dom_char R.cod_char isomorphic_implies_equicomposable by auto show ?thesis using a' \ \ \\ 1 2 preserves_arr H\<^sub>R_def R.dom_simp R.cod_simp R.comp_char R.seq_char R.inclusion whisker_right by metis qed qed qed lemma H\<^sub>L_preserved_along_iso: assumes "weak_unit a" and "a \ a'" shows "endofunctor (Left a) (H\<^sub>L a')" proof - have a: "ide a \ weak_unit a" using assms isomorphic_def by auto have a': "ide a'" using assms isomorphic_def by auto (* TODO: The following interpretation re-introduces unwanted notation for "in_hom" *) interpret L: subcategory V \left a\ using a left_hom_is_subcategory by simp have *: "\\. L.arr \ \ H\<^sub>L a' \ = a' \ \" using assms H\<^sub>L_def by simp have preserves_arr: "\\. L.arr \ \ L.arr (H\<^sub>L a' \)" using assms a' * L.arr_char left_def weak_unit_def weak_unit_self_composable isomorphic_implies_equicomposable L.ide_char match_4 hcomp_simps\<^sub>W\<^sub>C(1) null_agreement by metis show "endofunctor L.comp (H\<^sub>L a')" proof show "\\. \ L.arr \ \ H\<^sub>L a' \ = L.null" using assms L.arr_char L.null_char left_def H\<^sub>L_def null_agreement left_respects_isomorphic by metis fix \ assume "L.arr \" hence \: "L.arr \ \ arr \ \ left a \ \ left a' \ \ a \ \ \ null \ a' \ \ \ null" using assms L.arr_char left_respects_isomorphic composable_implies_arr null_agreement left_def by metis show "L.arr (H\<^sub>L a' \)" using \ preserves_arr by blast show "L.dom (H\<^sub>L a' \) = H\<^sub>L a' (L.dom \)" using a' \ * L.arr_char L.dom_char preserves_arr hom_connected(2) left_def by simp show "L.cod (H\<^sub>L a' \) = H\<^sub>L a' (L.cod \)" using a' \ * L.arr_char L.cod_char preserves_arr hom_connected(4) left_def by simp next fix \ \ assume \\: "L.seq \ \" have "L.arr \" using \\ by (elim L.seqE, auto) hence \: "L.arr \ \ arr \ \ left a \ \ left a' \ \ a \ \ \ null \ a' \ \ \ null" using assms L.arr_char left_respects_isomorphic composable_implies_arr null_agreement left_def by metis have \: "\\ : L.cod \ \ L.cod \\ \ arr \ \ left a \ \ a \ \ \ null \ left a' \ \ a' \ \ \ null" - using assms \\ left_def left_respects_isomorphic isomorphic_implies_equicomposable - L.hom_char L.dom_char L.cod_char L.inclusion in_homI L.arrE - by (elim L.seqE L.arrE) auto + by (metis (mono_tags, opaque_lifting) L.arrE L.cod_simp L.seq_char \\ assms(2) + in_homI seqE left_def left_respects_isomorphic) show "H\<^sub>L a' (L.comp \ \) = L.comp (H\<^sub>L a' \) (H\<^sub>L a' \)" proof - have 1: "L.arr (H\<^sub>L a' \)" using \ preserves_arr by blast have 2: "seq (a' \ \) (a' \ \)" using a' \ \ L.arr_char L.inclusion L.dom_char L.cod_char isomorphic_implies_equicomposable by auto show ?thesis using a' \ \ \\ 1 2 preserves_arr H\<^sub>L_def L.dom_simp L.cod_simp L.comp_char L.seq_char L.inclusion whisker_left by metis qed qed qed end subsection "Regularity" text \ We call a weak composition \emph{regular} if \f \ a \ f\ whenever \a\ is a source of 1-cell \f\, and \b \ f \ f\ whenever \b\ is a target of \f\. A consequence of regularity is that horizontal composability of 2-cells is fully determined by their sets of sources and targets. \ locale regular_weak_composition = weak_composition + assumes comp_ide_source: "\ a \ sources f; ide f \ \ f \ a \ f" and comp_target_ide: "\ b \ targets f; ide f \ \ b \ f \ f" begin lemma sources_determine_composability: assumes "a \ sources \" shows "\ \ \ \ null \ a \ \ \ null" proof - have *: "\\. ide \ \ a \ sources \ \ \ \ \ \ null \ a \ \ \ null" using assms comp_ide_source isomorphic_implies_equicomposable match_1 match_3 by (meson sourcesD(3)) show ?thesis proof - have "arr \" using assms composable_implies_arr by auto thus ?thesis using assms * [of "dom \"] hom_connected(1) by auto qed qed lemma targets_determine_composability: assumes "b \ targets \" shows "\ \ \ \ null \ \ \ b \ null" proof - have *: "\\. ide \ \ b \ targets \ \ \ \ \ \ null \ \ \ b \ null" using assms comp_target_ide isomorphic_implies_equicomposable match_2 match_4 by (meson targetsD(3)) show ?thesis proof - have "arr \" using assms composable_implies_arr by auto thus ?thesis using assms * [of "dom \"] hom_connected(2) by auto qed qed lemma composable_if_connected: assumes "sources \ \ targets \ \ {}" shows "\ \ \ \ null" using assms targets_determine_composability by blast lemma connected_if_composable: assumes "\ \ \ \ null" shows "sources \ = targets \" using assms sources_determine_composability targets_determine_composability by blast lemma iso_hcomp\<^sub>R\<^sub>W\<^sub>C: assumes "iso \" and "iso \" and "sources \ \ targets \ \ {}" shows "iso (\ \ \)" and "inverse_arrows (\ \ \) (inv \ \ inv \)" proof - have \: "arr \ \ \\ : dom \ \ cod \\ \ iso \ \ \inv \ : cod \ \ dom \\" using assms inv_in_hom arr_iff_in_hom iso_is_arr by auto have \: "arr \ \ \\ : dom \ \ cod \\ \ iso \ \ \inv \ : cod \ \ dom \\" using assms inv_in_hom by blast have 1: "sources (inv \) \ targets (inv \) \ {}" using assms \ \ sources_dom sources_cod targets_dom targets_cod arr_inv cod_inv by metis show "inverse_arrows (\ \ \) (inv \ \ inv \)" using assms 1 \ \ inv_in_hom inv_is_inverse comp_inv_arr interchange [of "inv \" \ "inv \" \] composable_if_connected ide_hcomp\<^sub>W\<^sub>C sources_dom targets_dom interchange [of \ "inv \" \ "inv \"] ide_hcomp\<^sub>W\<^sub>C sources_cod targets_cod ide_compE ide_dom comp_arr_inv' inverse_arrowsE seqI' inverse_arrowsI by metis thus "iso (\ \ \)" by auto qed lemma inv_hcomp\<^sub>R\<^sub>W\<^sub>C: assumes "iso \" and "iso \" and "sources \ \ targets \ \ {}" shows "inv (\ \ \) = inv \ \ inv \" using assms iso_hcomp\<^sub>R\<^sub>W\<^sub>C(2) [of \ \] inverse_arrow_unique [of "H \ \"] inv_is_inverse by auto end subsection "Associativity" text \ An \emph{associative weak composition} consists of a weak composition that has been equipped with an \emph{associator} isomorphism: \\\[f, g, h] : (f \ g) \ h \ f \ g \ h\\ for each composable triple \(f, g, h)\ of 1-cells, subject to naturality and coherence conditions. \ locale associative_weak_composition = weak_composition + fixes \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") assumes assoc_in_vhom\<^sub>A\<^sub>W\<^sub>C: "\ ide f; ide g; ide h; f \ g \ null; g \ h \ null \ \ \\[f, g, h] : (f \ g) \ h \ f \ g \ h\" and assoc_naturality\<^sub>A\<^sub>W\<^sub>C: "\ \ \ \ \ null; \ \ \ \ null \ \ \[cod \, cod \, cod \] \ ((\ \ \) \ \) = (\ \ \ \ \) \ \[dom \, dom \, dom \]" and iso_assoc\<^sub>A\<^sub>W\<^sub>C: "\ ide f; ide g; ide h; f \ g \ null; g \ h \ null \ \ iso \[f, g, h]" and pentagon\<^sub>A\<^sub>W\<^sub>C: "\ ide f; ide g; ide h; ide k; sources f \ targets g \ {}; sources g \ targets h \ {}; sources h \ targets k \ {} \ \ (f \ \[g, h, k]) \ \[f, g \ h, k] \ (\[f, g, h] \ k) = \[f, g, h \ k] \ \[f \ g, h, k]" begin lemma assoc_in_hom\<^sub>A\<^sub>W\<^sub>C: assumes "ide f" and "ide g" and "ide h" and "f \ g \ null" and "g \ h \ null" shows "sources \[f, g, h] = sources h" and "targets \[f, g, h] = targets f" and "\\[f, g, h] : (f \ g) \ h \ f \ g \ h\" proof - show 1: "\\[f, g, h] : (f \ g) \ h \ f \ g \ h\" using assms assoc_in_vhom\<^sub>A\<^sub>W\<^sub>C by simp show "sources \[f, g, h] = sources h" using assms 1 sources_dom [of "\[f, g, h]"] sources_hcomp match_3 by (elim in_homE, auto) show "targets \[f, g, h] = targets f" using assms 1 targets_cod [of "\[f, g, h]"] targets_hcomp match_4 by (elim in_homE, auto) qed lemma assoc_simps\<^sub>A\<^sub>W\<^sub>C [simp]: assumes "ide f" and "ide g" and "ide h" and "f \ g \ null" and "g \ h \ null" shows "arr \[f, g, h]" and "dom \[f, g, h] = (f \ g) \ h" and "cod \[f, g, h] = f \ g \ h" proof - have 1: "\\[f, g, h] : (f \ g) \ h \ f \ g \ h\" using assms assoc_in_hom\<^sub>A\<^sub>W\<^sub>C by auto show "arr \[f, g, h]" using 1 by auto show "dom \[f, g, h] = (f \ g) \ h" using 1 by auto show "cod \[f, g, h] = f \ g \ h" using 1 by auto qed lemma assoc'_in_hom\<^sub>A\<^sub>W\<^sub>C: assumes "ide f" and "ide g" and "ide h" and "f \ g \ null" and "g \ h \ null" shows "sources (inv \[f, g, h]) = sources h" and "targets (inv \[f, g, h]) = targets f" and "\inv \[f, g, h] : f \ g \ h \ (f \ g) \ h\" proof - show 1: "\inv \[f, g, h] : f \ g \ h \ (f \ g) \ h\" using assms assoc_in_hom\<^sub>A\<^sub>W\<^sub>C iso_assoc\<^sub>A\<^sub>W\<^sub>C inv_in_hom by auto show "sources (inv \[f, g, h]) = sources h" using assms 1 sources_hcomp [of "f \ g" h] sources_cod match_3 null_agreement by (elim in_homE, metis) show "targets (inv \[f, g, h]) = targets f" using assms 1 targets_hcomp [of f "g \ h"] targets_dom match_4 null_agreement by (elim in_homE, metis) qed lemma assoc'_simps\<^sub>A\<^sub>W\<^sub>C [simp]: assumes "ide f" and "ide g" and "ide h" and "f \ g \ null" and "g \ h \ null" shows "arr (inv \[f, g, h])" and "dom (inv \[f, g, h]) = f \ g \ h" and "cod (inv \[f, g, h]) = (f \ g) \ h" proof - have 1: "\inv \[f, g, h] : f \ g \ h \ (f \ g) \ h\" using assms assoc'_in_hom\<^sub>A\<^sub>W\<^sub>C by auto show "arr (inv \[f, g, h])" using 1 by auto show "dom (inv \[f, g, h]) = f \ g \ h" using 1 by auto show "cod (inv \[f, g, h]) = (f \ g) \ h" using 1 by auto qed lemma assoc'_naturality\<^sub>A\<^sub>W\<^sub>C: assumes "\ \ \ \ null" and "\ \ \ \ null" shows "inv \[cod \, cod \, cod \] \ (\ \ \ \ \) = ((\ \ \) \ \) \ inv \[dom \, dom \, dom \]" proof - have \\\: "arr \ \ arr \ \ arr \" using assms composable_implies_arr by simp have 0: "dom \ \ dom \ \ null \ dom \ \ dom \ \ null \ cod \ \ cod \ \ null \ cod \ \ cod \ \ null" using assms \\\ hom_connected by simp have 1: "\\ \ \ \ \ : dom \ \ dom \ \ dom \ \ cod \ \ cod \ \ cod \\" using assms match_4 by auto have 2: "\(\ \ \) \ \ : (dom \ \ dom \) \ dom \ \ (cod \ \ cod \) \ cod \\" using assms match_3 by auto have "(inv \[cod \, cod \, cod \] \ (\ \ \ \ \)) \ \[dom \, dom \, dom \] = (\ \ \) \ \" proof - have "(\ \ \) \ \ = (inv \[cod \, cod \, cod \] \ \[cod \, cod \, cod \]) \ ((\ \ \) \ \)" using 0 2 \\\ assoc_in_hom\<^sub>A\<^sub>W\<^sub>C iso_assoc\<^sub>A\<^sub>W\<^sub>C comp_inv_arr inv_is_inverse comp_cod_arr by auto thus ?thesis using assms \\\ 0 2 assoc_naturality\<^sub>A\<^sub>W\<^sub>C comp_assoc by metis qed thus ?thesis using 0 1 2 \\\ iso_assoc\<^sub>A\<^sub>W\<^sub>C assoc'_in_hom\<^sub>A\<^sub>W\<^sub>C inv_in_hom invert_side_of_triangle(2) by auto qed end subsection "Unitors" text \ For an associative weak composition with a chosen unit isomorphism \\ : a \ a \ a\, where \a\ is a weak unit, horizontal composition on the right by \a\ is a fully faithful endofunctor \R\ of the subcategory of arrows composable on the right with \a\, and is consequently an endo-equivalence of that subcategory. This equivalence, together with the associator isomorphisms and unit isomorphism \\\, canonically associate, with each identity arrow \f\ composable on the right with \a\, a \emph{right unit} isomorphism \\\[f] : f \ a \ f\\. These isomorphisms are the components of a natural isomorphism from \R\ to the identity functor. \ locale right_hom_with_unit = associative_weak_composition V H \ + right_hom V H a for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: 'a and a :: 'a + assumes weak_unit_a: "weak_unit a" and \_in_hom: "\\ : a \ a \ a\" and iso_\: "iso \" begin abbreviation R where "R \ H\<^sub>R a" interpretation R: endofunctor S.comp R using weak_unit_a weak_unit_self_composable endofunctor_H\<^sub>R by simp interpretation R: fully_faithful_functor S.comp S.comp R using weak_unit_a weak_unit_def by simp lemma fully_faithful_functor_R: shows "fully_faithful_functor S.comp S.comp R" .. definition runit ("\[_]") where "runit f \ THE \. \\ : R f \\<^sub>S f\ \ R \ = (f \ \) \\<^sub>S \[f, a, a]" lemma iso_unit: shows "S.iso \" and "\\ : a \ a \\<^sub>S a\" proof - show "\\ : a \ a \\<^sub>S a\" using weak_unit_a S.ide_char S.arr_char right_def weak_unit_self_composable S.ideD(1) R.preserves_arr H\<^sub>R_def S.in_hom_char S.arr_char right_def \_in_hom S.ideD(1) hom_connected(3) in_homE by metis thus "S.iso \" using iso_\ iso_char by blast qed lemma characteristic_iso: assumes "S.ide f" shows "\\[f, a, a] : (f \ a) \ a \\<^sub>S f \ a \ a\" and "\f \ \ : f \ a \ a \\<^sub>S f \ a\" and "\(f \ \) \\<^sub>S \[f, a, a] : R (R f) \\<^sub>S R f\" and "S.iso ((f \ \) \\<^sub>S \[f, a, a])" proof - have f: "S.ide f \ ide f" using assms S.ide_char by simp have a: "weak_unit a \ ide a \ S.ide a" using weak_unit_a S.ide_char weak_unit_def S.arr_char right_def weak_unit_self_composable by metis have fa: "f \ a \ null \ (f \ a) \ a \ null \ ((f \ a) \ a) \ a \ null" using assms S.ideD(1) R.preserves_arr H\<^sub>R_def S.not_arr_null S.null_char by metis have aa: "a \ a \ null" using a S.ideD(1) R.preserves_arr H\<^sub>R_def S.not_arr_null weak_unit_self_composable by auto have f_ia: "f \ \ \ null" using assms S.ide_char right_def S.arr_char hom_connected(4) \_in_hom by auto show assoc_in_hom: "\\[f, a, a] : (f \ a) \ a \\<^sub>S f \ a \ a\" using a f fa hom_connected(1) [of "\[f, a, a]" a] S.arr_char right_def match_3 match_4 S.in_hom_char weak_unit_self_composable by auto show 1: "\f \ \ : f \ a \ a \\<^sub>S f \ a\" using a f fa iso_unit left_hcomp_closed by (simp add: f_ia ide_in_hom) have unit_part: "\f \ \ : f \ a \ a \\<^sub>S f \ a\ \ S.iso (f \ \)" proof - have "S.iso (f \ \)" using a f fa f_ia 1 VoV.arr_char VxV.inv_simp inv_in_hom hom_connected(2) [of f "inv \"] VoV.arr_char VoV.iso_char preserves_iso iso_char iso_\ S.dom_simp S.cod_simp by auto thus ?thesis using 1 by blast qed show "S.iso ((f \ \) \\<^sub>S \[f, a, a])" using assms a f fa aa hom_connected(1) [of "\[f, a, a]" a] right_def iso_assoc\<^sub>A\<^sub>W\<^sub>C iso_char S.arr_char unit_part assoc_in_hom isos_compose S.isos_compose S.seqI' by auto show "\(f \ \) \\<^sub>S \[f, a, a] : R (R f) \\<^sub>S R f\" unfolding H\<^sub>R_def using unit_part assoc_in_hom by blast qed lemma runit_char: assumes "S.ide f" shows "\\[f] : R f \\<^sub>S f\" and "R \[f] = (f \ \) \\<^sub>S \[f, a, a]" and "\!\. \\ : R f \\<^sub>S f\ \ R \ = (f \ \) \\<^sub>S \[f, a, a]" proof - let ?P = "\\. \\ : R f \\<^sub>S f\ \ R \ = (f \ \) \\<^sub>S \[f, a, a]" show "\!\. ?P \" proof - have "\\. ?P \" using assms S.ide_char S.arr_char R.preserves_ide characteristic_iso(3) R.is_full by auto moreover have "\\ \'. ?P \ \ ?P \' \ \ = \'" using R.is_faithful S.in_homE by metis ultimately show ?thesis by blast qed hence "?P (THE \. ?P \)" using theI' [of ?P] by fastforce hence 1: "?P \[f]" unfolding runit_def by blast show "\\[f] : R f \\<^sub>S f\" using 1 by fast show "R \[f] = (f \ \) \\<^sub>S \[f, a, a]" using 1 by fast qed lemma iso_runit: assumes "S.ide f" shows "S.iso \[f]" using assms characteristic_iso(4) runit_char R.reflects_iso by metis lemma runit_eqI: assumes "\f : a \\<^sub>S b\" and "\\ : R f \\<^sub>S f\" and "R \ = ((f \ \) \\<^sub>S \[f, a, a])" shows "\ = \[f]" using assms S.ide_cod runit_char [of f] by blast lemma runit_naturality: assumes "S.arr \" shows "\[S.cod \] \\<^sub>S R \ = \ \\<^sub>S \[S.dom \]" proof - have 1: "\\[S.cod \] \\<^sub>S R \ : R (S.dom \) \\<^sub>S S.cod \\" using assms runit_char(1) S.ide_cod by blast have 2: "S.par (\[S.cod \] \\<^sub>S R \) (\ \\<^sub>S \[S.dom \])" using assms 1 S.ide_dom runit_char(1) by (metis S.comp_in_homI' S.in_homE) moreover have "R (\[S.cod \] \\<^sub>S R \) = R (\ \\<^sub>S \[S.dom \])" proof - have 3: "\\ \ a \ a : S.dom \ \ a \ a \\<^sub>S S.cod \ \ a \ a\" using assms weak_unit_a R.preserves_hom H\<^sub>R_def S.arr_iff_in_hom S.arr_char by (metis match_4 weak_unit_in_vhom weak_unit_self_right S.in_hom_char left_hcomp_closed S.not_arr_null S.null_char) have 4: "R (\[S.cod \] \\<^sub>S R \) = R \[S.cod \] \\<^sub>S R (R \)" using assms 1 R.as_nat_trans.preserves_comp_2 by blast also have 5: "... = ((S.cod \ \ \) \\<^sub>S \[S.cod \, a, a]) \\<^sub>S ((\ \ a) \ a)" using assms R.preserves_arr runit_char S.ide_cod H\<^sub>R_def by auto also have 6: "... = (S.cod \ \ \) \\<^sub>S \[S.cod \, a, a] \\<^sub>S ((\ \ a) \ a)" using assms S.comp_assoc by simp also have "... = (S.cod \ \ \) \\<^sub>S (\ \ a \ a) \\<^sub>S \[S.dom \, a, a]" proof - have "(\ \ a \ a) \\<^sub>S \[S.dom \, a, a] = \[S.cod \, a, a] \\<^sub>S ((\ \ a) \ a)" proof - have "(\ \ a \ a) \\<^sub>S \[S.dom \, a, a] = (\ \ a \ a) \ \[S.dom \, a, a]" using assms 3 S.ide_dom characteristic_iso(1) S.in_hom_char S.comp_char S.dom_simp by fastforce also have "... = \[S.cod \, a, a] \ ((\ \ a) \ a)" using assms weak_unit_a assoc_naturality\<^sub>A\<^sub>W\<^sub>C [of \ a a] S.dom_simp S.cod_simp weak_unit_self_composable S.arr_char right_def by simp also have "... = \[S.cod \, a, a] \\<^sub>S ((\ \ a) \ a)" using S.in_hom_char S.comp_char by (metis 2 4 5 6 R.preserves_arr S.seq_char) finally show ?thesis by blast qed thus ?thesis by argo qed also have "... = ((S.cod \ \ \) \\<^sub>S (\ \ a \ a)) \\<^sub>S \[S.dom \, a, a]" using S.comp_assoc by auto also have "... = ((\ \ a) \\<^sub>S (S.dom \ \ \)) \\<^sub>S \[S.dom \, a, a]" proof - have "\ \ a \ a \ null" using 3 S.not_arr_null by auto moreover have "S.dom \ \ \ \ null" using assms S.not_arr_null by (metis S.dom_char \_in_hom calculation hom_connected(1-2) in_homE) ultimately have "(S.cod \ \ \) \\<^sub>S (\ \ a \ a) = (\ \ a) \\<^sub>S (S.dom \ \ \)" using assms weak_unit_a iso_unit S.comp_arr_dom S.comp_cod_arr interchange [of "S.cod \" \ \ "a \ a"] interchange [of \ "S.dom \" a \] by auto thus ?thesis by argo qed also have "... = (\ \ a) \\<^sub>S (S.dom \ \ \) \\<^sub>S \[S.dom \, a, a]" using S.comp_assoc by auto also have "... = R \ \\<^sub>S R \[S.dom \]" using assms runit_char(2) S.ide_dom H\<^sub>R_def by auto also have "... = R (\ \\<^sub>S \[S.dom \])" using assms S.arr_iff_in_hom [of \] runit_char(1) S.ide_dom by fastforce finally show ?thesis by blast qed ultimately show "\[S.cod \] \\<^sub>S (R \) = \ \\<^sub>S \[S.dom \]" using R.is_faithful by blast qed abbreviation \ where "\ \ \ if S.arr \ then \ \\<^sub>S \[S.dom \] else null" interpretation \: natural_transformation S.comp S.comp R S.map \ proof - interpret \: transformation_by_components S.comp S.comp R S.map runit using runit_char(1) runit_naturality by unfold_locales simp_all have "\.map = \" using \.is_extensional \.map_def \.naturality \.map_simp_ide S.ide_dom S.ide_cod S.map_def by auto thus "natural_transformation S.comp S.comp R S.map \" using \.natural_transformation_axioms by auto qed lemma natural_transformation_\: shows "natural_transformation S.comp S.comp R S.map \" .. interpretation \: natural_isomorphism S.comp S.comp R S.map \ using S.ide_is_iso iso_runit runit_char(1) S.isos_compose by unfold_locales force lemma natural_isomorphism_\: shows "natural_isomorphism S.comp S.comp R S.map \" .. interpretation R: equivalence_functor S.comp S.comp R using natural_isomorphism_\ R.isomorphic_to_identity_is_equivalence by blast lemma equivalence_functor_R: shows "equivalence_functor S.comp S.comp R" .. lemma runit_commutes_with_R: assumes "S.ide f" shows "\[R f] = R \[f]" using assms runit_char(1) R.preserves_hom [of "\[f]" "R f" f] runit_naturality iso_runit S.iso_is_section S.section_is_mono S.monoE by (metis S.in_homE S.seqI') end text \ Symmetric results hold for the subcategory of all arrows composable on the left with a specified weak unit \b\. This yields the \emph{left unitors}. \ locale left_hom_with_unit = associative_weak_composition V H \ + left_hom V H b for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: 'a and b :: 'a + assumes weak_unit_b: "weak_unit b" and \_in_hom: "\\ : b \ b \ b\" and iso_\: "iso \" begin abbreviation L where "L \ H\<^sub>L b" interpretation L: endofunctor S.comp L using weak_unit_b weak_unit_self_composable endofunctor_H\<^sub>L by simp interpretation L: fully_faithful_functor S.comp S.comp L using weak_unit_b weak_unit_def by simp lemma fully_faithful_functor_L: shows "fully_faithful_functor S.comp S.comp L" .. definition lunit ("\[_]") where "lunit f \ THE \. \\ : L f \\<^sub>S f\ \ L \ = (\ \ f) \\<^sub>S (inv \[b, b, f])" lemma iso_unit: shows "S.iso \" and "\\ : b \ b \\<^sub>S b\" proof - show "\\ : b \ b \\<^sub>S b\" using weak_unit_b S.ide_char S.arr_char left_def weak_unit_self_composable S.ideD(1) L.preserves_arr H\<^sub>L_def S.in_hom_char S.arr_char left_def \_in_hom S.ideD(1) hom_connected(4) in_homE by metis thus "S.iso \" using iso_\ iso_char by blast qed lemma characteristic_iso: assumes "S.ide f" shows "\inv \[b, b, f] : b \ b \ f \\<^sub>S (b \ b) \ f\" and "\\ \ f : (b \ b) \ f \\<^sub>S b \ f\" and "\(\ \ f) \\<^sub>S inv \[b, b, f] : L (L f) \\<^sub>S L f\" and "S.iso ((\ \ f) \\<^sub>S inv \[b, b, f])" proof - have f: "S.ide f \ ide f" using assms S.ide_char by simp have b: "weak_unit b \ ide b \ S.ide b" using weak_unit_b S.ide_char weak_unit_def S.arr_char left_def weak_unit_self_composable by metis have bf: "b \ f \ null \ b \ b \ b \ f \ null" using assms S.ideD(1) L.preserves_arr H\<^sub>L_def S.not_arr_null S.null_char by metis have bb: "b \ b \ null" using b S.ideD(1) L.preserves_arr H\<^sub>L_def S.not_arr_null weak_unit_self_composable by auto have ib_f: "\ \ f \ null" using assms S.ide_char left_def S.arr_char hom_connected(3) \_in_hom by auto show assoc_in_hom: "\inv \[b, b, f] : b \ b \ f \\<^sub>S (b \ b) \ f\" using b f bf bb hom_connected(2) [of b "inv \[b, b, f]"] left_def by (metis S.arrI S.cod_closed S.in_hom_char assoc'_in_hom\<^sub>A\<^sub>W\<^sub>C(3) assoc'_simps\<^sub>A\<^sub>W\<^sub>C(2-3)) show 1: "\\ \ f : (b \ b) \ f \\<^sub>S b \ f\" using b f bf right_hcomp_closed by (simp add: ib_f ide_in_hom iso_unit(2)) have unit_part: "\\ \ f : (b \ b) \ f \\<^sub>S b \ f\ \ S.iso (\ \ f)" proof - have "S.iso (\ \ f)" using b f bf ib_f 1 VoV.arr_char VxV.inv_simp inv_in_hom hom_connected(1) [of "inv \" f] VoV.arr_char VoV.iso_char preserves_iso iso_char iso_\ S.dom_simp S.cod_simp by auto thus ?thesis using 1 by blast qed show "S.iso ((\ \ f) \\<^sub>S inv \[b, b, f])" using assms b f bf bb hom_connected(2) [of b "inv \[b, b, f]"] left_def iso_assoc\<^sub>A\<^sub>W\<^sub>C iso_char S.arr_char unit_part assoc_in_hom isos_compose S.isos_compose S.seqI' by auto show "\(\ \ f) \\<^sub>S inv \[b, b, f] : L (L f) \\<^sub>S L f\" unfolding H\<^sub>L_def using unit_part assoc_in_hom by blast qed lemma lunit_char: assumes "S.ide f" shows "\\[f] : L f \\<^sub>S f\" and "L \[f] = (\ \ f) \\<^sub>S inv \[b, b, f]" and "\!\. \\ : L f \\<^sub>S f\ \ L \ = (\ \ f) \\<^sub>S inv \[b, b, f]" proof - let ?P = "\\. \\ : L f \\<^sub>S f\ \ L \ = (\ \ f) \\<^sub>S inv \[b, b, f]" show "\!\. ?P \" proof - have "\\. ?P \" proof - have 1: "S.ide f" using assms S.ide_char S.arr_char by simp moreover have "S.ide (L f)" using 1 L.preserves_ide by simp ultimately show ?thesis using assms characteristic_iso(3) L.is_full by blast qed moreover have "\\ \'. ?P \ \ ?P \' \ \ = \'" using L.is_faithful S.in_homE by metis ultimately show ?thesis by blast qed hence "?P (THE \. ?P \)" using theI' [of ?P] by fastforce hence 1: "?P \[f]" unfolding lunit_def by blast show "\\[f] : L f \\<^sub>S f\" using 1 by fast show "L \[f] = (\ \ f) \\<^sub>S inv \[b, b, f]" using 1 by fast qed lemma iso_lunit: assumes "S.ide f" shows "S.iso \[f]" using assms characteristic_iso(4) lunit_char L.reflects_iso by metis lemma lunit_eqI: assumes "\f : a \\<^sub>S b\" and "\\ : L f \\<^sub>S f\" and "L \ = ((\ \ f) \\<^sub>S inv \[b, b, f])" shows "\ = \[f]" using assms S.ide_cod lunit_char [of f] by blast lemma lunit_naturality: assumes "S.arr \" shows "\[S.cod \] \\<^sub>S L \ = \ \\<^sub>S \[S.dom \]" proof - have 1: "\\[S.cod \] \\<^sub>S L \ : L (S.dom \) \\<^sub>S S.cod \\" using assms lunit_char(1) [of "S.cod \"] S.ide_cod by blast have "S.par (\[S.cod \] \\<^sub>S L \) (\ \\<^sub>S \[S.dom \])" using assms 1 S.ide_dom lunit_char(1) by (metis S.comp_in_homI' S.in_homE) moreover have "L (\[S.cod \] \\<^sub>S L \) = L (\ \\<^sub>S \[S.dom \])" proof - have 2: "\b \ b \ \ : b \ b \ S.dom \ \\<^sub>S b \ b \ S.cod \\" using assms weak_unit_b L.preserves_hom H\<^sub>L_def S.arr_iff_in_hom [of \] S.arr_char by simp have 3: "\(b \ b) \ \ : (b \ b) \ S.dom \ \\<^sub>S (b \ b) \ S.cod \\" using assms weak_unit_b L.preserves_hom H\<^sub>L_def S.arr_iff_in_hom S.arr_char by (metis match_3 weak_unit_in_vhom weak_unit_self_left S.in_hom_char S.not_arr_null S.null_char right_hcomp_closed) have "L (\[S.cod \] \\<^sub>S L \) = L \[S.cod \] \\<^sub>S L (L \)" using assms 1 L.as_nat_trans.preserves_comp_2 by blast also have "... = ((\ \ S.cod \) \\<^sub>S inv \[b, b, S.cod \]) \\<^sub>S (b \ b \ \)" using assms L.preserves_arr lunit_char S.ide_cod H\<^sub>L_def by auto also have "... = (\ \ S.cod \) \\<^sub>S inv \[b, b, S.cod \] \\<^sub>S (b \ b \ \)" using S.comp_assoc by auto also have "... = (\ \ S.cod \) \\<^sub>S ((b \ b) \ \) \\<^sub>S inv \[b, b, S.dom \]" proof - have "inv \[b, b, S.cod \] \\<^sub>S (b \ b \ \) = ((b \ b) \ \) \\<^sub>S inv \[b, b, S.dom \]" proof - have "((b \ b) \ \) \\<^sub>S inv \[b, b, S.dom \] = ((b \ b) \ \) \ inv \[b, b, S.dom \]" using assms 3 S.in_hom_char S.comp_char [of "(b \ b) \ \" "inv \[b, b, S.dom \]"] by (metis S.ide_dom characteristic_iso(1) ext) also have "... = inv \[b, b, S.cod \] \ (b \ b \ \)" using assms weak_unit_b assoc'_naturality\<^sub>A\<^sub>W\<^sub>C [of b b \] S.dom_simp S.cod_simp weak_unit_self_composable S.arr_char left_def by simp also have "... = inv \[b, b, S.cod \] \\<^sub>S (b \ b \ \)" using assms 2 S.in_hom_char S.comp_char by (metis S.comp_simp S.ide_cod S.seqI' characteristic_iso(1)) finally show ?thesis by argo qed thus ?thesis by argo qed also have "... = ((\ \ S.cod \) \\<^sub>S ((b \ b) \ \)) \\<^sub>S inv \[b, b, S.dom \]" using S.comp_assoc by auto also have "... = ((b \ \) \\<^sub>S (\ \ S.dom \)) \\<^sub>S inv \[b, b, S.dom \]" proof - have "(b \ b) \ \ \ null" using 3 S.not_arr_null by (elim S.in_homE, auto) moreover have "\ \ S.dom \ \ null" using assms S.not_arr_null by (metis S.dom_char \_in_hom calculation hom_connected(1-2) in_homE) ultimately have "(\ \ S.cod \) \\<^sub>S ((b \ b) \ \) = (b \ \) \\<^sub>S (\ \ S.dom \)" using assms weak_unit_b iso_unit S.comp_arr_dom S.comp_cod_arr interchange [of \ "b \ b" "S.cod \" \ ] interchange [of b \ \ "S.dom \"] by auto thus ?thesis by argo qed also have "... = (b \ \) \\<^sub>S (\ \ S.dom \) \\<^sub>S inv \[b, b, S.dom \]" using S.comp_assoc by auto also have "... = L \ \\<^sub>S L \[S.dom \]" using assms lunit_char(2) S.ide_dom H\<^sub>L_def by auto also have "... = L (\ \\<^sub>S \[S.dom \])" using assms S.arr_iff_in_hom [of \] lunit_char(1) S.ide_dom S.seqI by fastforce finally show ?thesis by blast qed ultimately show "\[S.cod \] \\<^sub>S L \ = \ \\<^sub>S \[S.dom \]" using L.is_faithful by blast qed abbreviation \ where "\ \ \ if S.arr \ then \ \\<^sub>S \[S.dom \] else null" interpretation \: natural_transformation S.comp S.comp L S.map \ proof - interpret \: transformation_by_components S.comp S.comp L S.map lunit using lunit_char(1) lunit_naturality by (unfold_locales, simp_all) have "\.map = \" using \.is_extensional \.map_def \.naturality \.map_simp_ide S.ide_dom S.ide_cod S.map_def by auto thus "natural_transformation S.comp S.comp L S.map \" using \.natural_transformation_axioms by auto qed lemma natural_transformation_\: shows "natural_transformation S.comp S.comp L S.map \" .. interpretation \: natural_isomorphism S.comp S.comp L S.map \ using S.ide_is_iso iso_lunit lunit_char(1) S.isos_compose by (unfold_locales, force) lemma natural_isomorphism_\: shows "natural_isomorphism S.comp S.comp L S.map \" .. interpretation L: equivalence_functor S.comp S.comp L using natural_isomorphism_\ L.isomorphic_to_identity_is_equivalence by blast lemma equivalence_functor_L: shows "equivalence_functor S.comp S.comp L" .. lemma lunit_commutes_with_L: assumes "S.ide f" shows "\[L f] = L \[f]" using assms lunit_char(1) L.preserves_hom [of "\[f]" "L f" f] lunit_naturality iso_lunit S.iso_is_section S.section_is_mono S.monoE by (metis S.in_homE S.seqI') end subsection "Prebicategories" text \ A \emph{prebicategory} is an associative weak composition satisfying the additional assumption that every arrow has a source and a target. \ locale prebicategory = associative_weak_composition + assumes arr_has_source: "arr \ \ sources \ \ {}" and arr_has_target: "arr \ \ targets \ \ {}" begin lemma arr_iff_has_src: shows "arr \ \ sources \ \ {}" using arr_has_source composable_implies_arr by auto lemma arr_iff_has_trg: shows "arr \ \ targets \ \ {}" using arr_has_target composable_implies_arr by auto end text \ The horizontal composition of a prebicategory is regular. \ sublocale prebicategory \ regular_weak_composition V H proof show "\a f. a \ sources f \ ide f \ f \ a \ f" proof - fix a f assume a: "a \ sources f" and f: "ide f" interpret Right_a: subcategory V \right a\ using a right_hom_is_subcategory weak_unit_self_composable by force interpret Right_a: right_hom_with_unit V H \ \some_unit a\ a using a iso_some_unit by (unfold_locales, auto) show "f \ a \ f" proof - have "Right_a.ide f" using a f Right_a.ide_char Right_a.arr_char right_def by auto hence "Right_a.iso (Right_a.runit f) \ (Right_a.runit f) \ Right_a.hom (f \ a) f" using Right_a.iso_runit Right_a.runit_char(1) H\<^sub>R_def by simp hence "iso (Right_a.runit f) \ (Right_a.runit f) \ hom (f \ a) f" using Right_a.iso_char Right_a.hom_char by auto thus ?thesis using f isomorphic_def by auto qed qed show "\b f. b \ targets f \ ide f \ b \ f \ f" proof - fix b f assume b: "b \ targets f" and f: "ide f" interpret Left_b: subcategory V \left b\ using b left_hom_is_subcategory weak_unit_self_composable by force interpret Left_b: left_hom_with_unit V H \ \some_unit b\ b using b iso_some_unit by (unfold_locales, auto) show "b \ f \ f" proof - have "Left_b.ide f" using b f Left_b.ide_char Left_b.arr_char left_def by auto hence "Left_b.iso (Left_b.lunit f) \ (Left_b.lunit f) \ Left_b.hom (b \ f) f" using b f Left_b.iso_lunit Left_b.lunit_char(1) H\<^sub>L_def by simp hence "iso (Left_b.lunit f) \ (Left_b.lunit f) \ hom (b \ f) f" using Left_b.iso_char Left_b.hom_char by auto thus ?thesis using isomorphic_def by auto qed qed qed text \ The regularity allows us to show that, in a prebicategory, all sources of a given arrow are isomorphic, and similarly for targets. \ context prebicategory begin lemma sources_are_isomorphic: assumes "a \ sources \" and "a' \ sources \" shows "a \ a'" proof - have \: "arr \" using assms composable_implies_arr by auto have "\f. \ ide f; a \ sources f; a' \ sources f \ \ a \ a'" using \ assms(1) comp_ide_source comp_target_ide [of a a'] weak_unit_self_composable(1) [of a] weak_unit_self_composable(1) [of a'] isomorphic_transitive isomorphic_symmetric sources_determine_composability sourcesD(2-3) by (metis (full_types) connected_if_composable) moreover have "ide (dom \) \ a \ sources (dom \) \ a' \ sources (dom \)" using assms \ sources_dom by auto ultimately show ?thesis by auto qed lemma targets_are_isomorphic: assumes "b \ targets \" and "b' \ targets \" shows "b \ b'" proof - have \: "arr \" using assms composable_implies_arr by auto have "\f. \ ide f; b \ targets f; b' \ targets f \ \ b \ b'" - using \ assms(1) comp_ide_source comp_target_ide [of b b'] - weak_unit_self_composable(1) [of b] weak_unit_self_composable(1) [of b'] - isomorphic_transitive isomorphic_symmetric - sources_determine_composability targetsD(2-3) - by (metis (full_types) connected_if_composable) + by (metis connected_if_composable sources_are_isomorphic targetsD(3)) moreover have "ide (dom \) \ b \ targets (dom \) \ b' \ targets (dom \)" using assms \ targets_dom [of \] by auto ultimately show ?thesis by auto qed text \ In fact, we now show that the sets of sources and targets of a 2-cell are isomorphism-closed, and hence are isomorphism classes. We first show that the notion ``weak unit'' is preserved under isomorphism. \ interpretation H: partial_magma H using is_partial_magma by auto lemma isomorphism_respects_weak_units: assumes "weak_unit a" and "a \ a'" shows "weak_unit a'" proof - obtain \ where \: "iso \ \ \\ : a \ a'\" using assms by auto interpret Left_a: subcategory V \left a\ using assms left_hom_is_subcategory by fastforce interpret Left_a: left_hom_with_unit V H \ \some_unit a\ a using assms iso_some_unit weak_unit_self_composable apply unfold_locales by auto interpret Right_a: subcategory V "right a" using assms right_hom_is_subcategory by fastforce interpret Right_a: right_hom_with_unit V H \ \some_unit a\ a using assms iso_some_unit weak_unit_self_composable apply unfold_locales by auto have a': "ide a' \ a \ a' \ null \ a' \ a \ null \ a' \ a' \ null \ \ \ a' \ null \ Left_a.ide a'" - using assms \ weak_unit_self_composable hom_connected weak_unit_self_composable(3) - Left_a.ide_char Left_a.arr_char left_def in_homE - by auto metis+ + by (metis (no_types, lifting) Left_a.left_hom_axioms Right_a.weak_unit_a \ assms(2) + ide_cod hom_connected(1) in_homE isomorphic_implies_equicomposable(1) + left_def left_hom_def subcategory.ideI isomorphic_implies_equicomposable(2) + weak_unit_self_composable(3)) have iso: "a' \ a' \ a'" proof - have 1: "Right a' = Right a" using assms right_respects_isomorphic by simp interpret Right_a': subcategory V \right a'\ using assms right_hom_is_subcategory by fastforce (* TODO: The previous interpretation brings in unwanted notation for in_hom. *) interpret Ra': endofunctor \Right a'\ \H\<^sub>R a'\ using assms a' endofunctor_H\<^sub>R by auto let ?\ = "Left_a.lunit a' \ inv (\ \ a')" have "iso ?\ \ \?\ : a' \ a' \ a'\" proof - have "iso (Left_a.lunit a') \ \Left_a.lunit a' : a \ a' \ a'\" using a' Left_a.lunit_char(1) Left_a.iso_lunit Left_a.iso_char Left_a.in_hom_char H\<^sub>L_def by auto moreover have "iso (\ \ a') \ \\ \ a' : a \ a' \ a' \ a'\" using a' \ 1 Right_a'.arr_char Right_a'.in_hom_char Right_a.iso_char right_def Ra'.preserves_iso Ra'.preserves_hom Right_a'.iso_char Ra'.preserves_dom Ra'.preserves_cod Right_a'.arr_iff_in_hom H\<^sub>R_def by metis ultimately show ?thesis using isos_compose by blast qed thus ?thesis using isomorphic_def by auto qed text \ We show that horizontal composition on the left and right by @{term a'} is naturally isomorphic to the identity functor. This follows from the fact that if @{term a} is isomorphic to @{term a'}, then horizontal composition with @{term a} is naturally isomorphic to horizontal composition with @{term a'}, hence the latter is naturally isomorphic to the identity if the former is. This is conceptually simple, but there are tedious composability details to handle. \ have 1: "Left a' = Left a \ Right a' = Right a" using assms left_respects_isomorphic right_respects_isomorphic by simp interpret L: fully_faithful_functor \Left a\ \Left a\ \H\<^sub>L a\ using assms weak_unit_def by simp interpret L': endofunctor \Left a\ \H\<^sub>L a'\ using a' 1 endofunctor_H\<^sub>L [of a'] by auto interpret \: natural_isomorphism \Left a\ \Left a\ \H\<^sub>L a\ \H\<^sub>L a'\ \H\<^sub>L \\ proof fix \ show "\ Left_a.arr \ \ H\<^sub>L \ \ = Left_a.null" using left_def \ H\<^sub>L_def hom_connected(1) Left_a.null_char null_agreement Left_a.arr_char by auto assume "Left_a.arr \" hence \: "Left_a.arr \ \ arr \ \ a \ \ \ null" using Left_a.arr_char left_def composable_implies_arr by simp have 2: "\ \ \ \ null" using assms \ \ Left_a.arr_char left_def hom_connected by auto show "Left_a.dom (H\<^sub>L \ \) = H\<^sub>L a (Left_a.dom \)" using assms 2 \ \ Left_a.arr_char left_def hom_connected(2) [of a \] weak_unit_self_composable match_4 Left_a.dom_char H\<^sub>L_def by auto show "Left_a.cod (H\<^sub>L \ \) = H\<^sub>L a' (Left_a.cod \)" using assms 2 \ \ Left_a.arr_char left_def hom_connected(2) [of a \] weak_unit_self_composable match_4 Left_a.cod_char H\<^sub>L_def by auto show "Left_a.comp (H\<^sub>L a' \) (H\<^sub>L \ (Left_a.dom \)) = H\<^sub>L \ \" proof - have "Left_a.comp (H\<^sub>L a' \) (H\<^sub>L \ (Left_a.dom \)) = Left_a.comp (a' \ \) (\ \ dom \)" using assms 1 2 \ \ Left_a.dom_char left_def H\<^sub>L_def by simp also have "... = (a' \ \) \ (\ \ dom \)" proof - have "Left_a.seq (a' \ \) (\ \ dom \)" proof (intro Left_a.seqI) show 3: "Left_a.arr (\ \ dom \)" using assms 2 \ \ Left_a.arr_char left_def by (metis H\<^sub>L_def L'.preserves_arr hcomp_simps\<^sub>W\<^sub>C(1) in_homE right_connected paste_1) show 4: "Left_a.arr (a' \ \)" using \ H\<^sub>L_def L'.preserves_arr by auto show "Left_a.dom (a' \ \) = Left_a.cod (\ \ dom \)" using a' \ \ 2 3 4 Left_a.dom_char Left_a.cod_char by (metis Left_a.seqE Left_a.seq_char hcomp_simps\<^sub>W\<^sub>C(1) in_homE paste_1) qed thus ?thesis using Left_a.comp_char Left_a.arr_char left_def by auto qed also have "... = a' \ \ \ \ \ dom \" using a' \ \ interchange hom_connected by auto also have "... = \ \ \" using \ \ comp_arr_dom comp_cod_arr by auto finally show ?thesis using H\<^sub>L_def by simp qed show "Left_a.comp (H\<^sub>L \ (Left_a.cod \)) (Left_a.L \) = H\<^sub>L \ \" proof - have "Left_a.comp (H\<^sub>L \ (Left_a.cod \)) (Left_a.L \) = Left_a.comp (\ \ cod \) (a \ \)" using assms 1 2 \ \ Left_a.cod_char left_def H\<^sub>L_def by simp also have "... = (\ \ cod \) \ (a \ \)" proof - have "Left_a.seq (\ \ cod \) (a \ \)" proof (intro Left_a.seqI) show 3: "Left_a.arr (\ \ cod \)" using \ \ 2 Left_a.arr_char left_def by (metis (no_types, lifting) H\<^sub>L_def L.preserves_arr hcomp_simps\<^sub>W\<^sub>C(1) in_homE right_connected paste_2) show 4: "Left_a.arr (a \ \)" using assms \ Left_a.arr_char left_def using H\<^sub>L_def L.preserves_arr by auto show "Left_a.dom (\ \ cod \) = Left_a.cod (a \ \)" using assms \ \ 2 3 4 Left_a.dom_char Left_a.cod_char by (metis Left_a.seqE Left_a.seq_char hcomp_simps\<^sub>W\<^sub>C(1) in_homE paste_2) qed thus ?thesis using Left_a.comp_char Left_a.arr_char left_def by auto qed also have "... = \ \ a \ cod \ \ \" using \ \ interchange hom_connected by auto also have "... = \ \ \" using \ \ comp_arr_dom comp_cod_arr by auto finally show ?thesis using H\<^sub>L_def by simp qed next fix \ assume \: "Left_a.ide \" have 1: "\ \ \ \ null" using assms \ \ Left_a.ide_char Left_a.arr_char left_def hom_connected by auto show "Left_a.iso (H\<^sub>L \ \)" proof - have "iso (\ \ \)" proof - have "a \ sources \ \ targets \" using assms \ \ 1 hom_connected weak_unit_self_composable Left_a.ide_char Left_a.arr_char left_def connected_if_composable by auto thus ?thesis using \ \ Left_a.ide_char ide_is_iso iso_hcomp\<^sub>R\<^sub>W\<^sub>C(1) by blast qed moreover have "left a (\ \ \)" using assms 1 \ weak_unit_self_composable hom_connected(2) [of a \] left_def match_4 null_agreement by auto ultimately show ?thesis using Left_a.iso_char Left_a.arr_char left_iff_left_inv Left_a.inv_char H\<^sub>L_def by metis qed qed interpret L': equivalence_functor \Left a'\ \Left a'\ \H\<^sub>L a'\ proof - have "naturally_isomorphic (Left a') (Left a') (H\<^sub>L a') (identity_functor.map (Left a'))" proof - have "naturally_isomorphic (Left a) (Left a) (H\<^sub>L a') (identity_functor.map (Left a))" - using assms Left_a.natural_isomorphism_\ naturally_isomorphic_def - \.natural_isomorphism_axioms naturally_isomorphic_symmetric - naturally_isomorphic_transitive - by blast + by (meson Left_a.natural_isomorphism_\ \.natural_isomorphism_axioms + naturally_isomorphic_def naturally_isomorphic_symmetric + naturally_isomorphic_transitive) thus ?thesis using 1 by auto qed thus "equivalence_functor (Left a') (Left a') (H\<^sub>L a')" using 1 L'.isomorphic_to_identity_is_equivalence naturally_isomorphic_def by fastforce qed text \ Now we do the same for \R'\. \ interpret R: fully_faithful_functor \Right a\ \Right a\ \H\<^sub>R a\ using assms weak_unit_def by simp interpret R': endofunctor \Right a\ \H\<^sub>R a'\ using a' 1 endofunctor_H\<^sub>R [of a'] by auto interpret \: natural_isomorphism \Right a\ \Right a\ \H\<^sub>R a\ \H\<^sub>R a'\ \H\<^sub>R \\ proof fix \ show "\ Right_a.arr \ \ H\<^sub>R \ \ = Right_a.null" using right_def \ H\<^sub>R_def hom_connected Right_a.null_char Right_a.arr_char by auto assume "Right_a.arr \" hence \: "Right_a.arr \ \ arr \ \ \ \ a \ null" using Right_a.arr_char right_def composable_implies_arr by simp have 2: "\ \ \ \ null" using assms \ \ Right_a.arr_char right_def hom_connected by auto show "Right_a.dom (H\<^sub>R \ \) = H\<^sub>R a (Right_a.dom \)" using assms 2 \ \ Right_a.arr_char right_def hom_connected(1) [of \ a] weak_unit_self_composable match_3 Right_a.dom_char H\<^sub>R_def by auto show "Right_a.cod (H\<^sub>R \ \) = H\<^sub>R a' (Right_a.cod \)" using assms 2 a' \ \ Right_a.arr_char right_def hom_connected(3) [of \ a] weak_unit_self_composable match_3 Right_a.cod_char H\<^sub>R_def by auto show "Right_a.comp (H\<^sub>R a' \) (H\<^sub>R \ (Right_a.dom \)) = H\<^sub>R \ \" proof - have "Right_a.comp (H\<^sub>R a' \) (H\<^sub>R \ (Right_a.dom \)) = Right_a.comp (\ \ a') (dom \ \ \)" using assms 1 2 \ \ Right_a.dom_char right_def H\<^sub>R_def by simp also have "... = (\ \ a') \ (dom \ \ \)" proof - have "Right_a.seq (\ \ a') (dom \ \ \)" proof (intro Right_a.seqI) show 3: "Right_a.arr (dom \ \ \)" using assms 2 \ \ Right_a.arr_char right_def by (metis H\<^sub>R_def R'.preserves_arr hcomp_simps\<^sub>W\<^sub>C(1) in_homE left_connected paste_2) show 4: "Right_a.arr (\ \ a')" using \ H\<^sub>R_def R'.preserves_arr by auto show "Right_a.dom (\ \ a') = Right_a.cod (dom \ \ \)" using a' \ \ 2 3 4 Right_a.dom_char Right_a.cod_char by (metis Right_a.seqE Right_a.seq_char hcomp_simps\<^sub>W\<^sub>C(1) in_homE paste_2) qed thus ?thesis using Right_a.comp_char Right_a.arr_char right_def by auto qed also have "... = \ \ dom \ \ a' \ \" using a' \ \ interchange hom_connected by auto also have "... = \ \ \" using \ \ comp_arr_dom comp_cod_arr by auto finally show ?thesis using H\<^sub>R_def by simp qed show "Right_a.comp (H\<^sub>R \ (Right_a.cod \)) (Right_a.R \) = H\<^sub>R \ \" proof - have "Right_a.comp (H\<^sub>R \ (Right_a.cod \)) (Right_a.R \) = Right_a.comp (cod \ \ \) (\ \ a)" using assms 1 2 \ \ Right_a.cod_char right_def H\<^sub>R_def by simp also have "... = (cod \ \ \) \ (\ \ a)" proof - have "Right_a.seq (cod \ \ \) (\ \ a)" proof (intro Right_a.seqI) show 3: "Right_a.arr (cod \ \ \)" using \ \ 2 Right_a.arr_char right_def by (metis (no_types, lifting) H\<^sub>R_def R.preserves_arr hcomp_simps\<^sub>W\<^sub>C(1) in_homE left_connected paste_1) show 4: "Right_a.arr (\ \ a)" using assms \ Right_a.arr_char right_def using H\<^sub>R_def R.preserves_arr by auto show "Right_a.dom (cod \ \ \) = Right_a.cod (\ \ a)" using assms \ \ 2 3 4 Right_a.dom_char Right_a.cod_char by (metis Right_a.seqE Right_a.seq_char hcomp_simps\<^sub>W\<^sub>C(1) in_homE paste_1) qed thus ?thesis using Right_a.comp_char Right_a.arr_char right_def by auto qed also have "... = cod \ \ \ \ \ \ a" using \ \ interchange hom_connected by auto also have "... = \ \ \" using \ \ comp_arr_dom comp_cod_arr by auto finally show ?thesis using H\<^sub>R_def by simp qed next fix \ assume \: "Right_a.ide \" have 1: "\ \ \ \ null" using assms \ \ Right_a.ide_char Right_a.arr_char right_def hom_connected by auto show "Right_a.iso (H\<^sub>R \ \)" proof - have "iso (\ \ \)" proof - have "a \ targets \ \ sources \" using assms \ \ 1 hom_connected weak_unit_self_composable Right_a.ide_char Right_a.arr_char right_def connected_if_composable by (metis (full_types) IntI targetsI) thus ?thesis using \ \ Right_a.ide_char ide_is_iso iso_hcomp\<^sub>R\<^sub>W\<^sub>C(1) by blast qed moreover have "right a (\ \ \)" using assms 1 \ weak_unit_self_composable hom_connected(1) [of \ a] right_def match_3 null_agreement by auto ultimately show ?thesis using Right_a.iso_char Right_a.arr_char right_iff_right_inv Right_a.inv_char H\<^sub>R_def by metis qed qed interpret R': equivalence_functor \Right a'\ \Right a'\ \H\<^sub>R a'\ proof - have "naturally_isomorphic (Right a') (Right a') (H\<^sub>R a') (identity_functor.map (Right a'))" - using assms Right_a.natural_isomorphism_\ naturally_isomorphic_def - \.natural_isomorphism_axioms naturally_isomorphic_symmetric - naturally_isomorphic_transitive proof - have "naturally_isomorphic (Right a) (Right a) (H\<^sub>R a') Right_a.map" - using assms Right_a.natural_isomorphism_\ naturally_isomorphic_def - \.natural_isomorphism_axioms naturally_isomorphic_symmetric - naturally_isomorphic_transitive - by blast + by (meson Right_a.natural_isomorphism_\ \.natural_isomorphism_axioms + naturally_isomorphic_def naturally_isomorphic_symmetric + naturally_isomorphic_transitive) thus ?thesis using 1 by auto qed thus "equivalence_functor (Right a') (Right a') (H\<^sub>R a')" using 1 R'.isomorphic_to_identity_is_equivalence naturally_isomorphic_def by fastforce qed show "weak_unit a'" using weak_unit_def iso L'.fully_faithful_functor_axioms R'.fully_faithful_functor_axioms by blast qed lemma sources_iso_closed: assumes "a \ sources \" and "a \ a'" shows "a' \ sources \" using assms isomorphism_respects_weak_units isomorphic_implies_equicomposable by blast lemma targets_iso_closed: assumes "a \ targets \" and "a \ a'" shows "a' \ targets \" using assms isomorphism_respects_weak_units isomorphic_implies_equicomposable by blast lemma sources_eqI: assumes "sources \ \ sources \ \ {}" shows "sources \ = sources \" using assms sources_iso_closed sources_are_isomorphic by blast lemma targets_eqI: assumes "targets \ \ targets \ \ {}" shows "targets \ = targets \" using assms targets_iso_closed targets_are_isomorphic by blast text \ The sets of sources and targets of a weak unit are isomorphism classes. \ lemma sources_char: assumes "weak_unit a" shows "sources a = {x. x \ a}" using assms sources_iso_closed weak_unit_iff_self_source sources_are_isomorphic isomorphic_symmetric by blast lemma targets_char: assumes "weak_unit a" shows "targets a = {x. x \ a}" using assms targets_iso_closed weak_unit_iff_self_target targets_are_isomorphic isomorphic_symmetric by blast end section "Horizontal Homs" text \ Here we define a locale that axiomatizes a (vertical) category \V\ that has been punctuated into ``horizontal homs'' by the choice of idempotent endofunctors \src\ and \trg\ that assign a specific ``source'' and ``target'' 1-cell to each of its arrows. The functors \src\ and \trg\ are also subject to further conditions that constrain how they commute with each other. \ locale horizontal_homs = category V + src: endofunctor V src + trg: endofunctor V trg for V :: "'a comp" (infixr "\" 55) and src :: "'a \ 'a" and trg :: "'a \ 'a" + assumes ide_src [simp]: "arr \ \ ide (src \)" and ide_trg [simp]: "arr \ \ ide (trg \)" and src_src [simp]: "arr \ \ src (src \) = src \" and trg_trg [simp]: "arr \ \ trg (trg \) = trg \" and trg_src [simp]: "arr \ \ trg (src \) = src \" and src_trg [simp]: "arr \ \ src (trg \) = trg \" begin no_notation in_hom ("\_ : _ \ _\") notation in_hom ("\_ : _ \ _\") text \ We define an \emph{object} to be an arrow that is its own source (or equivalently, its own target). \ definition obj where "obj a \ arr a \ src a = a" lemma obj_def': shows "obj a \ arr a \ trg a = a" using trg_src src_trg obj_def by metis lemma objI_src: assumes "arr a" and "src a = a" shows "obj a" using assms obj_def by simp lemma objI_trg: assumes "arr a" and "trg a = a" shows "obj a" using assms obj_def' by simp lemma objE [elim]: assumes "obj a" and "\ ide a; src a = a; trg a = a \ \ T" shows T using assms obj_def obj_def' ide_src ide_trg by metis (* * I believe I have sorted out the looping issues that were making these less than * useful, but do not make them default simps because it slows things down too much * and they are not used all that often. *) lemma obj_simps (* [simp] *): assumes "obj a" shows "arr a" and "src a = a" and "trg a = a" and "dom a = a" and "cod a = a" using assms by auto lemma obj_src [intro, simp]: assumes "arr \" shows "obj (src \)" using assms objI_src by auto lemma obj_trg [intro, simp]: assumes "arr \" shows "obj (trg \)" using assms objI_trg by auto definition in_hhom ("\_ : _ \ _\") where "in_hhom \ a b \ arr \ \ src \ = a \ trg \ = b" abbreviation hhom where "hhom a b \ {\. \\ : a \ b\}" abbreviation (input) hseq\<^sub>H\<^sub>H where "hseq\<^sub>H\<^sub>H \ \\ \. arr \ \ arr \ \ src \ = trg \" lemma in_hhomI [intro, simp]: assumes "arr \" and "src \ = a" and "trg \ = b" shows "\\ : a \ b\" using assms in_hhom_def by auto lemma in_hhomE [elim]: assumes "\\ : a \ b\" and "\ arr \; obj a; obj b; src \ = a; trg \ = b \ \ T" shows "T" using assms in_hhom_def by auto (* * TODO: I tried removing the second assertion here, thinking that it should already * be covered by the category locale, but in fact it breaks some proofs in * SpanBicategory that ought to be trivial. So it seems that the presence of * this introduction rule adds something, and I should consider whether this rule * should be added to the category locale. *) lemma ide_in_hom [intro]: assumes "ide f" shows "\f : src f \ trg f\" and "\f : f \ f\" using assms by auto lemma src_dom [simp]: shows "src (dom \) = src \" by (metis arr_dom_iff_arr obj_simps(4) obj_src src.is_extensional src.preserves_dom) lemma src_cod [simp]: shows "src (cod \) = src \" by (metis arr_cod_iff_arr obj_simps(5) obj_src src.is_extensional src.preserves_cod) lemma trg_dom [simp]: shows "trg (dom \) = trg \" by (metis arr_dom_iff_arr ide_char ide_trg trg.is_extensional trg.preserves_dom) lemma trg_cod [simp]: shows "trg (cod \) = trg \" by (metis arr_cod_iff_arr ide_char ide_trg trg.is_extensional trg.preserves_cod) (* * TODO: In theory, the following simps should already be available from the fact * that src and trg are endofunctors. But they seem not to get used. *) lemma dom_src [simp]: shows "dom (src \) = src \" by (metis dom_null ideD(2) ide_src src.is_extensional) lemma cod_src [simp]: shows "cod (src \) = src \" by (metis cod_null ideD(3) ide_src src.is_extensional) lemma dom_trg [simp]: shows "dom (trg \) = trg \" by (metis dom_null ideD(2) ide_trg trg.is_extensional) lemma cod_trg [simp]: shows "cod (trg \) = trg \" by (metis cod_null ideD(3) ide_trg trg.is_extensional) lemma vcomp_in_hhom [intro, simp]: assumes "seq \ \" and "src \ = a" and "trg \ = b" shows "\\ \ \ : a \ b\" using assms src_cod [of "\ \ \"] trg_cod [of "\ \ \"] by auto lemma src_vcomp [simp]: assumes "seq \ \" shows "src (\ \ \) = src \" using assms src_cod [of "\ \ \"] by auto lemma trg_vcomp [simp]: assumes "seq \ \" shows "trg (\ \ \) = trg \" using assms trg_cod [of "\ \ \"] by auto lemma vseq_implies_hpar: assumes "seq \ \" shows "src \ = src \" and "trg \ = trg \" using assms src_dom [of "\ \ \"] trg_dom [of "\ \ \"] src_cod [of "\ \ \"] trg_cod [of "\ \ \"] by auto lemma vconn_implies_hpar: assumes "\\ : f \ g\" shows "src \ = src f" and "trg \ = trg f" and "src g = src f" and "trg g = trg f" using assms by auto lemma src_inv [simp]: assumes "iso \" shows "src (inv \) = src \" using assms inv_in_hom iso_is_arr src_dom src_cod iso_inv_iso dom_inv by metis lemma trg_inv [simp]: assumes "iso \" shows "trg (inv \) = trg \" using assms inv_in_hom iso_is_arr trg_dom trg_cod iso_inv_iso cod_inv by metis lemma inv_in_hhom [intro, simp]: assumes "iso \" and "src \ = a" and "trg \ = b" shows "\inv \ : a \ b\" using assms iso_is_arr by simp lemma hhom_is_subcategory: shows "subcategory V (\\. \\ : a \ b\)" using src_dom trg_dom src_cod trg_cod by (unfold_locales, auto) lemma isomorphic_objects_are_equal: assumes "obj a" and "obj b" and "a \ b" shows "a = b" using assms isomorphic_def by (metis dom_inv in_homE objE src_dom src_inv) text \ Having the functors \src\ and \trg\ allows us to form categories VV and VVV of formally horizontally composable pairs and triples of arrows. \ sublocale VxV: product_category V V .. sublocale VV: subcategory VxV.comp \\\\. hseq\<^sub>H\<^sub>H (fst \\) (snd \\)\ by (unfold_locales, auto) lemma subcategory_VV: shows "subcategory VxV.comp (\\\. hseq\<^sub>H\<^sub>H (fst \\) (snd \\))" .. sublocale VxVxV: product_category V VxV.comp .. sublocale VVV: subcategory VxVxV.comp \\\\\. arr (fst \\\) \ VV.arr (snd \\\) \ src (fst \\\) = trg (fst (snd \\\))\ using VV.arr_char by (unfold_locales, auto) lemma subcategory_VVV: shows "subcategory VxVxV.comp (\\\\. arr (fst \\\) \ VV.arr (snd \\\) \ src (fst \\\) = trg (fst (snd \\\)))" .. end subsection "Prebicategories with Homs" text \ A \emph{weak composition with homs} consists of a weak composition that is equipped with horizontal homs in such a way that the chosen source and target of each 2-cell \\\ in fact lie in the set of sources and targets, respectively, of \\\, such that horizontal composition respects the chosen sources and targets, and such that if 2-cells \\\ and \\\ are horizontally composable, then the chosen target of \\\ coincides with the chosen source of \\\. \ locale weak_composition_with_homs = weak_composition + horizontal_homs + assumes src_in_sources: "arr \ \ src \ \ sources \" and trg_in_targets: "arr \ \ trg \ \ targets \" and src_hcomp': "\ \ \ \ null \ src (\ \ \) = src \" and trg_hcomp': "\ \ \ \ null \ trg (\ \ \) = trg \" and seq_if_composable: "\ \ \ \ null \ src \ = trg \" locale prebicategory_with_homs = prebicategory + weak_composition_with_homs begin lemma composable_char\<^sub>P\<^sub>B\<^sub>H: shows "\ \ \ \ null \ arr \ \ arr \ \ src \ = trg \" using trg_in_targets src_in_sources composable_if_connected sourcesD(3) targets_determine_composability seq_if_composable composable_implies_arr by metis lemma hcomp_in_hom\<^sub>P\<^sub>B\<^sub>H: assumes "\\ : a \\<^sub>W\<^sub>C b\" and "\\ : b \\<^sub>W\<^sub>C c\" shows "\\ \ \ : a \\<^sub>W\<^sub>C c\" and "\\ \ \ : dom \ \ dom \ \ cod \ \ cod \\" proof - show "\\ \ \ : a \\<^sub>W\<^sub>C c\" using assms sources_determine_composability sources_hcomp targets_hcomp by auto thus "\\ \ \ : dom \ \ dom \ \ cod \ \ cod \\" using assms by auto qed text \ In a prebicategory with homs, if \a\ is an object (i.e. \src a = a\ and \trg a = a\), then \a\ is a weak unit. The converse need not hold: there can be weak units that the \src\ and \trg\ mappings send to other 1-cells in the same isomorphism class. \ lemma obj_is_weak_unit: assumes "obj a" shows "weak_unit a" proof - have "a \ sources a" using assms objE src_in_sources ideD(1) by metis thus ?thesis by auto qed end subsection "Choosing Homs" text \ Every prebicategory extends to a prebicategory with homs, by choosing an arbitrary representative of each isomorphism class of weak units to serve as an object. ``The source'' of a 2-cell is defined to be the chosen representative of the set of all its sources (which is an isomorphism class), and similarly for ``the target''. \ context prebicategory begin definition rep where "rep f \ SOME f'. f' \ { f'. f \ f' }" definition some_src where "some_src \ \ if arr \ then rep (SOME a. a \ sources \) else null" definition some_trg where "some_trg \ \ if arr \ then rep (SOME b. b \ targets \) else null" lemma isomorphic_ide_rep: assumes "ide f" shows "f \ rep f" proof - have "\f'. f' \ { f'. f \ f' }" using assms isomorphic_reflexive by blast thus ?thesis using rep_def someI_ex by simp qed lemma rep_rep: assumes "ide f" shows "rep (rep f) = rep f" proof - have "rep f \ { f'. f \ f' }" using assms isomorphic_ide_rep by simp have "{ f'. f \ f' } = { f'. rep f \ f' }" proof - have "\f'. f \ f' \ rep f \ f'" proof fix f' assume f': "f \ f'" show "rep f \ f'" proof - obtain \ where \: "\ \ hom f f' \ iso \" using f' by auto obtain \ where \: "\ \ hom f (rep f) \ iso \" using assms isomorphic_ide_rep by blast have "inv \ \ hom (rep f) f \ iso (inv \)" using \ by simp hence "iso (V \ (inv \)) \ V \ (inv \) \ hom (rep f) f'" using \ isos_compose by auto thus ?thesis using isomorphic_def by auto qed next fix f' assume f': "rep f \ f'" show "f \ f'" using assms f' isomorphic_ide_rep isos_compose isomorphic_def by (meson isomorphic_transitive) qed thus ?thesis by auto qed hence "rep (rep f) = (SOME f'. f' \ { f'. f \ f' })" using assms rep_def by fastforce also have "... = rep f" using assms rep_def by simp finally show ?thesis by simp qed lemma some_src_in_sources: assumes "arr \" shows "some_src \ \ sources \" proof - have 1: "(SOME a. a \ sources \) \ sources \" using assms arr_iff_has_src someI_ex [of "\a. a \ sources \"] by blast moreover have "ide (SOME a. a \ sources \)" using 1 weak_unit_self_composable by auto ultimately show ?thesis using assms 1 some_src_def sources_iso_closed isomorphic_ide_rep by metis qed lemma some_trg_in_targets: assumes "arr \" shows "some_trg \ \ targets \" proof - have 1: "(SOME a. a \ targets \) \ targets \" using assms arr_iff_has_trg someI_ex [of "\a. a \ targets \"] by blast moreover have "ide (SOME a. a \ targets \)" using 1 weak_unit_self_composable by auto ultimately show ?thesis using assms 1 some_trg_def targets_iso_closed isomorphic_ide_rep by metis qed lemma some_src_dom: assumes "arr \" shows "some_src (dom \) = some_src \" using assms some_src_def sources_dom by simp lemma some_src_cod: assumes "arr \" shows "some_src (cod \) = some_src \" using assms some_src_def sources_cod by simp lemma some_trg_dom: assumes "arr \" shows "some_trg (dom \) = some_trg \" using assms some_trg_def targets_dom by simp lemma some_trg_cod: assumes "arr \" shows "some_trg (cod \) = some_trg \" using assms some_trg_def targets_cod by simp lemma ide_some_src: assumes "arr \" shows "ide (some_src \)" using assms some_src_in_sources weak_unit_self_composable by blast lemma ide_some_trg: assumes "arr \" shows "ide (some_trg \)" using assms some_trg_in_targets weak_unit_self_composable by blast lemma some_src_composable: assumes "arr \" shows "\ \ \ \ null \ some_src \ \ \ \ null" using assms some_src_in_sources sources_determine_composability by blast lemma some_trg_composable: assumes "arr \" shows "\ \ \ \ null \ \ \ some_trg \ \ null" using assms some_trg_in_targets targets_determine_composability by blast lemma sources_some_src: assumes "arr \" shows "sources (some_src \) = sources \" using assms sources_determine_composability some_src_in_sources by blast lemma targets_some_trg: assumes "arr \" shows "targets (some_trg \) = targets \" using assms targets_determine_composability some_trg_in_targets by blast lemma src_some_src: assumes "arr \" shows "some_src (some_src \) = some_src \" using assms some_src_def ide_some_src sources_some_src by force lemma trg_some_trg: assumes "arr \" shows "some_trg (some_trg \) = some_trg \" using assms some_trg_def ide_some_trg targets_some_trg by force lemma sources_char': assumes "arr \" shows "a \ sources \ \ some_src \ \ a" using assms some_src_in_sources sources_iso_closed sources_are_isomorphic by meson lemma targets_char': assumes "arr \" shows "a \ targets \ \ some_trg \ \ a" - using assms some_trg_in_targets targets_iso_closed targets_are_isomorphic by blast + using assms some_trg_in_targets targets_iso_closed targets_are_isomorphic by meson text \ An arbitrary choice of sources and targets in a prebicategory results in a notion of formal composability that coincides with the actual horizontal composability of the prebicategory. \ lemma composable_char\<^sub>P\<^sub>B: shows "\ \ \ \ null \ arr \ \ arr \ \ some_src \ = some_trg \" proof assume \\: "\ \ \ \ null" show "arr \ \ arr \ \ some_src \ = some_trg \" using \\ composable_implies_arr connected_if_composable some_src_def some_trg_def by force next assume \\: "arr \ \ arr \ \ some_src \ = some_trg \" show "\ \ \ \ null" using \\ some_src_in_sources some_trg_composable by force qed text \ A 1-cell is its own source if and only if it is its own target. \ lemma self_src_iff_self_trg: assumes "ide a" shows "a = some_src a \ a = some_trg a" proof assume a: "a = some_src a" have "weak_unit a \ a \ a \ null" using assms a some_src_in_sources [of a] by force thus "a = some_trg a" using a composable_char\<^sub>P\<^sub>B by simp next assume a: "a = some_trg a" have "weak_unit a \ a \ a \ null" using assms a some_trg_in_targets [of a] by force thus "a = some_src a" using a composable_char\<^sub>P\<^sub>B by simp qed lemma some_trg_some_src: assumes "arr \" shows "some_trg (some_src \) = some_src \" using assms ide_some_src some_src_def some_trg_def some_src_in_sources sources_char targets_char sources_some_src by force lemma src_some_trg: assumes "arr \" shows "some_src (some_trg \) = some_trg \" using assms ide_some_trg some_src_def some_trg_def some_trg_in_targets sources_char targets_char targets_some_trg by force lemma some_src_eqI: assumes "a \ sources \" and "some_src a = a" shows "some_src \ = a" using assms sources_char' some_src_def some_src_in_sources sources_are_isomorphic isomorphic_symmetric isomorphic_transitive by (metis composable_char\<^sub>P\<^sub>B sourcesD(3)) lemma some_trg_eqI: assumes "b \ targets \" and "some_trg b = b" shows "some_trg \ = b" using assms targets_char' some_trg_def some_trg_in_targets targets_are_isomorphic isomorphic_symmetric isomorphic_transitive by (metis composable_char\<^sub>P\<^sub>B targetsD(3)) lemma some_src_comp: assumes "\ \ \ \ null" shows "some_src (\ \ \) = some_src \" proof (intro some_src_eqI [of "some_src \" "\ \ \"]) show "some_src (some_src \) = some_src \" using assms src_some_src composable_implies_arr by simp show "some_src \ \ sources (H \ \)" using assms some_src_in_sources composable_char\<^sub>P\<^sub>B by (simp add: sources_hcomp) qed lemma some_trg_comp: assumes "\ \ \ \ null" shows "some_trg (\ \ \) = some_trg \" proof (intro some_trg_eqI [of "some_trg \" "\ \ \"]) show "some_trg (some_trg \) = some_trg \" using assms trg_some_trg composable_implies_arr by simp show "some_trg \ \ targets (H \ \)" using assms some_trg_in_targets composable_char\<^sub>P\<^sub>B by (simp add: targets_hcomp) qed text \ The mappings that take an arrow to its chosen source or target are endofunctors of the vertical category, which commute with each other in the manner required for horizontal homs. \ interpretation S: endofunctor V some_src using some_src_def ide_some_src some_src_dom some_src_cod apply unfold_locales apply auto[4] proof - fix \ \ assume \\: "seq \ \" show "some_src (\ \ \) = some_src \ \ some_src \" using \\ some_src_dom [of "\ \ \"] some_src_dom some_src_cod [of "\ \ \"] some_src_cod ide_some_src by auto qed interpretation T: endofunctor V some_trg using some_trg_def ide_some_trg some_trg_dom some_trg_cod apply unfold_locales apply auto[4] proof - fix \ \ assume \\: "seq \ \" show "some_trg (\ \ \) = some_trg \ \ some_trg \" using \\ some_trg_dom [of "\ \ \"] some_trg_dom some_trg_cod [of "\ \ \"] some_trg_cod ide_some_trg by auto qed interpretation weak_composition_with_homs V H some_src some_trg apply unfold_locales using some_src_in_sources some_trg_in_targets src_some_src trg_some_trg src_some_trg some_trg_some_src some_src_comp some_trg_comp composable_char\<^sub>P\<^sub>B ide_some_src ide_some_trg by simp_all proposition extends_to_weak_composition_with_homs: shows "weak_composition_with_homs V H some_src some_trg" .. proposition extends_to_prebicategory_with_homs: shows "prebicategory_with_homs V H \ some_src some_trg" .. end subsection "Choosing Units" text \ A \emph{prebicategory with units} is a prebicategory equipped with a choice, for each weak unit \a\, of a ``unit isomorphism'' \\\[a] : a \ a \ a\\. \ locale prebicategory_with_units = prebicategory V H \ + weak_composition V H for V :: "'a comp" (infixr "\" 55) and H :: "'a comp" (infixr "\" 53) and \ :: "'a \ 'a \ 'a \ 'a" ("\[_, _, _]") and \ :: "'a \ 'a" ("\[_]") + assumes unit_in_vhom\<^sub>P\<^sub>B\<^sub>U: "weak_unit a \ \\[a] : a \ a \ a\" and iso_unit\<^sub>P\<^sub>B\<^sub>U: "weak_unit a \ iso \[a]" begin lemma unit_in_hom\<^sub>P\<^sub>B\<^sub>U: assumes "weak_unit a" shows "\\[a] : a \\<^sub>W\<^sub>C a\" and "\\[a] : a \ a \ a\" proof - show 1: "\\[a] : a \ a \ a\" using assms unit_in_vhom\<^sub>P\<^sub>B\<^sub>U by auto show "\\[a] : a \\<^sub>W\<^sub>C a\" using assms 1 weak_unit_iff_self_source weak_unit_iff_self_target sources_cod [of "\[a]"] targets_cod [of "\[a]"] by (elim in_homE, auto) qed lemma unit_simps [simp]: assumes "weak_unit a" shows "arr \[a]" and "dom \[a] = a \ a" and "cod \[a] = a" using assms unit_in_vhom\<^sub>P\<^sub>B\<^sub>U by auto end text \ Every prebicategory extends to a prebicategory with units, simply by choosing the unit isomorphisms arbitrarily. \ context prebicategory begin proposition extends_to_prebicategory_with_units: shows "prebicategory_with_units V H \ some_unit" using iso_some_unit by (unfold_locales, auto) end subsection "Horizontal Composition" text \ The following locale axiomatizes a (vertical) category \V\ with horizontal homs, which in addition has been equipped with a functorial operation \H\ of horizontal composition from \VV\ to \V\, assumed to preserve source and target. \ locale horizontal_composition = horizontal_homs V src trg + H: "functor" VV.comp V \\\\. H (fst \\) (snd \\)\ for V :: "'a comp" (infixr "\" 55) and H :: "'a \ 'a \ 'a" (infixr "\" 53) and src :: "'a \ 'a" and trg :: "'a \ 'a" + assumes src_hcomp: "arr (\ \ \) \ src (\ \ \) = src \" and trg_hcomp: "arr (\ \ \) \ trg (\ \ \) = trg \" begin (* TODO: Why does this get re-introduced? *) no_notation in_hom ("\_ : _ \ _\") text \ \H\ is a partial magma, which shares its null with \V\. \ lemma is_partial_magma: shows "partial_magma H" and "partial_magma.null H = null" proof - have 1: "\f. null \ f = null \ f \ null = null" using H.is_extensional VV.arr_char not_arr_null by auto interpret H: partial_magma H using 1 VV.arr_char H.is_extensional not_arr_null by unfold_locales metis show "partial_magma H" .. show "H.null = null" using 1 H.null_def the1_equality [of "\n. \f. n \ f = n \ f \ n = n"] by metis qed text \ \textbf{Note:} The following is ``almost'' \H.seq\, but for that we would need \H.arr = V.arr\. This would be unreasonable to expect, in general, as the definition of \H.arr\ is based on ``strict'' units rather than weak units. Later we will show that we do have \H.arr = V.arr\ if the vertical category is discrete. \ abbreviation hseq where "hseq \ \ \ arr (\ \ \)" lemma hseq_char: shows "hseq \ \ \ arr \ \ arr \ \ src \ = trg \" proof - have "hseq \ \ \ VV.arr (\, \)" using H.is_extensional H.preserves_arr by force also have "... \ arr \ \ arr \ \ src \ = trg \" using VV.arr_char by force finally show ?thesis by blast qed lemma hseq_char': shows "hseq \ \ \ \ \ \ \ null" using VV.arr_char H.preserves_arr H.is_extensional hseq_char [of \ \] by auto lemma hseqI' [intro, simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "hseq \ \" using assms hseq_char by simp lemma hseqI: assumes "\\ : a \ b\" and "\\ : b \ c\" shows "hseq \ \" using assms hseq_char by auto lemma hseqE [elim]: assumes "hseq \ \" and "arr \ \ arr \ \ src \ = trg \ \ T" shows "T" using assms hseq_char by simp lemma hcomp_simps [simp]: assumes "hseq \ \" shows "src (\ \ \) = src \" and "trg (\ \ \) = trg \" and "dom (\ \ \) = dom \ \ dom \" and "cod (\ \ \) = cod \ \ cod \" using assms VV.arr_char src_hcomp apply blast using assms VV.arr_char trg_hcomp apply blast using assms VV.arr_char H.preserves_dom VV.dom_simp apply force using assms VV.arr_char H.preserves_cod VV.cod_simp by force lemma ide_hcomp [intro, simp]: assumes "ide \" and "ide \" and "src \ = trg \" shows "ide (\ \ \)" using assms VV.ide_char VV.arr_char H.preserves_ide [of "(\, \)"] by auto lemma hcomp_in_hhom [intro]: assumes "\\ : a \ b\" and "\\ : b \ c\" shows "\\ \ \ : a \ c\" using assms hseq_char by fastforce lemma hcomp_in_hhom' (* [simp] *): assumes "arr \" and "arr \" and "src \ = a" and "trg \ = c" and "src \ = trg \" shows "\\ \ \ : a \ c\" using assms hseq_char by fastforce lemma hcomp_in_hhomE [elim]: assumes "\\ \ \ : a \ c\" and "\ arr \; arr \; src \ = trg \; src \ = a; trg \ = c \ \ T" shows T using assms in_hhom_def by fastforce lemma hcomp_in_vhom [intro]: assumes "\\ : f \ g\" and "\\ : h \ k\" and "src h = trg f" shows "\\ \ \ : h \ f \ k \ g\" using assms by fastforce lemma hcomp_in_vhom' (* [simp] *): assumes "hseq \ \" and "dom \ = f" and "dom \ = h" and "cod \ = g" and "cod \ = k" assumes "\\ : f \ g\" and "\\ : h \ k\" and "src h = trg f" shows "\\ \ \ : h \ f \ k \ g\" using assms by fastforce lemma hcomp_in_vhomE [elim]: assumes "\\ \ \ : f \ g\" and "\ arr \; arr \; src \ = trg \; src \ = src f; src \ = src g; trg \ = trg f; trg \ = trg g \ \ T" shows T using assms in_hom_def by (metis in_homE hseqE src_cod src_dom src_hcomp trg_cod trg_dom trg_hcomp) text \ A horizontal composition yields a weak composition by simply forgetting the \src\ and \trg\ functors. \ lemma match_1: assumes "\ \ \ \ null" and "(\ \ \) \ \ \ null" shows "\ \ \ \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto lemma match_2: assumes "\ \ (\ \ \) \ null" and "\ \ \ \ null" shows "\ \ \ \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto lemma match_3: assumes "\ \ \ \ null" and "\ \ \ \ null" shows "(\ \ \) \ \ \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto lemma match_4: assumes "\ \ \ \ null" and "\ \ \ \ null" shows "\ \ (\ \ \) \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char hseq_char' by auto lemma left_connected: assumes "seq \ \'" shows "\ \ \ \ null \ \' \ \ \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char' by (metis hseq_char seqE vseq_implies_hpar(1)) lemma right_connected: assumes "seq \ \'" shows "H \ \ \ null \ H \ \' \ null" using assms H.is_extensional not_arr_null VV.arr_char hseq_char' by (metis hseq_char seqE vseq_implies_hpar(2)) proposition is_weak_composition: shows "weak_composition V H" proof - have 1: "(\\\. fst \\ \ snd \\ \ null) = (\\\. arr (fst \\) \ arr (snd \\) \ src (fst \\) = trg (snd \\))" using hseq_char' by auto interpret VoV: subcategory VxV.comp \\\\. fst \\ \ snd \\ \ null\ using 1 VV.subcategory_axioms by simp interpret H: "functor" VoV.comp V \\\\. fst \\ \ snd \\\ using H.functor_axioms 1 by simp show ?thesis using match_1 match_2 match_3 match_4 left_connected right_connected by (unfold_locales, metis) qed interpretation weak_composition V H using is_weak_composition by auto text \ It can be shown that \arr ((\ \ \) \ (\ \ \)) \ (\ \ \) \ (\ \ \) = (\ \ \) \ (\ \ \)\. However, we do not have \arr ((\ \ \) \ (\ \ \)) \ (\ \ \) \ (\ \ \) = (\ \ \) \ (\ \ \)\, because it does not follow from \arr ((\ \ \) \ (\ \ \))\ that \dom \ = cod \\ and \dom \ = cod \\, only that \dom \ \ dom \ = cod \ \ cod \\. So we don't get interchange unconditionally. \ lemma interchange: assumes "seq \ \" and "seq \ \" shows "(\ \ \) \ (\ \ \) = (\ \ \) \ (\ \ \)" using assms interchange by simp lemma whisker_right: assumes "ide f" and "seq \ \" shows "(\ \ \) \ f = (\ \ f) \ (\ \ f)" using assms whisker_right by simp lemma whisker_left: assumes "ide f" and "seq \ \" shows "f \ (\ \ \) = (f \ \) \ (f \ \)" using assms whisker_left by simp lemma inverse_arrows_hcomp: assumes "iso \" and "iso \" and "src \ = trg \" shows "inverse_arrows (\ \ \) (inv \ \ inv \)" proof - show "inverse_arrows (\ \ \) (inv \ \ inv \)" proof show "ide ((inv \ \ inv \) \ (\ \ \))" proof - have "(inv \ \ inv \) \ (\ \ \) = dom \ \ dom \" using assms interchange iso_is_arr comp_inv_arr' by (metis arr_dom) thus ?thesis using assms iso_is_arr by simp qed show "ide ((\ \ \) \ (inv \ \ inv \))" proof - have "(\ \ \) \ (inv \ \ inv \) = cod \ \ cod \" using assms interchange iso_is_arr comp_arr_inv' by (metis arr_cod) thus ?thesis using assms iso_is_arr by simp qed qed qed lemma iso_hcomp [intro, simp]: assumes "iso \" and "iso \" and "src \ = trg \" shows "iso (\ \ \)" using assms inverse_arrows_hcomp by auto (* * TODO: Maybe a good idea to change hcomp_in_vhom hypotheses to match this * and iso_hcomp. *) lemma hcomp_iso_in_hom [intro]: assumes "iso_in_hom \ f g" and "iso_in_hom \ h k" and "src \ = trg \" shows "iso_in_hom (\ \ \) (h \ f) (k \ g)" unfolding iso_in_hom_def using assms hcomp_in_vhom iso_hcomp iso_in_hom_def vconn_implies_hpar(1-2) by auto lemma isomorphic_implies_ide: assumes "f \ g" shows "ide f" and "ide g" using assms isomorphic_def by auto lemma hcomp_ide_isomorphic: assumes "ide f" and "g \ h" and "src f = trg g" shows "f \ g \ f \ h" proof - obtain \ where \: "iso \ \ \\ : g \ h\" using assms isomorphic_def by auto have "iso (f \ \) \ \f \ \ : f \ g \ f \ h\" using assms \ iso_hcomp by auto thus ?thesis using isomorphic_def by auto qed lemma hcomp_isomorphic_ide: assumes "f \ g" and "ide h" and "src f = trg h" shows "f \ h \ g \ h" proof - obtain \ where \: "iso \ \ \\ : f \ g\" using assms isomorphic_def by auto have "iso (\ \ h) \ \\ \ h : f \ h \ g \ h\" using assms \ iso_hcomp by auto thus ?thesis using isomorphic_def by auto qed lemma isomorphic_implies_hpar: assumes "f \ f'" shows "ide f" and "ide f'" and "src f = src f'" and "trg f = trg f'" using assms isomorphic_def by auto lemma inv_hcomp [simp]: assumes "iso \" and "iso \" and "src \ = trg \" shows "inv (\ \ \) = inv \ \ inv \" using assms inverse_arrow_unique [of "\ \ \"] inv_is_inverse inverse_arrows_hcomp by auto text \ The following define the two ways of using horizontal composition to compose three arrows. \ definition HoHV where "HoHV \ \ if VVV.arr \ then (fst \ \ fst (snd \)) \ snd (snd \) else null" definition HoVH where "HoVH \ \ if VVV.arr \ then fst \ \ fst (snd \) \ snd (snd \) else null" lemma functor_HoHV: shows "functor VVV.comp V HoHV" apply unfold_locales using VVV.arr_char VV.arr_char VVV.dom_char VVV.cod_char VVV.comp_char HoHV_def apply auto[4] proof - fix f g assume fg: "VVV.seq g f" show "HoHV (VVV.comp g f) = HoHV g \ HoHV f" proof - have "VxVxV.comp g f = (fst g \ fst f, fst (snd g) \ fst (snd f), snd (snd g) \ snd (snd f))" using fg VVV.seq_char VVV.arr_char VV.arr_char VxVxV.comp_char VxV.comp_char by (metis (no_types, lifting) VxV.seqE VxVxV.seqE) hence "HoHV (VVV.comp g f) = (fst g \ fst f \ fst (snd g) \ fst (snd f)) \ snd (snd g) \ snd (snd f)" using HoHV_def VVV.comp_simp fg by auto also have "... = ((fst g \ fst (snd g)) \ snd (snd g)) \ ((fst f \ fst (snd f)) \ snd (snd f))" using fg VVV.seq_char VVV.arr_char VV.arr_char interchange by (metis (no_types, lifting) VxV.seqE VxVxV.seqE hseqI' src_vcomp trg_vcomp) also have "... = HoHV g \ HoHV f" using HoHV_def fg by auto finally show ?thesis by simp qed qed sublocale HoHV: "functor" VVV.comp V HoHV using functor_HoHV by simp lemma functor_HoVH: shows "functor VVV.comp V HoVH" apply unfold_locales using VVV.arr_char VV.arr_char VVV.dom_char VVV.cod_char VVV.comp_char HoHV_def HoVH_def apply auto[4] proof - fix f g assume fg: "VVV.seq g f" show "HoVH (VVV.comp g f) = HoVH g \ HoVH f" proof - have "VxVxV.comp g f = (fst g \ fst f, fst (snd g) \ fst (snd f), snd (snd g) \ snd (snd f))" using fg VVV.seq_char VVV.arr_char VV.arr_char VxVxV.comp_char VxV.comp_char by (metis (no_types, lifting) VxV.seqE VxVxV.seqE) hence "HoVH (VVV.comp g f) = fst g \ fst f \ fst (snd g) \ fst (snd f) \ snd (snd g) \ snd (snd f)" using HoVH_def VVV.comp_simp fg by auto also have "... = (fst g \ fst (snd g) \ snd (snd g)) \ (fst f \ fst (snd f) \ snd (snd f))" using fg VVV.seq_char VVV.arr_char VV.arr_char interchange by (metis (no_types, lifting) VxV.seqE VxVxV.seqE hseqI' src_vcomp trg_vcomp) also have "... = HoVH g \ HoVH f" using fg VVV.seq_char VVV.arr_char HoVH_def VVV.comp_char VV.arr_char by (metis (no_types, lifting)) finally show ?thesis by simp qed qed sublocale HoVH: "functor" VVV.comp V HoVH using functor_HoVH by simp text \ The following define horizontal composition of an arrow on the left by its target and on the right by its source. \ abbreviation L where "L \ \\. if arr \ then trg \ \ \ else null" abbreviation R where "R \ \\. if arr \ then \ \ src \ else null" sublocale L: endofunctor V L using vseq_implies_hpar(2) whisker_left by (unfold_locales, auto) lemma endofunctor_L: shows "endofunctor V L" .. sublocale R: endofunctor V R using vseq_implies_hpar(1) whisker_right by (unfold_locales, auto) lemma endofunctor_R: shows "endofunctor V R" .. end end diff --git a/thys/Bicategory/Pseudofunctor.thy b/thys/Bicategory/Pseudofunctor.thy --- a/thys/Bicategory/Pseudofunctor.thy +++ b/thys/Bicategory/Pseudofunctor.thy @@ -1,3131 +1,2960 @@ (* Title: Pseudofunctor Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) section "Pseudofunctors" theory Pseudofunctor imports MonoidalCategory.MonoidalFunctor Bicategory Subbicategory InternalEquivalence Coherence begin text \ The traditional definition of a pseudofunctor \F : C \ D\ between bicategories \C\ and \D\ is in terms of two maps: an ``object map'' \F\<^sub>o\ that takes objects of \C\ to objects of \D\ and an ``arrow map'' \F\<^sub>a\ that assigns to each pair of objects \a\ and \b\ of \C\ a functor \F\<^sub>a a b\ from the hom-category \hom\<^sub>C a b\ to the hom-category \hom\<^sub>D (F\<^sub>o a) (F\<^sub>o b)\. In addition, there is assigned to each object \a\ of \C\ an invertible 2-cell \\\ a : F\<^sub>o a \\<^sub>D (F\<^sub>a a a) a\\, and to each pair \(f, g)\ of composable 1-cells of C there is assigned an invertible 2-cell \\\ (f, g) : F g \ F f \ F (g \ f)\\, all subject to naturality and coherence conditions. In keeping with the ``object-free'' style in which we have been working, we do not wish to adopt a definition of pseudofunctor that distinguishes between objects and other arrows. Instead, we would like to understand a pseudofunctor as an ordinary functor between (vertical) categories that weakly preserves horizontal composition in a suitable sense. So, we take as a starting point that a pseudofunctor \F : C \ D\ is a functor from \C\ to \D\, when these are regarded as ordinary categories with respect to vertical composition. Next, \F\ should preserve source and target, but only ``weakly'' (up to isomorphism, rather than ``on the nose''). Weak preservation of horizontal composition is expressed by specifying, for each horizontally composable pair of vertical identities \(f, g)\ of \C\, a ``compositor'' \\\ (f, g) : F g \ F f \ F (g \ f)\\ in \D\, such that the \\ (f, g)\ are the components of a natural isomorphism. Associators must also be weakly preserved by F; this is expressed by a coherence condition that relates an associator \\\<^sub>C[f, g, h]\ in \C\, its image \F \\<^sub>C[f, g, h]\, the associator \\\<^sub>D[F f, F g, F h]\ in \D\ and compositors involving \f\, \g\, and \h\. As regards the weak preservation of unitors, just as for monoidal functors, which are in fact pseudofunctors between one-object bicategories, it is only necessary to assume that \F \\<^sub>C[a]\ and \\\<^sub>D[F a]\ are isomorphic in \D\ for each object \a\ of \C\, for there is then a canonical way to obtain, for each \a\, an isomorphism \\\ a : src (F a) \ F a\\ that satisfies the usual coherence conditions relating the unitors and the associators. Note that the map \a \ src (F a)\ amounts to the traditional ``object map'' \F\<^sub>o\, so that this becomes a derived notion, rather than a primitive one. \ subsection "Weak Arrows of Homs" text \ We begin with a locale that defines a functor between ``horizontal homs'' that preserves source and target up to isomorphism. \ locale weak_arrow_of_homs = C: horizontal_homs C src\<^sub>C trg\<^sub>C + D: horizontal_homs D src\<^sub>D trg\<^sub>D + "functor" C D F for C :: "'c comp" (infixr "\\<^sub>C" 55) and src\<^sub>C :: "'c \ 'c" and trg\<^sub>C :: "'c \ 'c" and D :: "'d comp" (infixr "\\<^sub>D" 55) and src\<^sub>D :: "'d \ 'd" and trg\<^sub>D :: "'d \ 'd" and F :: "'c \ 'd" + assumes weakly_preserves_src: "\\. C.arr \ \ D.isomorphic (F (src\<^sub>C \)) (src\<^sub>D (F \))" and weakly_preserves_trg: "\\. C.arr \ \ D.isomorphic (F (trg\<^sub>C \)) (trg\<^sub>D (F \))" begin lemma isomorphic_src: assumes "C.obj a" shows "D.isomorphic (src\<^sub>D (F a)) (F a)" using assms weakly_preserves_src [of a] D.isomorphic_symmetric by auto lemma isomorphic_trg: assumes "C.obj a" shows "D.isomorphic (trg\<^sub>D (F a)) (F a)" using assms weakly_preserves_trg [of a] D.isomorphic_symmetric by auto abbreviation (input) hseq\<^sub>C where "hseq\<^sub>C \ \ \ C.arr \ \ C.arr \ \ src\<^sub>C \ = trg\<^sub>C \" abbreviation (input) hseq\<^sub>D where "hseq\<^sub>D \ \ \ D.arr \ \ D.arr \ \ src\<^sub>D \ = trg\<^sub>D \" lemma preserves_hseq: assumes "hseq\<^sub>C \ \" shows "hseq\<^sub>D (F \) (F \)" - proof - - have "src\<^sub>C \ = trg\<^sub>C \" - using assms by auto - hence "D.isomorphic (F (src\<^sub>C \)) (trg\<^sub>D (F \))" - using assms weakly_preserves_trg by auto - moreover have "D.isomorphic (src\<^sub>D (F \)) (F (src\<^sub>C \))" - using assms weakly_preserves_src D.isomorphic_symmetric by blast - ultimately have "D.isomorphic (src\<^sub>D (F \)) (trg\<^sub>D (F \))" - using D.isomorphic_transitive by blast - hence "src\<^sub>D (F \) = trg\<^sub>D (F \)" - using assms D.isomorphic_objects_are_equal by auto - thus ?thesis - using assms by auto - qed + by (metis D.isomorphic_def D.src_src D.src_trg D.vconn_implies_hpar(3) + assms preserves_reflects_arr weakly_preserves_src weakly_preserves_trg) text \ Though \F\ does not preserve objects ``on the nose'', we can recover from it the usual ``object map'', which does. It is slightly confusing at first to get used to the idea that applying the object map of a weak arrow of homs to an object does not give the same thing as applying the underlying functor, but rather only something isomorphic to it. The following defines the object map associated with \F\. \ definition map\<^sub>0 where "map\<^sub>0 a \ src\<^sub>D (F a)" lemma map\<^sub>0_simps [simp]: assumes "C.obj a" shows "D.obj (map\<^sub>0 a)" and "src\<^sub>D (map\<^sub>0 a) = map\<^sub>0 a" and "trg\<^sub>D (map\<^sub>0 a) = map\<^sub>0 a" and "D.dom (map\<^sub>0 a) = map\<^sub>0 a" and "D.cod (map\<^sub>0 a) = map\<^sub>0 a" using assms map\<^sub>0_def by auto lemma preserves_src [simp]: assumes "C.arr \" shows "src\<^sub>D (F \) = map\<^sub>0 (src\<^sub>C \)" using assms by (metis C.src.preserves_arr C.src_src C.trg_src map\<^sub>0_def preserves_hseq) lemma preserves_trg [simp]: assumes "C.arr \" shows "trg\<^sub>D (F \) = map\<^sub>0 (trg\<^sub>C \)" using assms map\<^sub>0_def preserves_hseq C.src_trg C.trg.preserves_arr by presburger lemma preserves_hhom [intro]: assumes "C.arr \" shows "D.in_hhom (F \) (map\<^sub>0 (src\<^sub>C \)) (map\<^sub>0 (trg\<^sub>C \))" using assms by simp text \ We define here the lifting of \F\ to a functor \FF: CC \ DD\. We need this to define the domains and codomains of the compositors. \ definition FF where "FF \ \\\. if C.VV.arr \\ then (F (fst \\), F (snd \\)) else D.VV.null" sublocale FF: "functor" C.VV.comp D.VV.comp FF proof - have 1: "\\\. C.VV.arr \\ \ D.VV.arr (FF \\)" unfolding FF_def using C.VV.arr_char D.VV.arr_char preserves_hseq by simp show "functor C.VV.comp D.VV.comp FF" proof fix \\ show "\ C.VV.arr \\ \ FF \\ = D.VV.null" using FF_def by simp show "C.VV.arr \\ \ D.VV.arr (FF \\)" using 1 by simp assume \\: "C.VV.arr \\" show "D.VV.dom (FF \\) = FF (C.VV.dom \\)" using \\ 1 FF_def C.VV.arr_char D.VV.arr_char C.VV.dom_simp D.VV.dom_simp by simp show "D.VV.cod (FF \\) = FF (C.VV.cod \\)" using \\ 1 FF_def C.VV.arr_char D.VV.arr_char C.VV.cod_simp D.VV.cod_simp by simp next fix \\ \\ assume 2: "C.VV.seq \\ \\" show "FF (C.VV.comp \\ \\) = D.VV.comp (FF \\) (FF \\)" proof - have "FF (C.VV.comp \\ \\) = (F (fst \\) \\<^sub>D F (fst \\), F (snd \\) \\<^sub>D F (snd \\))" using 1 2 FF_def C.VV.comp_char C.VxV.comp_char C.VV.arr_char by (metis (no_types, lifting) C.VV.seq_char C.VxV.seqE fst_conv as_nat_trans.preserves_comp_2 snd_conv) also have "... = D.VV.comp (FF \\) (FF \\)" using 1 2 FF_def D.VV.comp_char D.VxV.comp_char C.VV.arr_char D.VV.arr_char C.VV.seq_char C.VxV.seqE preserves_seq by (simp, meson) finally show ?thesis by simp qed qed qed lemma functor_FF: shows "functor C.VV.comp D.VV.comp FF" .. end subsection "Definition of Pseudofunctors" text \ I don't much like the term "pseudofunctor", which is suggestive of something that is ``not really'' a functor. In the development here we can see that a pseudofunctor is really a \emph{bona fide} functor with respect to vertical composition, which happens to have in addition a weak preservation property with respect to horizontal composition. This weak preservation of horizontal composition is captured by extra structure, the ``compositors'', which are the components of a natural transformation. So ``pseudofunctor'' is really a misnomer; it's an actual functor that has been equipped with additional structure relating to horizontal composition. I would use the term ``bifunctor'' for such a thing, but it seems to not be generally accepted and also tends to conflict with the usage of that term to refer to an ordinary functor of two arguments; which I have called a ``binary functor''. Sadly, there seem to be no other plausible choices of terminology, other than simply ``functor'' (recommended on n-Lab @{url \https://ncatlab.org/nlab/show/pseudofunctor\}), but that is not workable here because we need a name that does not clash with that used for an ordinary functor between categories. \ locale 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 + weak_arrow_of_homs V\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D src\<^sub>D trg\<^sub>D F + FoH\<^sub>C: composite_functor C.VV.comp V\<^sub>C V\<^sub>D \\\\. H\<^sub>C (fst \\) (snd \\)\ F + H\<^sub>DoFF: composite_functor C.VV.comp D.VV.comp V\<^sub>D FF \\\\. H\<^sub>D (fst \\) (snd \\)\ + \: natural_isomorphism C.VV.comp V\<^sub>D H\<^sub>DoFF.map FoH\<^sub>C.map \ 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 \ :: "'c * 'c \ 'd" + assumes assoc_coherence: "\ C.ide f; C.ide g; C.ide h; src\<^sub>C f = trg\<^sub>C g; src\<^sub>C g = trg\<^sub>C h \ \ F \\<^sub>C[f, g, h] \\<^sub>D \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) = \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h]" begin no_notation C.in_hom ("\_ : _ \\<^sub>C _\") no_notation D.in_hom ("\_ : _ \\<^sub>D _\") notation C.in_hhom ("\_ : _ \\<^sub>C _\") notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hhom ("\_ : _ \\<^sub>D _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") notation C.lunit ("\\<^sub>C[_]") notation C.runit ("\\<^sub>C[_]") notation C.lunit' ("\\<^sub>C\<^sup>-\<^sup>1[_]") notation C.runit' ("\\<^sub>C\<^sup>-\<^sup>1[_]") notation C.\' ("\\<^sub>C\<^sup>-\<^sup>1[_, _, _]") notation D.lunit ("\\<^sub>D[_]") notation D.runit ("\\<^sub>D[_]") notation D.lunit' ("\\<^sub>D\<^sup>-\<^sup>1[_]") notation D.runit' ("\\<^sub>D\<^sup>-\<^sup>1[_]") notation D.\' ("\\<^sub>D\<^sup>-\<^sup>1[_, _, _]") lemma weakly_preserves_objects: assumes "C.obj a" shows "D.isomorphic (map\<^sub>0 a) (F a)" using assms weakly_preserves_src [of a] D.isomorphic_symmetric by auto lemma cmp_in_hom [intro]: assumes "C.ide a" and "C.ide b" and "src\<^sub>C a = trg\<^sub>C b" shows "\\ (a, b) : map\<^sub>0 (src\<^sub>C b) \\<^sub>D map\<^sub>0 (trg\<^sub>C a)\" and "\\ (a, b) : F a \\<^sub>D F b \\<^sub>D F (a \\<^sub>C b)\" proof - show "\\ (a, b) : F a \\<^sub>D F b \\<^sub>D F (a \\<^sub>C b)\" - proof - - have "H\<^sub>DoFF.map (a, b) = F a \\<^sub>D F b" - using assms C.VV.ide_char FF_def by auto - moreover have "FoH\<^sub>C.map (a, b) = F (a \\<^sub>C b)" - using assms C.VV.ide_char by simp - ultimately show ?thesis - using assms C.VV.ide_char \.preserves_hom - apply simp - by (metis (no_types, lifting) C.VV.ideI C.VV.ide_in_hom C.VxV.ide_Ide C.ideD(1) - fst_conv snd_conv) - qed - show "\\ (a, b) : map\<^sub>0 (src\<^sub>C b) \\<^sub>D map\<^sub>0 (trg\<^sub>C a)\" - proof - - have "C.hseq a b" - by (simp add: assms(1-3)) - thus ?thesis - by (metis C.src_hcomp C.trg_hcomp D.in_hhom_def D.in_homE D.src_cod D.trg_cod - \\\ (a, b) : F a \\<^sub>D F b \\<^sub>D F (a \\<^sub>C b)\\ preserves_src preserves_trg) - qed + using assms C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def by auto + thus "\\ (a, b) : map\<^sub>0 (src\<^sub>C b) \\<^sub>D map\<^sub>0 (trg\<^sub>C a)\" + using assms D.vconn_implies_hpar by auto qed lemma cmp_simps [simp]: assumes "C.ide f" and "C.ide g" and "src\<^sub>C f = trg\<^sub>C g" shows "D.arr (\ (f, g))" and "src\<^sub>D (\ (f, g)) = src\<^sub>D (F g)" and "trg\<^sub>D (\ (f, g)) = trg\<^sub>D (F f)" and "D.dom (\ (f, g)) = F f \\<^sub>D F g" and "D.cod (\ (f, g)) = F (f \\<^sub>C g)" - proof - - show "D.arr (\ (f, g))" - using assms cmp_in_hom by auto - show "src\<^sub>D (\ (f, g)) = src\<^sub>D (F g)" - using assms cmp_in_hom by auto - show "trg\<^sub>D (\ (f, g)) = trg\<^sub>D (F f)" - using assms cmp_in_hom by auto - show "D.dom (\ (f, g)) = F f \\<^sub>D F g" - using assms cmp_in_hom by blast - show "D.cod (\ (f, g)) = F (f \\<^sub>C g)" - using assms cmp_in_hom by blast - qed + using assms cmp_in_hom by simp_all blast+ lemma cmp_in_hom': assumes "C.arr \" and "C.arr \" and "src\<^sub>C \ = trg\<^sub>C \" shows "\\ (\, \) : map\<^sub>0 (src\<^sub>C \) \\<^sub>D map\<^sub>0 (trg\<^sub>C \)\" and "\\ (\, \) : F (C.dom \) \\<^sub>D F (C.dom \) \\<^sub>D F (C.cod \ \\<^sub>C C.cod \)\" proof - - show 1: "\\ (\, \) : F (C.dom \) \\<^sub>D F (C.dom \) \\<^sub>D F (C.cod \ \\<^sub>C C.cod \)\" - using assms C.VV.arr_char \.preserves_hom FF_def C.VV.dom_simp C.VV.cod_simp - by auto - show "\\ (\, \) : map\<^sub>0 (src\<^sub>C \) \\<^sub>D map\<^sub>0 (trg\<^sub>C \)\" - using assms 1 D.src_dom [of "\ (\, \)"] D.trg_dom [of "\ (\, \)"] - C.VV.dom_simp C.VV.cod_simp D.vconn_implies_hpar(1-2) - by auto + show "\\ (\, \) : F (C.dom \) \\<^sub>D F (C.dom \) \\<^sub>D F (C.cod \ \\<^sub>C C.cod \)\" + using assms C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def by auto + thus "\\ (\, \) : map\<^sub>0 (src\<^sub>C \) \\<^sub>D map\<^sub>0 (trg\<^sub>C \)\" + using assms D.vconn_implies_hpar by auto qed lemma cmp_simps': assumes "C.arr \" and "C.arr \" and "src\<^sub>C \ = trg\<^sub>C \" shows "D.arr (\ (\, \))" and "src\<^sub>D (\ (\, \)) = map\<^sub>0 (src\<^sub>C \)" and "trg\<^sub>D (\ (\, \)) = map\<^sub>0 (trg\<^sub>C \)" and "D.dom (\ (\, \)) = F (C.dom \) \\<^sub>D F (C.dom \)" and "D.cod (\ (\, \)) = F (C.cod \ \\<^sub>C C.cod \)" - proof - - show "D.arr (\ (\, \))" - using assms cmp_in_hom by auto - show "src\<^sub>D (\ (\, \)) = map\<^sub>0 (src\<^sub>C \)" - using assms cmp_in_hom' by auto - show "trg\<^sub>D (\ (\, \)) = map\<^sub>0 (trg\<^sub>C \)" - using assms cmp_in_hom' by auto - show "D.dom (\ (\, \)) = F (C.dom \) \\<^sub>D F (C.dom \)" - using assms cmp_in_hom' by blast - show "D.cod (\ (\, \)) = F (C.cod \ \\<^sub>C C.cod \)" - using assms cmp_in_hom' by blast - qed + using assms cmp_in_hom' by simp_all blast+ lemma cmp_components_are_iso [simp]: assumes "C.ide f" and "C.ide g" and "src\<^sub>C f = trg\<^sub>C g" shows "D.iso (\ (f, g))" using assms C.VV.ide_char C.VV.arr_char by simp lemma weakly_preserves_hcomp: assumes "C.ide f" and "C.ide g" and "src\<^sub>C f = trg\<^sub>C g" shows "D.isomorphic (F f \\<^sub>D F g) (F (f \\<^sub>C g))" using assms D.isomorphic_def by auto end context pseudofunctor begin text \ The following defines the image of the unit isomorphism \\\<^sub>C[a]\ under \F\. We will use \(F a, \[a])\ as an ``alternate unit'', to substitute for \(src\<^sub>D (F a), \\<^sub>D[src\<^sub>D (F a)])\. \ abbreviation (input) \ ("\[_]") where "\[a] \ F \\<^sub>C[a] \\<^sub>D \ (a, a)" lemma \_in_hom [intro]: assumes "C.obj a" shows "\F \\<^sub>C[a] \\<^sub>D \ (a, a) : map\<^sub>0 a \\<^sub>D map\<^sub>0 a\" and "\\[a] : F a \\<^sub>D F a \\<^sub>D F a\" proof (unfold map\<^sub>0_def) show "\F \\<^sub>C[a] \\<^sub>D \ (a, a) : F a \\<^sub>D F a \\<^sub>D F a\" using assms preserves_hom cmp_in_hom by (intro D.comp_in_homI, auto) show "\F \\<^sub>C[a] \\<^sub>D \ (a, a) : src\<^sub>D (F a) \\<^sub>D src\<^sub>D (F a)\" using assms C.VV.arr_char C.VV.dom_simp C.VV.cod_simp by (intro D.vcomp_in_hhom D.seqI, auto) qed lemma \_simps [simp]: assumes "C.obj a" shows "D.arr (\ a)" and "src\<^sub>D \[a] = map\<^sub>0 a" and "trg\<^sub>D \[a] = map\<^sub>0 a" and "D.dom \[a] = F a \\<^sub>D F a" and "D.cod \[a] = F a" - proof - - show "src\<^sub>D \[a] = map\<^sub>0 a" - unfolding map\<^sub>0_def - using assms \_in_hom D.src_cod [of "F a"] - by (metis C.unit_simps(1) C.unit_simps(5) D.arrI D.src_vcomp D.vseq_implies_hpar(1) - as_nat_trans.is_natural_2 preserves_arr) - show "trg\<^sub>D \[a] = map\<^sub>0 a" - unfolding map\<^sub>0_def - using assms \_in_hom D.trg_cod [of "F a"] - by (metis C.obj_def C.unit_simps(1) C.unit_simps(3) D.arrI D.trg_vcomp preserves_hseq) - show "D.arr \[a]" - using assms \_in_hom by auto - show "D.dom \[a] = F a \\<^sub>D F a" - using assms \_in_hom by auto - show "D.cod \[a] = F a" - using assms \_in_hom by auto - qed + using assms \_in_hom by auto lemma iso_\: assumes "C.obj a" shows "D.iso \[a]" - proof - - have "D.iso (\ (a, a))" - using assms by auto - moreover have "D.iso (F \\<^sub>C[a])" - using assms C.iso_unit by simp - moreover have "D.seq (F \\<^sub>C[a]) (\ (a, a))" - using assms C.obj_self_composable(1) C.seq_if_composable by simp - ultimately show ?thesis by auto - qed + using assms C.iso_unit C.obj_self_composable(1) C.seq_if_composable + by (meson C.objE D.isos_compose \_simps(1) cmp_components_are_iso preserves_iso) text \ If \a\ is an object of \C\ and we have an isomorphism \\\ (a, a) : F a \\<^sub>D F a \\<^sub>D F (a \\<^sub>C a)\\, then there is a canonical way to define a compatible isomorphism \\\ a : map\<^sub>0 a \\<^sub>D F a\\. Specifically, we take \\ a\ to be the unique isomorphism \\\ : map\<^sub>0 a \\<^sub>D F a\\ such that \\ \\<^sub>D \\<^sub>D[map\<^sub>0 a] = \[a] \\<^sub>D (\ \\<^sub>D \)\. \ definition unit where "unit a \ THE \. \\ : map\<^sub>0 a \\<^sub>D F a\ \ D.iso \ \ \ \\<^sub>D \\<^sub>D[map\<^sub>0 a] = \[a] \\<^sub>D (\ \\<^sub>D \)" lemma unit_char: assumes "C.obj a" shows "\unit a : map\<^sub>0 a \\<^sub>D F a\" and "D.iso (unit a)" and "unit a \\<^sub>D \\<^sub>D[map\<^sub>0 a] = \[a] \\<^sub>D (unit a \\<^sub>D unit a)" and "\!\. \\ : map\<^sub>0 a \\<^sub>D F a\ \ D.iso \ \ \ \\<^sub>D \\<^sub>D[map\<^sub>0 a] = \[a] \\<^sub>D (\ \\<^sub>D \)" proof - let ?P = "\\. \\ : map\<^sub>0 a \\<^sub>D F a\ \ D.iso \ \ \ \\<^sub>D \\<^sub>D[map\<^sub>0 a] = \[a] \\<^sub>D (\ \\<^sub>D \)" show "\!\. ?P \" proof - have "D.obj (map\<^sub>0 a)" using assms by simp moreover have "D.isomorphic (map\<^sub>0 a) (F a)" unfolding map\<^sub>0_def using assms isomorphic_src by simp ultimately show ?thesis using assms D.unit_unique_upto_unique_iso \.preserves_hom \_in_hom iso_\ by simp qed hence 1: "?P (unit a)" using assms unit_def the1I2 [of ?P ?P] by simp show "\unit a : map\<^sub>0 a \\<^sub>D F a\" using 1 by simp show "D.iso (unit a)" using 1 by simp show "unit a \\<^sub>D \\<^sub>D[map\<^sub>0 a] = \[a] \\<^sub>D (unit a \\<^sub>D unit a)" using 1 by simp qed lemma unit_simps [simp]: assumes "C.obj a" shows "D.arr (unit a)" and "src\<^sub>D (unit a) = map\<^sub>0 a" and "trg\<^sub>D (unit a) = map\<^sub>0 a" and "D.dom (unit a) = map\<^sub>0 a" and "D.cod (unit a) = F a" - proof - - show "D.arr (unit a)" - using assms unit_char(1) by auto - show 1: "D.dom (unit a) = map\<^sub>0 a" - using assms unit_char(1) by auto - show 2: "D.cod (unit a) = F a" - using assms unit_char(1) by auto - show "src\<^sub>D (unit a) = map\<^sub>0 a" - using assms 1 D.src_dom - unfolding map\<^sub>0_def - by (metis C.obj_def D.src_src preserves_arr) - show "trg\<^sub>D (unit a) = map\<^sub>0 a" - unfolding map\<^sub>0_def - using assms 2 unit_char - by (metis 1 D.trg_dom map\<^sub>0_def map\<^sub>0_simps(3)) - qed + using assms unit_char(1) + apply auto + apply (metis D.vconn_implies_hpar(1) map\<^sub>0_simps(2)) + by (metis D.vconn_implies_hpar(2) map\<^sub>0_simps(3)) lemma unit_in_hom [intro]: assumes "C.obj a" shows "\unit a : map\<^sub>0 a \\<^sub>D map\<^sub>0 a\" and "\unit a : map\<^sub>0 a \\<^sub>D F a\" using assms by auto lemma unit_eqI: assumes "C.obj a" and "\\: map\<^sub>0 a \\<^sub>D F a\" and "D.iso \" and "\ \\<^sub>D \\<^sub>D[map\<^sub>0 a] = \ a \\<^sub>D (\ \\<^sub>D \)" shows "\ = unit a" using assms unit_def unit_char the1_equality [of "\\. \\ : map\<^sub>0 a \\<^sub>D F a\ \ D.iso \ \ \ \\<^sub>D \\<^sub>D[map\<^sub>0 a] = \[a] \\<^sub>D (\ \\<^sub>D \)" \] by simp text \ The following defines the unique isomorphism satisfying the characteristic conditions for the left unitor \\\<^sub>D[trg\<^sub>D (F f)]\, but using the ``alternate unit'' \\[trg\<^sub>C f]\ instead of \\\<^sub>D[trg\<^sub>D (F f)]\, which is used to define \\\<^sub>D[trg\<^sub>D (F f)]\. \ definition lF where "lF f \ THE \. \\ : F (trg\<^sub>C f) \\<^sub>D F f \\<^sub>D F f\ \ F (trg\<^sub>C f) \\<^sub>D \ =(\[trg\<^sub>C f] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]" lemma lF_char: assumes "C.ide f" shows "\lF f : F (trg\<^sub>C f) \\<^sub>D F f \\<^sub>D F f\" and "F (trg\<^sub>C f) \\<^sub>D lF f = (\[trg\<^sub>C f] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]" and "\!\. \\ : F (trg\<^sub>C f) \\<^sub>D F f \\<^sub>D F f\ \ F (trg\<^sub>C f) \\<^sub>D \ = (\[trg\<^sub>C f] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]" proof - let ?P = "\\. \\ : F (trg\<^sub>C f) \\<^sub>D F f \\<^sub>D F f\ \ F (trg\<^sub>C f) \\<^sub>D \ = (\[trg\<^sub>C f] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]" show "\!\. ?P \" proof - interpret Df: prebicategory \(\\<^sub>D)\ \(\\<^sub>D)\ \\<^sub>D using D.is_prebicategory by simp interpret S: subcategory \(\\<^sub>D)\ \Df.left (F (trg\<^sub>C f))\ using assms Df.left_hom_is_subcategory by simp interpret Df: left_hom \(\\<^sub>D)\ \(\\<^sub>D)\ \F (trg\<^sub>C f)\ using assms D.weak_unit_char - apply unfold_locales by simp + by unfold_locales simp interpret Df: left_hom_with_unit \(\\<^sub>D)\ \(\\<^sub>D)\ \\<^sub>D \\[trg\<^sub>C f]\ \F (trg\<^sub>C f)\ - proof - show "Df.weak_unit (F (trg\<^sub>C f))" - using assms D.weak_unit_char - by (metis C.ideD(1) C.trg.preserves_reflects_arr C.trg_trg weakly_preserves_trg) - show "\\[trg\<^sub>C f] : F (trg\<^sub>C f) \\<^sub>D F (trg\<^sub>C f) \\<^sub>D F (trg\<^sub>C f)\" - using assms \_in_hom by simp - show "D.iso \[trg\<^sub>C f]" - using assms iso_\ by simp - qed + using assms \_in_hom iso_\ D.weak_unit_char(1) assms weakly_preserves_trg + by unfold_locales auto have "\!\. \\ : Df.L (F f) \\<^sub>S F f\ \ Df.L \ = (\[trg\<^sub>C f] \\<^sub>D F f) \\<^sub>S \\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]" proof - have "Df.left (F (trg\<^sub>C f)) (F f)" using assms weakly_preserves_src D.isomorphic_def D.hseq_char D.hseq_char' Df.left_def by fastforce thus ?thesis using assms Df.lunit_char(3) S.ide_char S.arr_char by simp qed moreover have "Df.L (F f) = F (trg\<^sub>C f) \\<^sub>D F f" using assms by (simp add: Df.H\<^sub>L_def) moreover have "\\. Df.L \ = F (trg\<^sub>C f) \\<^sub>D \" using Df.H\<^sub>L_def by simp moreover have "(\[trg\<^sub>C f] \\<^sub>D F f) \\<^sub>S \\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f] = (\[trg\<^sub>C f] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]" - by (metis (no_types, lifting) D.assoc'_eq_inv_assoc D.ext D.hseqE D.seqE - D.vconn_implies_hpar(1) D.vconn_implies_hpar(3) Df.characteristic_iso(4) - Df.equivalence_functor_L Df.iso_unit(2) S.comp_simp S.ext S.ide_char - S.in_hom_char S.iso_is_arr S.null_char calculation(1) category.ide_cod - category.in_homE equivalence_functor_def) + by (metis (no_types, lifting) D.arrI D.ext D.hseqI' D.hseq_char' D.seqE + D.seq_if_composable D.vconn_implies_hpar(1) D.vconn_implies_hpar(2-3) + D.vconn_implies_hpar(4) Df.\_in_hom Df.arr_\ S.comp_char S.in_hom_char + calculation(1,3)) moreover have "\\. \\ : F (trg\<^sub>C f) \\<^sub>D F f \\<^sub>D F f\ \ \\ : F (trg\<^sub>C f) \\<^sub>D F f \\<^sub>S F f\" using assms S.in_hom_char S.arr_char - by (metis D.in_homE Df.hom_connected(2) Df.left_def calculation(1) calculation(2)) + by (metis D.in_homE Df.hom_connected(2) Df.left_def calculation(1-2)) ultimately show ?thesis by simp qed hence 1: "?P (lF f)" using lF_def the1I2 [of ?P ?P] by simp show "\lF f : F (trg\<^sub>C f) \\<^sub>D F f \\<^sub>D F f\" using 1 by simp show "F (trg\<^sub>C f) \\<^sub>D lF f = (\[trg\<^sub>C f] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]" using 1 by simp qed lemma lF_simps [simp]: assumes "C.ide f" shows "D.arr (lF f)" and "src\<^sub>D (lF f) = map\<^sub>0 (src\<^sub>C f)" and "trg\<^sub>D (lF f) = map\<^sub>0 (trg\<^sub>C f)" and "D.dom (lF f) = F (trg\<^sub>C f) \\<^sub>D F f" and "D.cod (lF f) = F f" using assms lF_char(1) apply auto[5] unfolding map\<^sub>0_def using assms apply (metis C.ideD(1) D.vconn_implies_hpar(1,3) map\<^sub>0_def preserves_src) by (metis C.ideD(1) C.src_trg C.trg.preserves_arr D.in_homE D.trg_cod preserves_src preserves_trg) text \ \sloppypar The next two lemmas generalize the eponymous results from @{theory MonoidalCategory.MonoidalFunctor}. See the proofs of those results for diagrams. \ lemma lunit_coherence1: assumes "C.ide f" shows "\\<^sub>D[F f] \\<^sub>D D.inv (unit (trg\<^sub>C f) \\<^sub>D F f) = lF f" proof - let ?b = "trg\<^sub>C f" have 1: "trg\<^sub>D (F f) = map\<^sub>0 ?b" using assms by simp have "lF f \\<^sub>D (unit ?b \\<^sub>D F f) = \\<^sub>D[F f]" proof - have "D.par (lF f \\<^sub>D (unit ?b \\<^sub>D F f)) \\<^sub>D[F f]" using assms 1 D.lunit_in_hom unit_char(1-2) lF_char(1) D.ideD(1) by auto moreover have "map\<^sub>0 ?b \\<^sub>D (lF f \\<^sub>D (unit ?b \\<^sub>D F f)) = map\<^sub>0 ?b \\<^sub>D \\<^sub>D[F f]" proof - have "map\<^sub>0 ?b \\<^sub>D (lF f \\<^sub>D (unit ?b \\<^sub>D F f)) = (map\<^sub>0 ?b \\<^sub>D lF f) \\<^sub>D (map\<^sub>0 ?b \\<^sub>D unit ?b \\<^sub>D F f)" using assms D.objE [of "map\<^sub>0 (trg\<^sub>C f)"] D.whisker_left [of "map\<^sub>0 ?b" "lF f" "unit ?b \\<^sub>D F f"] by auto also have "... = (map\<^sub>0 ?b \\<^sub>D lF f) \\<^sub>D (D.inv (unit ?b) \\<^sub>D F ?b \\<^sub>D F f) \\<^sub>D (unit ?b \\<^sub>D unit ?b \\<^sub>D F f)" proof - have "(D.inv (unit ?b) \\<^sub>D F ?b \\<^sub>D F f) \\<^sub>D (unit ?b \\<^sub>D unit ?b \\<^sub>D F f) = D.inv (unit ?b) \\<^sub>D unit ?b \\<^sub>D F ?b \\<^sub>D unit ?b \\<^sub>D F f \\<^sub>D F f" using assms unit_char(1-2) D.interchange [of "F ?b" "unit ?b" "F f" "F f"] D.interchange [of "D.inv (unit ?b)" "unit ?b" "F ?b \\<^sub>D F f" "unit ?b \\<^sub>D F f"] by simp also have "... = map\<^sub>0 ?b \\<^sub>D unit ?b \\<^sub>D F f" using assms unit_char(1-2) [of ?b] D.comp_arr_dom D.comp_cod_arr D.comp_inv_arr by (simp add: D.inv_is_inverse) finally show ?thesis by simp qed also have "... = (D.inv (unit ?b) \\<^sub>D F f) \\<^sub>D (F ?b \\<^sub>D lF f) \\<^sub>D (unit ?b \\<^sub>D unit ?b \\<^sub>D F f)" proof - have "(map\<^sub>0 ?b \\<^sub>D lF f) \\<^sub>D (D.inv (unit ?b) \\<^sub>D F ?b \\<^sub>D F f) = (D.inv (unit ?b) \\<^sub>D F f) \\<^sub>D (F ?b \\<^sub>D lF f)" proof - have "(map\<^sub>0 ?b \\<^sub>D lF f) \\<^sub>D (D.inv (unit ?b) \\<^sub>D F ?b \\<^sub>D F f) = D.inv (unit ?b) \\<^sub>D lF f" using assms unit_char(1-2) lF_char(1) D.comp_arr_dom D.comp_cod_arr D.interchange [of "map\<^sub>0 ?b" "D.inv (unit ?b)" "lF f" "F ?b \\<^sub>D F f"] by simp also have "... = D.inv (unit ?b) \\<^sub>D F ?b \\<^sub>D F f \\<^sub>D lF f" using assms unit_char(1-2) lF_char(1) D.comp_arr_dom D.comp_cod_arr D.inv_in_hom by auto also have "... = (D.inv (unit ?b) \\<^sub>D F f) \\<^sub>D (F ?b \\<^sub>D lF f)" using assms unit_char(1-2) lF_char(1) D.inv_in_hom D.interchange [of "D.inv (unit ?b)" "F ?b" "F f" "lF f"] by simp finally show ?thesis by simp qed thus ?thesis using assms unit_char(1-2) D.inv_in_hom D.comp_assoc by metis qed also have "... = (D.inv (unit ?b) \\<^sub>D F f) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f] \\<^sub>D (unit ?b \\<^sub>D unit ?b \\<^sub>D F f)" using assms unit_char(1-2) lF_char(2) D.comp_assoc by auto also have "... = ((D.inv (unit ?b) \\<^sub>D F f) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D ((unit ?b \\<^sub>D unit ?b) \\<^sub>D F f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 ?b, map\<^sub>0 ?b, F f]" using assms unit_char(1-2) D.assoc'_naturality [of "unit ?b" "unit ?b" "F f"] D.comp_assoc by (simp add: \trg\<^sub>D (F f) = map\<^sub>0 (trg\<^sub>C f)\) also have "... = (\\<^sub>D[map\<^sub>0 ?b] \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 ?b, map\<^sub>0 ?b, F f]" proof - have "((D.inv (unit ?b) \\<^sub>D F f) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D ((unit ?b \\<^sub>D unit ?b) \\<^sub>D F f)) = \\<^sub>D[map\<^sub>0 ?b] \\<^sub>D F f" proof - have "((D.inv (unit ?b) \\<^sub>D F f) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D ((unit ?b \\<^sub>D unit ?b) \\<^sub>D F f)) = D.inv (unit ?b) \\<^sub>D unit ?b \\<^sub>D \\<^sub>D[map\<^sub>0 ?b] \\<^sub>D F f" using assms 1 D.unit_in_hom D.whisker_right [of "F f"] unit_char(2-3) D.invert_side_of_triangle(1) by (metis C.ideD(1) C.obj_trg D.seqI' map\<^sub>0_simps(1) unit_in_hom(2) preserves_ide) also have "... = \\<^sub>D[map\<^sub>0 ?b] \\<^sub>D F f" proof - have "(D.inv (unit (trg\<^sub>C f)) \\<^sub>D unit (trg\<^sub>C f)) \\<^sub>D \\<^sub>D[map\<^sub>0 ?b] = \\<^sub>D[map\<^sub>0 ?b]" by (simp add: D.comp_cod_arr D.comp_inv_arr D.inv_is_inverse unit_char(2) assms) thus ?thesis by (simp add: D.comp_assoc) qed finally show ?thesis by blast qed thus ?thesis by simp qed also have "... = map\<^sub>0 ?b \\<^sub>D \\<^sub>D[F f]" using assms D.lunit_char [of "F f"] \trg\<^sub>D (F f) = map\<^sub>0 ?b\ by simp finally show ?thesis by blast qed ultimately show ?thesis using assms D.L.is_faithful by (metis D.trg_cod D.trg_vcomp D.vseq_implies_hpar(2) lF_simps(3)) qed thus ?thesis using assms 1 unit_char(1-2) C.ideD(1) C.obj_trg D.inverse_arrows_hcomp(1) D.invert_side_of_triangle(2) D.lunit_simps(1) unit_simps(2) preserves_ide D.iso_hcomp as_nat_iso.components_are_iso by metis qed lemma lunit_coherence2: assumes "C.ide f" shows "lF f = F \\<^sub>C[f] \\<^sub>D \ (trg\<^sub>C f, f)" proof - let ?b = "trg\<^sub>C f" have "D.par (F \\<^sub>C[f] \\<^sub>D \ (?b, f)) (lF f)" - using assms cmp_in_hom [of ?b f] lF_simps by auto + using assms cmp_simps'(1) cmp_simps(4-5) by force moreover have "F ?b \\<^sub>D F \\<^sub>C[f] \\<^sub>D \ (?b, f) = F ?b \\<^sub>D lF f" proof - have "F ?b \\<^sub>D F \\<^sub>C[f] \\<^sub>D \ (?b, f) = (F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D (F ?b \\<^sub>D \ (?b, f))" using assms cmp_in_hom D.whisker_left [of "F ?b" "F \\<^sub>C[f]" "\ (?b, f)"] by (simp add: calculation) also have "... = F ?b \\<^sub>D lF f" proof - have "(F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D (F ?b \\<^sub>D \ (?b, f)) = (F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D F \\<^sub>C[?b, ?b, f] \\<^sub>D \ (?b \\<^sub>C ?b, f) \\<^sub>D (\ (?b, ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]" proof - have 1: "D.seq (F \\<^sub>C[trg\<^sub>C f, trg\<^sub>C f, f]) (\ (trg\<^sub>C f \\<^sub>C trg\<^sub>C f, f) \\<^sub>D (\ (trg\<^sub>C f, trg\<^sub>C f) \\<^sub>D F f))" using assms by fastforce hence 2: "D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D F \\<^sub>C[?b, ?b, f] \\<^sub>D \ (?b \\<^sub>C ?b, f) \\<^sub>D (\ (?b, ?b) \\<^sub>D F f) = (F ?b \\<^sub>D \ (?b, f)) \\<^sub>D \\<^sub>D[F ?b, F ?b, F f]" using assms cmp_in_hom assoc_coherence cmp_components_are_iso D.invert_side_of_triangle(1) [of "F \\<^sub>C[?b, ?b, f] \\<^sub>D \ (?b \\<^sub>C ?b, f) \\<^sub>D (\ (?b, ?b) \\<^sub>D F f)" "\ (?b, ?b \\<^sub>C f)" "(F ?b \\<^sub>D \ (?b, f)) \\<^sub>D \\<^sub>D[F ?b, F ?b, F f]"] C.ideD(1) C.ide_hcomp C.trg_hcomp C.trg_trg C.src_trg C.trg.preserves_ide by metis hence "F ?b \\<^sub>D \ (?b, f) = (D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D F \\<^sub>C[?b, ?b, f] \\<^sub>D \ (?b \\<^sub>C ?b, f) \\<^sub>D (\ (?b, ?b) \\<^sub>D F f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]" proof - have "D.seq (D.inv (\ (trg\<^sub>C f, trg\<^sub>C f \\<^sub>C f))) (F \\<^sub>C[trg\<^sub>C f, trg\<^sub>C f, f] \\<^sub>D \ (trg\<^sub>C f \\<^sub>C trg\<^sub>C f, f) \\<^sub>D (\ (trg\<^sub>C f, trg\<^sub>C f) \\<^sub>D F f))" using assms 1 D.hseq_char by auto moreover have "(F (trg\<^sub>C f) \\<^sub>D \ (trg\<^sub>C f, f)) \\<^sub>D \\<^sub>D[F (trg\<^sub>C f), F (trg\<^sub>C f), F f] = D.inv (\ (trg\<^sub>C f, trg\<^sub>C f \\<^sub>C f)) \\<^sub>D F \\<^sub>C[trg\<^sub>C f, trg\<^sub>C f, f] \\<^sub>D \ (trg\<^sub>C f \\<^sub>C trg\<^sub>C f, f) \\<^sub>D (\ (trg\<^sub>C f, trg\<^sub>C f) \\<^sub>D F f)" using assms 2 by simp ultimately show ?thesis using assms D.invert_side_of_triangle(2) [of "D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D F \\<^sub>C[?b, ?b, f] \\<^sub>D \ (?b \\<^sub>C ?b, f) \\<^sub>D (\ (?b, ?b) \\<^sub>D F f)" "F ?b \\<^sub>D \ (?b, f)" "\\<^sub>D[F ?b, F ?b, F f]"] by fastforce qed thus ?thesis using D.comp_assoc by simp qed also have "... = (F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D (D.inv (F (?b \\<^sub>C \\<^sub>C[f])) \\<^sub>D F (\\<^sub>C[?b] \\<^sub>C f)) \\<^sub>D \ (?b \\<^sub>C ?b, f) \\<^sub>D (\ (?b, ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]" proof - have 1: "F (?b \\<^sub>C \\<^sub>C[f]) = F (\\<^sub>C[?b] \\<^sub>C f) \\<^sub>D D.inv (F \\<^sub>C[?b, ?b, f])" using assms C.lunit_char(1-2) C.unit_in_hom preserves_inv by auto have "F \\<^sub>C[?b, ?b, f] = D.inv (F (?b \\<^sub>C \\<^sub>C[f])) \\<^sub>D F (\\<^sub>C[?b] \\<^sub>C f)" proof - have "F \\<^sub>C[?b, ?b, f] \\<^sub>D D.inv (F (\\<^sub>C[?b] \\<^sub>C f)) = D.inv (F (\\<^sub>C[?b] \\<^sub>C f) \\<^sub>D D.inv (F \\<^sub>C[?b, ?b, f]))" using assms by (simp add: C.iso_unit D.inv_comp) thus ?thesis using assms 1 D.invert_side_of_triangle D.iso_inv_iso by (metis C.iso_hcomp C.ideD(1) C.ide_is_iso C.iso_lunit C.iso_unit C.lunit_simps(3) C.obj_trg C.src_trg C.trg.as_nat_iso.components_are_iso C.unit_simps(2) D.arr_inv D.inv_inv preserves_iso) qed thus ?thesis by argo qed also have "... = (F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D D.inv (F (?b \\<^sub>C \\<^sub>C[f])) \\<^sub>D (F (\\<^sub>C[?b] \\<^sub>C f) \\<^sub>D \ (?b \\<^sub>C ?b, f)) \\<^sub>D (\ (?b, ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]" using D.comp_assoc by auto also have "... = (F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D D.inv (F (?b \\<^sub>C \\<^sub>C[f])) \\<^sub>D (\ (?b, f) \\<^sub>D (F \\<^sub>C[?b] \\<^sub>D F f)) \\<^sub>D (\ (?b, ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]" using assms \.naturality [of "(\\<^sub>C[?b], f)"] FF_def C.VV.arr_char C.VV.cod_char C.VV.dom_char by simp also have "... = (F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D D.inv (F (?b \\<^sub>C \\<^sub>C[f])) \\<^sub>D \ (?b, f) \\<^sub>D ((F \\<^sub>C[?b] \\<^sub>D F f)) \\<^sub>D (\ (?b, ?b) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]" using D.comp_assoc by auto also have "... = (F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D D.inv (F (?b \\<^sub>C \\<^sub>C[f])) \\<^sub>D \ (?b, f) \\<^sub>D (\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]" using assms by (simp add: D.comp_assoc D.whisker_right) also have "... = (F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D (D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D D.inv (F (?b \\<^sub>C \\<^sub>C[f])) \\<^sub>D \ (?b, f)) \\<^sub>D (F ?b \\<^sub>D lF f)" - proof - - have "(\ ?b \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f] = F ?b \\<^sub>D lF f" - using assms lF_char by auto - thus ?thesis - using assms D.inv_is_inverse \_in_hom cmp_in_hom D.invert_side_of_triangle(2) - by (simp add: D.comp_assoc) - qed + using D.comp_assoc assms lF_char(2) by presburger also have "... = (F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D D.inv (F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D (F ?b \\<^sub>D lF f)" proof - have "D.inv (F ?b \\<^sub>D F \\<^sub>C[f]) = D.inv (((F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D D.inv (\ (?b, ?b \\<^sub>C f))) \\<^sub>D \ (?b, ?b \\<^sub>C f))" using assms D.comp_inv_arr D.comp_inv_arr' cmp_simps(4) D.comp_arr_dom D.comp_assoc by simp also have "... = D.inv (D.inv (\ (?b, f)) \\<^sub>D F (?b \\<^sub>C \\<^sub>C[f]) \\<^sub>D \ (?b, ?b \\<^sub>C f))" proof - have 1: "\ (?b, f) \\<^sub>D (F ?b \\<^sub>D F \\<^sub>C[f]) = F (?b \\<^sub>C \\<^sub>C[f]) \\<^sub>D \ (?b, ?b \\<^sub>C f)" using assms \.naturality [of "(?b, \\<^sub>C[f])"] FF_def C.VV.arr_char C.VV.cod_char D.VV.null_char C.VV.dom_simp by simp have "(F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D D.inv (\ (?b, ?b \\<^sub>C f)) = D.inv (\ (?b, f)) \\<^sub>D F (?b \\<^sub>C \\<^sub>C[f])" proof - have "D.seq (\ (?b, f)) (F ?b \\<^sub>D F \\<^sub>C[f])" using assms cmp_in_hom(2) [of ?b f] by auto moreover have "D.iso (\ (?b, f)) \ D.iso (\ (?b, ?b \\<^sub>C f))" using assms by simp ultimately show ?thesis using 1 D.invert_opposite_sides_of_square by simp qed thus ?thesis using D.comp_assoc by auto qed also have "... = D.inv (F (?b \\<^sub>C \\<^sub>C[f]) \\<^sub>D \ (?b, ?b \\<^sub>C f)) \\<^sub>D \ (?b, f)" proof - have "D.iso (F (?b \\<^sub>C \\<^sub>C[f]) \\<^sub>D \ (?b, ?b \\<^sub>C f))" using assms D.isos_compose C.VV.arr_char C.iso_lunit C.VV.dom_simp C.VV.cod_simp by simp moreover have "D.iso (D.inv (\ (?b, f)))" using assms by simp moreover have "D.seq (D.inv (\ (?b, f))) (F (?b \\<^sub>C \\<^sub>C[f]) \\<^sub>D \ (?b, ?b \\<^sub>C f))" using assms C.VV.arr_char C.VV.dom_simp C.VV.cod_simp by simp ultimately show ?thesis using assms D.inv_comp by simp qed also have "... = D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D D.inv (F (?b \\<^sub>C \\<^sub>C[f])) \\<^sub>D \ (?b, f)" - proof - - have "D.inv (F (?b \\<^sub>C \\<^sub>C[f]) \\<^sub>D \ (?b, ?b \\<^sub>C f)) = - D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D D.inv (F (?b \\<^sub>C \\<^sub>C[f]))" - using assms D.isos_compose C.VV.arr_char C.iso_lunit D.inv_comp - C.VV.dom_simp C.VV.cod_simp - by simp - thus ?thesis using D.comp_assoc by simp - qed + using D.comp_assoc D.inv_comp assms cmp_simps'(1) cmp_simps(5) by force finally have "D.inv (F ?b \\<^sub>D F \\<^sub>C[f]) = D.inv (\ (?b, ?b \\<^sub>C f)) \\<^sub>D D.inv (F (?b \\<^sub>C \\<^sub>C[f])) \\<^sub>D \ (?b, f)" by blast thus ?thesis by argo qed also have "... = ((F ?b \\<^sub>D F \\<^sub>C[f]) \\<^sub>D D.inv (F ?b \\<^sub>D F \\<^sub>C[f])) \\<^sub>D (F ?b \\<^sub>D lF f)" using D.comp_assoc by simp also have "... = F ?b \\<^sub>D lF f" using assms D.comp_arr_inv' [of "F ?b \\<^sub>D F \\<^sub>C[f]"] D.comp_cod_arr by simp finally show ?thesis by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis using assms D.L.is_faithful by (metis D.in_homI lF_char(2-3) lF_simps(4-5)) qed lemma lunit_coherence: assumes "C.ide f" shows "\\<^sub>D[F f] = F \\<^sub>C[f] \\<^sub>D \ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f)" proof - - have 1: "\\<^sub>D[F f] \\<^sub>D D.inv (unit (trg\<^sub>C f) \\<^sub>D F f) = F \\<^sub>C[f] \\<^sub>D \ (trg\<^sub>C f, f)" - using assms lunit_coherence1 lunit_coherence2 by simp have "\\<^sub>D[F f] = (F \\<^sub>C[f] \\<^sub>D \ (trg\<^sub>C f, f)) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f)" - proof - - have "D.seq (F \\<^sub>C[f]) (\ (trg\<^sub>C f, f))" - using assms cmp_in_hom [of "trg\<^sub>C f" f] C.unit_in_vhom by auto - moreover have "D.iso (D.inv (unit (trg\<^sub>C f) \\<^sub>D F f))" - using assms unit_char unit_char(1-2) - by (simp add: preserves_hseq) - ultimately show ?thesis - using assms 1 unit_char(2) cmp_in_hom D.inv_inv - D.invert_side_of_triangle(2) - [of "F \\<^sub>C[f] \\<^sub>D \ (trg\<^sub>C f, f)" "\\<^sub>D[F f]" "D.inv (unit (trg\<^sub>C f) \\<^sub>D F f)"] - by auto - qed + by (metis C.ideD(1) C.obj_trg D.inv_inv D.invert_side_of_triangle(2) + D.iso_hcomp D.iso_inv_iso as_nat_iso.components_are_iso assms lF_simps(1) + lunit_coherence1 lunit_coherence2 preserves_trg unit_char(2) unit_simps(2)) thus ?thesis - using assms unit_char(1) D.comp_assoc by auto + using assms D.comp_assoc by simp qed text \ We postpone proving the dual version of this result until after we have developed the notion of the ``op bicategory'' in the next section. \ end subsection "Pseudofunctors and Opposite Bicategories" text \ There are three duals to a bicategory: \begin{enumerate} \item ``op'': sources and targets are exchanged; \item ``co'': domains and codomains are exchanged; \item ``co-op'': both sources and targets and domains and codomains are exchanged. \end{enumerate} Here we consider the "op" case. \ locale op_bicategory = B: bicategory V H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B for V :: "'a comp" (infixr "\" 55) and H\<^sub>B :: "'a comp" (infixr "\\<^sub>B" 53) and \\<^sub>B :: "'a \ 'a \ 'a \ 'a" ("\\<^sub>B[_, _, _]") and \\<^sub>B :: "'a \ 'a" ("\\<^sub>B[_]") and src\<^sub>B :: "'a \ 'a" and trg\<^sub>B :: "'a \ 'a" begin abbreviation H (infixr "\" 53) where "H f g \ H\<^sub>B g f" abbreviation \ ("\[_]") where "\ \ \\<^sub>B" abbreviation src where "src \ trg\<^sub>B" abbreviation trg where "trg \ src\<^sub>B" interpretation horizontal_homs V src trg by (unfold_locales, auto) interpretation H: "functor" VV.comp V \\\\. fst \\ \ snd \\\ using VV.arr_char VV.dom_simp VV.cod_simp apply unfold_locales apply (metis (no_types, lifting) B.hseqE B.hseq_char') apply auto[3] using VV.comp_char VV.seq_char VV.arr_char B.VxV.comp_char B.interchange by (metis (no_types, lifting) B.VxV.seqE fst_conv snd_conv) interpretation horizontal_composition V H src trg by (unfold_locales, auto) abbreviation UP where "UP \\\ \ if B.VVV.arr \\\ then (snd (snd \\\), fst (snd \\\), fst \\\) else VVV.null" abbreviation DN where "DN \\\ \ if VVV.arr \\\ then (snd (snd \\\), fst (snd \\\), fst \\\) else B.VVV.null" lemma VVV_arr_char: shows "VVV.arr \\\ \ B.VVV.arr (DN \\\)" using VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char B.VVV.not_arr_null by auto lemma VVV_ide_char: shows "VVV.ide \\\ \ B.VVV.ide (DN \\\)" proof - have "VVV.ide \\\ \ VVV.arr \\\ \ B.VxVxV.ide \\\" using VVV.ide_char by simp also have "... \ B.VVV.arr (DN \\\) \ B.VxVxV.ide (DN \\\)" using VVV_arr_char B.VxVxV.ide_char by auto also have "... \ B.VVV.ide (DN \\\)" using B.VVV.ide_char [of "DN \\\"] by blast finally show ?thesis by fast qed lemma VVV_dom_char: shows "VVV.dom \\\ = UP (B.VVV.dom (DN \\\))" proof (cases "VVV.arr \\\") show "\ VVV.arr \\\ \ VVV.dom \\\ = UP (B.VVV.dom (DN \\\))" using VVV.dom_def VVV.has_domain_iff_arr VVV_arr_char B.VVV.dom_null by auto show "VVV.arr \\\ \ VVV.dom \\\ = UP (B.VVV.dom (DN \\\))" proof - assume \\\: "VVV.arr \\\" have "VVV.dom \\\ = (B.dom (fst \\\), B.dom (fst (snd \\\)), B.dom (snd (snd \\\)))" using \\\ VVV.dom_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char by simp also have "... = UP (B.dom (snd (snd \\\)), B.dom (fst (snd \\\)), B.dom (fst \\\))" - using \\\ VVV_arr_char B.VV.arr_char B.VVV.arr_char B.arr_dom_iff_arr src_dom - trg_dom - apply simp - by (metis (no_types, lifting) VV.arrE VVV.arrE) + by (metis (no_types, lifting) B.VV.arrI B.VVV.arr_char B.arr_dom VV.arrE VVV.arrE + \\\ fst_conv snd_conv src_dom trg_dom) also have "... = UP (B.VVV.dom (DN \\\))" using \\\ B.VVV.dom_char B.VVV.arr_char B.VV.arr_char VVV.arr_char VV.arr_char by simp finally show ?thesis by blast qed qed lemma VVV_cod_char: shows "VVV.cod \\\ = UP (B.VVV.cod (DN \\\))" proof (cases "VVV.arr \\\") show "\ VVV.arr \\\ \ VVV.cod \\\ = UP (B.VVV.cod (DN \\\))" using VVV.cod_def VVV.has_codomain_iff_arr VVV_arr_char B.VVV.cod_null by auto show "VVV.arr \\\ \ VVV.cod \\\ = UP (B.VVV.cod (DN \\\))" proof - assume \\\: "VVV.arr \\\" have "VVV.cod \\\ = (B.cod (fst \\\), B.cod (fst (snd \\\)), B.cod (snd (snd \\\)))" using \\\ VVV.cod_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char by simp also have "... = UP (B.cod (snd (snd \\\)), B.cod (fst (snd \\\)), B.cod (fst \\\))" - using \\\ VVV_arr_char B.VV.arr_char B.VVV.arr_char - apply simp - by (metis (no_types, lifting) B.arr_cod_iff_arr VV.arrE VVV.arrE) + by (metis (no_types, lifting) B.VV.arrI B.VVV.arr_char B.arr_cod VV.arrE VVV.arrE + \\\ fst_conv snd_conv src_cod trg_cod) also have "... = UP (B.VVV.cod (DN \\\))" using \\\ B.VVV.cod_char B.VVV.arr_char B.VV.arr_char VVV.arr_char VV.arr_char by simp finally show ?thesis by blast qed qed lemma HoHV_char: shows "HoHV \\\ = B.HoVH (DN \\\)" using HoHV_def B.HoVH_def VVV_arr_char by simp lemma HoVH_char: shows "HoVH \\\ = B.HoHV (DN \\\)" using HoVH_def B.HoHV_def VVV_arr_char by simp definition \ ("\[_, _, _]") where "\[\, \, \] \ B.\' (DN (\, \, \))" interpretation natural_isomorphism VVV.comp \(\)\ HoHV HoVH \\\\\. \[fst \\\, fst (snd \\\), snd (snd \\\)]\ proof fix \\\ show "\ VVV.arr \\\ \ \[fst \\\, fst (snd \\\), snd (snd \\\)] = B.null" using VVV.arr_char B.VVV.arr_char \_def B.\'.is_extensional by auto assume \\\: "VVV.arr \\\" show "B.dom \[fst \\\, fst (snd \\\), snd (snd \\\)] = HoHV (VVV.dom \\\)" using \\\ \_def HoHV_def B.\'.preserves_dom VVV.arr_char VVV.dom_char VV.arr_char B.HoVH_def B.VVV.arr_char B.VV.arr_char B.VVV.dom_char by auto show "B.cod \[fst \\\, fst (snd \\\), snd (snd \\\)] = HoVH (VVV.cod \\\)" using \\\ \_def HoVH_def B.\'.preserves_cod VVV.arr_char VVV.cod_char VV.arr_char B.HoHV_def B.VVV.arr_char B.VV.arr_char B.VVV.cod_char by auto show "HoVH \\\ \ \[fst (VVV.dom \\\), fst (snd (VVV.dom \\\)), snd (snd (VVV.dom \\\))] = \[fst \\\, fst (snd \\\), snd (snd \\\)]" proof - have "HoVH \\\ \ \[fst (VVV.dom \\\), fst (snd (VVV.dom \\\)), snd (snd (VVV.dom \\\))] = HoVH \\\ \ B.\' (B.VVV.dom (snd (snd \\\), fst (snd \\\), fst \\\))" unfolding \_def using \\\ VVV.arr_char VV.arr_char VVV.dom_char B.VVV.arr_char B.VVV.dom_char by auto also have "... = B.\' (snd (snd \\\), fst (snd \\\), fst \\\)" using B.\'.is_natural_1 VVV_arr_char \\\ HoVH_char by presburger also have "... = \[fst \\\, fst (snd \\\), snd (snd \\\)]" using \\\ \_def by simp finally show ?thesis by blast qed show "\[fst (VVV.cod \\\), fst (snd (VVV.cod \\\)), snd (snd (VVV.cod \\\))] \ HoHV \\\ = \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))" proof - have "\[fst (VVV.cod \\\), fst (snd (VVV.cod \\\)), snd (snd (VVV.cod \\\))] \ HoHV \\\ = B.\' (B.VVV.cod (snd (snd \\\), fst (snd \\\), fst \\\)) \ HoHV \\\" unfolding \_def using \\\ VVV.arr_char VV.arr_char VVV.cod_char B.VVV.arr_char B.VVV.cod_char by auto also have "... = B.\' (snd (snd \\\), fst (snd \\\), fst \\\)" using B.\'.is_natural_2 VVV_arr_char \\\ HoHV_char by presburger also have "... = \[fst \\\, fst (snd \\\), snd (snd \\\)]" using \\\ \_def by simp finally show ?thesis by blast qed next fix \\\ assume \\\: "VVV.ide \\\" show "B.iso \[fst \\\, fst (snd \\\), snd (snd \\\)]" proof - have "B.VVV.ide (DN \\\)" using \\\ VVV_ide_char by blast thus ?thesis using \\\ \_def B.\'.components_are_iso by force qed qed sublocale bicategory V H \ \ src trg proof show "\a. obj a \ \\ a : H a a \\<^sub>B a\" using obj_def objE B.obj_def B.objE B.unit_in_hom by meson show "\a. obj a \ B.iso (\ a)" using obj_def objE B.obj_def B.objE B.iso_unit by meson show "\f g h k. \ B.ide f; B.ide g; B.ide h; B.ide k; src f = trg g; src g = trg h; src h = trg k \ \ (f \ \[g, h, k]) \ \[f, g \ h, k] \ (\[f, g, h] \ k) = \[f, g, h \ k] \ \[f \ g, h, k]" unfolding \_def using B.\'_def B.comp_assoc B.pentagon' VVV.arr_char VV.arr_char by simp qed proposition is_bicategory: shows "bicategory V H \ \ src trg" .. lemma assoc_ide_simp: assumes "B.ide f" and "B.ide g" and "B.ide h" and "src f = trg g" and "src g = trg h" shows "\[f, g, h] = B.\' h g f" using assms \_def B.\'_def by fastforce lemma lunit_ide_simp: assumes "B.ide f" shows "lunit f = B.runit f" proof (intro B.runit_eqI) show "B.ide f" by fact show "\lunit f : H (trg f) f \\<^sub>B f\" using assms by simp show "trg f \ lunit f = (\[trg f] \ f) \ \\<^sub>B[f, trg f, trg f]" proof - have "trg f \ lunit f = (\[trg f] \ f) \ \' (trg f) (trg f) f" using assms lunit_char(1-2) [of f] by simp moreover have "\' (trg f) (trg f) f = \\<^sub>B[f, trg f, trg f]" proof (unfold \'_def) have 1: "VVV.ide (trg f, trg f, f)" using assms VVV.ide_char VVV.arr_char VV.arr_char by simp have "\' (trg f, trg f, f) = B.inv \[trg f, trg f, f]" using 1 B.\'.inverts_components by simp also have "... = B.inv (B.\' (f, trg f, trg f))" unfolding \_def using 1 by simp also have "... = \\<^sub>B[f, trg f, trg f]" using 1 B.VVV.ide_char B.VVV.arr_char B.VV.arr_char VVV.ide_char VVV.arr_char VV.arr_char B.\'.inverts_components B.\_def by force finally show "\' (trg f, trg f, f) = \\<^sub>B[f, trg f, trg f]" by blast qed ultimately show ?thesis by simp qed qed lemma runit_ide_simp: assumes "B.ide f" shows "runit f = B.lunit f" using assms runit_char(1-2) [of f] B.\'_def assoc_ide_simp by (intro B.lunit_eqI) auto end context pseudofunctor begin interpretation C': op_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C .. interpretation D': op_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D .. notation C'.H (infixr "\\<^sub>C\<^sup>o\<^sup>p" 53) notation D'.H (infixr "\\<^sub>D\<^sup>o\<^sup>p" 53) interpretation F': weak_arrow_of_homs V\<^sub>C C'.src C'.trg V\<^sub>D D'.src D'.trg F apply unfold_locales using weakly_preserves_src weakly_preserves_trg by simp_all interpretation H\<^sub>D'oFF: composite_functor C'.VV.comp D'.VV.comp V\<^sub>D F'.FF \\\\. fst \\ \\<^sub>D\<^sup>o\<^sup>p snd \\\ .. interpretation FoH\<^sub>C': composite_functor C'.VV.comp V\<^sub>C V\<^sub>D \\\\. fst \\ \\<^sub>C\<^sup>o\<^sup>p snd \\\ F .. interpretation \': natural_isomorphism C'.VV.comp V\<^sub>D H\<^sub>D'oFF.map FoH\<^sub>C'.map \\f. \ (snd f, fst f)\ using C.VV.arr_char C'.VV.arr_char C'.VV.ide_char C.VV.ide_char FF_def F'.FF_def \.is_extensional \.is_natural_1 \.is_natural_2 C.VV.dom_simp C.VV.cod_simp C'.VV.dom_simp C'.VV.cod_simp - apply unfold_locales by auto + by unfold_locales auto interpretation F': pseudofunctor V\<^sub>C C'.H C'.\ \\<^sub>C C'.src C'.trg V\<^sub>D D'.H D'.\ \\<^sub>D D'.src D'.trg F \\f. \ (snd f, fst f)\ proof fix f g h assume f: "C.ide f" and g: "C.ide g" and h: "C.ide h" and fg: "C'.src f = C'.trg g" and gh: "C'.src g = C'.trg h" show "F (C'.\ f g h) \\<^sub>D \ (snd (f \\<^sub>C\<^sup>o\<^sup>p g, h), fst (f \\<^sub>C\<^sup>o\<^sup>p g, h)) \\<^sub>D (\ (snd (f, g), fst (f, g)) \\<^sub>D\<^sup>o\<^sup>p F h) = \ (snd (f, g \\<^sub>C\<^sup>o\<^sup>p h), fst (f, g \\<^sub>C\<^sup>o\<^sup>p h)) \\<^sub>D (F f \\<^sub>D\<^sup>o\<^sup>p \ (snd (g, h), fst (g, h))) \\<^sub>D D'.\ (F f) (F g) (F h)" using f g h fg gh C.VV.in_hom_char FF_def C.VV.arr_char C.VV.dom_simp C.VV.cod_simp C'.assoc_ide_simp D'.assoc_ide_simp preserves_inv assoc_coherence D.invert_opposite_sides_of_square [of "F (\\<^sub>C h g f)" "\ (g \\<^sub>C\<^sup>o\<^sup>p h, f) \\<^sub>D (F f \\<^sub>D\<^sup>o\<^sup>p \ (h, g))" "\ (h, f \\<^sub>C\<^sup>o\<^sup>p g) \\<^sub>D (\ (g, f) \\<^sub>D\<^sup>o\<^sup>p F h)" "\\<^sub>D (F h) (F g) (F f)"] D.comp_assoc by auto qed lemma induces_pseudofunctor_between_opposites: shows "pseudofunctor (\\<^sub>C) (\\<^sub>C\<^sup>o\<^sup>p) C'.\ \\<^sub>C C'.src C'.trg (\\<^sub>D) (\\<^sub>D\<^sup>o\<^sup>p) D'.\ \\<^sub>D D'.src D'.trg F (\f. \ (snd f, fst f))" .. text \ It is now easy to dualize the coherence condition for \F\ with respect to left unitors to obtain the corresponding condition for right unitors. \ lemma runit_coherence: assumes "C.ide f" shows "\\<^sub>D[F f] = F \\<^sub>C[f] \\<^sub>D \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f))" proof - have "C'.lunit f = \\<^sub>C[f]" using assms C'.lunit_ide_simp by simp moreover have "D'.lunit (F f) = \\<^sub>D[F f]" using assms D'.lunit_ide_simp by simp moreover have "F'.unit (src\<^sub>C f) = unit (src\<^sub>C f)" using assms F'.unit_char F'.preserves_trg - by (intro unit_eqI, auto) + by (intro unit_eqI) auto moreover have "D'.lunit (F f) = F (C'.lunit f) \\<^sub>D \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D F'.unit (src\<^sub>C f))" using assms F'.lunit_coherence by simp ultimately show ?thesis by simp qed end subsection "Preservation Properties" text \ The objective of this section is to establish explicit formulas for the result of applying a pseudofunctor to expressions of various forms. \ context pseudofunctor begin lemma preserves_lunit: assumes "C.ide f" shows "F \\<^sub>C[f] = \\<^sub>D[F f] \\<^sub>D (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F f) \\<^sub>D D.inv (\ (trg\<^sub>C f, f))" and "F \\<^sub>C\<^sup>-\<^sup>1[f] = \ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]" proof - - have "\\<^sub>D[F f] \\<^sub>D D.inv (\ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f)) = F \\<^sub>C[f]" + show 1: "F \\<^sub>C[f] = \\<^sub>D[F f] \\<^sub>D (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F f) \\<^sub>D D.inv (\ (trg\<^sub>C f, f))" proof - - have "D.arr \\<^sub>D[F f]" - using assms by simp - moreover have "\\<^sub>D[F f] = F \\<^sub>C[f] \\<^sub>D \ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f)" - using assms lunit_coherence by simp - moreover have "D.iso (\ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f))" - using assms unit_char cmp_components_are_iso - by (intro D.isos_compose D.seqI) auto - ultimately show ?thesis - using assms D.invert_side_of_triangle(2) by metis + have "\\<^sub>D[F f] \\<^sub>D D.inv (\ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f)) = F \\<^sub>C[f]" + proof - + have "D.arr \\<^sub>D[F f]" + using assms by simp + moreover have "\\<^sub>D[F f] = F \\<^sub>C[f] \\<^sub>D \ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f)" + using assms lunit_coherence by simp + moreover have "D.iso (\ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f))" + using assms unit_char cmp_components_are_iso + by (intro D.isos_compose D.seqI) auto + ultimately show ?thesis + using assms D.invert_side_of_triangle(2) by metis + qed + moreover have "D.inv (\ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f)) = + (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F f) \\<^sub>D D.inv (\ (trg\<^sub>C f, f))" + using assms C.VV.arr_char unit_char FF_def D.inv_comp C.VV.dom_simp by simp + ultimately show ?thesis by simp qed - moreover have "D.inv (\ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f)) = - (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F f) \\<^sub>D D.inv (\ (trg\<^sub>C f, f))" - using assms C.VV.arr_char unit_char FF_def D.inv_comp C.VV.dom_simp by simp - ultimately show "F \\<^sub>C[f] = - \\<^sub>D[F f] \\<^sub>D (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F f) \\<^sub>D D.inv (\ (trg\<^sub>C f, f))" - by simp - hence "F \\<^sub>C\<^sup>-\<^sup>1[f] = - D.inv (\\<^sub>D[F f] \\<^sub>D (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F f) \\<^sub>D D.inv (\ (trg\<^sub>C f, f)))" - using assms preserves_inv by simp - also have "... = \ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]" - using assms unit_char D.comp_assoc D.isos_compose - by (auto simp add: D.inv_comp) - finally show "F \\<^sub>C\<^sup>-\<^sup>1[f] = \ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]" - by simp + show "F \\<^sub>C\<^sup>-\<^sup>1[f] = \ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]" + proof - + have "F \\<^sub>C\<^sup>-\<^sup>1[f] = + D.inv (\\<^sub>D[F f] \\<^sub>D (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F f) \\<^sub>D D.inv (\ (trg\<^sub>C f, f)))" + using assms 1 preserves_inv by simp + also have "... = \ (trg\<^sub>C f, f) \\<^sub>D (unit (trg\<^sub>C f) \\<^sub>D F f) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]" + using assms unit_char D.comp_assoc D.isos_compose + by (auto simp add: D.inv_comp) + finally show ?thesis by simp + qed qed lemma preserves_runit: assumes "C.ide f" shows "F \\<^sub>C[f] = \\<^sub>D[F f] \\<^sub>D (F f \\<^sub>D D.inv (unit (src\<^sub>C f))) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" and "F \\<^sub>C\<^sup>-\<^sup>1[f] = \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]" proof - - have "F \\<^sub>C[f] = \\<^sub>D[F f] \\<^sub>D D.inv (\ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f)))" + show 1: "F \\<^sub>C[f] = \\<^sub>D[F f] \\<^sub>D (F f \\<^sub>D D.inv (unit (src\<^sub>C f))) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" proof - - have "\\<^sub>D[F f] = F \\<^sub>C[f] \\<^sub>D \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f))" - using assms runit_coherence by simp - moreover have "D.iso (\ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f)))" - using assms unit_char D.iso_hcomp FF_def - apply (intro D.isos_compose D.seqI) by auto - moreover have "D.arr \\<^sub>D[F f]" - using assms by simp - ultimately show ?thesis - using assms D.invert_side_of_triangle(2) by metis + have "F \\<^sub>C[f] = \\<^sub>D[F f] \\<^sub>D D.inv (\ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f)))" + proof - + have "\\<^sub>D[F f] = F \\<^sub>C[f] \\<^sub>D \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f))" + using assms runit_coherence by simp + moreover have "D.iso (\ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f)))" + using assms unit_char D.iso_hcomp FF_def + apply (intro D.isos_compose D.seqI) by auto + moreover have "D.arr \\<^sub>D[F f]" + using assms by simp + ultimately show ?thesis + using assms D.invert_side_of_triangle(2) by metis + qed + moreover have "D.inv (\ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f))) = + (F f \\<^sub>D D.inv (unit (src\<^sub>C f))) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" + using assms C.VV.arr_char unit_char D.iso_hcomp FF_def D.inv_comp C.VV.dom_simp + by simp + ultimately show ?thesis by simp qed - moreover have "D.inv (\ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f))) = - (F f \\<^sub>D D.inv (unit (src\<^sub>C f))) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" - using assms C.VV.arr_char unit_char D.iso_hcomp FF_def D.inv_comp C.VV.dom_simp - by simp - ultimately show "F \\<^sub>C[f] = - \\<^sub>D[F f] \\<^sub>D (F f \\<^sub>D D.inv (unit (src\<^sub>C f))) \\<^sub>D D.inv (\ (f, src\<^sub>C f))" - by simp - hence "F \\<^sub>C\<^sup>-\<^sup>1[f] = - D.inv (\\<^sub>D[F f] \\<^sub>D (F f \\<^sub>D D.inv (unit (src\<^sub>C f))) \\<^sub>D D.inv (\ (f, src\<^sub>C f)))" - using assms preserves_inv C.iso_runit by simp - also have "... = \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]" - using assms unit_char D.comp_assoc D.isos_compose - by (auto simp add: D.inv_comp) - finally show "F \\<^sub>C\<^sup>-\<^sup>1[f] = \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]" - by simp + show "F \\<^sub>C\<^sup>-\<^sup>1[f] = \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]" + proof - + have "F \\<^sub>C\<^sup>-\<^sup>1[f] = + D.inv (\\<^sub>D[F f] \\<^sub>D (F f \\<^sub>D D.inv (unit (src\<^sub>C f))) \\<^sub>D D.inv (\ (f, src\<^sub>C f)))" + using assms 1 preserves_inv C.iso_runit by simp + also have "... = \ (f, src\<^sub>C f) \\<^sub>D (F f \\<^sub>D unit (src\<^sub>C f)) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f]" + using assms unit_char D.comp_assoc D.isos_compose + by (auto simp add: D.inv_comp) + finally show ?thesis by simp + qed qed lemma preserves_assoc: assumes "C.ide f" and "C.ide g" and "C.ide h" and "src\<^sub>C f = trg\<^sub>C g" and "src\<^sub>C g = trg\<^sub>C h" shows "F \\<^sub>C[f, g, h] = \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h] \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" and "F \\<^sub>C\<^sup>-\<^sup>1[f, g, h] = \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F h] \\<^sub>D (F f \\<^sub>D D.inv (\ (g, h))) \\<^sub>D D.inv (\ (f, g \\<^sub>C h))" proof - - have "F \\<^sub>C[f, g, h] \\<^sub>D \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) = - \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h]" - using assms assoc_coherence [of f g h] by simp - moreover have "D.seq (\ (f, g \\<^sub>C h)) ((F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h])" - using assms C.VV.arr_char FF_def C.VV.dom_simp C.VV.cod_simp by auto - moreover have 1: "D.iso (\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h))" - using assms C.VV.arr_char FF_def C.VV.dom_simp C.VV.cod_simp by auto - moreover have "D.inv (\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h)) = - (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" - using assms 1 C.VV.arr_char FF_def D.inv_comp C.VV.dom_simp C.VV.cod_simp - by simp - ultimately show 1: "F \\<^sub>C[f, g, h] = - \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h] \\<^sub>D - (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" - using D.invert_side_of_triangle(2) - [of "\ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h]" - "F \\<^sub>C[f, g, h]" "\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h)"] - D.comp_assoc - by simp - hence "F \\<^sub>C\<^sup>-\<^sup>1[f, g, h] = - D.inv (\ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h] \\<^sub>D - (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h)))" - using assms preserves_inv by simp - also have "... = \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F h] \\<^sub>D - (F f \\<^sub>D D.inv (\ (g, h))) \\<^sub>D D.inv (\ (f, g \\<^sub>C h))" + show 1: "F \\<^sub>C[f, g, h] = + \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h] \\<^sub>D + (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" proof - - have "\\ (f, g \\<^sub>C h) : F f \\<^sub>D F (g \\<^sub>C h) \\<^sub>D F (f \\<^sub>C g \\<^sub>C h)\ \ D.iso (\ (f, g \\<^sub>C h))" - using assms by auto - moreover have "\F f \\<^sub>D \ (g, h) : F f \\<^sub>D F g \\<^sub>D F h \\<^sub>D F f \\<^sub>D F (g \\<^sub>C h)\ \ - D.iso (F f \\<^sub>D \ (g, h))" - using assms - by (intro conjI D.hcomp_in_vhom, auto) + have "F \\<^sub>C[f, g, h] \\<^sub>D \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) = + \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h]" + using assms assoc_coherence [of f g h] by simp + moreover have "D.seq (\ (f, g \\<^sub>C h)) ((F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h])" + using assms C.VV.arr_char FF_def C.VV.dom_simp C.VV.cod_simp by auto + moreover have 2: "D.iso (\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h))" + using assms C.VV.arr_char FF_def C.VV.dom_simp C.VV.cod_simp by auto + moreover have "D.inv (\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h)) = + (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" + using assms 2 C.VV.arr_char FF_def D.inv_comp C.VV.dom_simp C.VV.cod_simp + by simp ultimately show ?thesis - using assms D.isos_compose D.comp_assoc - by (elim conjE D.in_homE) (auto simp add: D.inv_comp) + using D.invert_side_of_triangle(2) + [of "\ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h]" + "F \\<^sub>C[f, g, h]" "\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h)"] + D.comp_assoc + by simp qed - finally show "F \\<^sub>C\<^sup>-\<^sup>1[f, g, h] = - \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F h] \\<^sub>D - (F f \\<^sub>D D.inv (\ (g, h))) \\<^sub>D D.inv (\ (f, g \\<^sub>C h))" - by simp + show "F \\<^sub>C\<^sup>-\<^sup>1[f, g, h] = + \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F h] \\<^sub>D + (F f \\<^sub>D D.inv (\ (g, h))) \\<^sub>D D.inv (\ (f, g \\<^sub>C h))" + proof - + have "F \\<^sub>C\<^sup>-\<^sup>1[f, g, h] = + D.inv (\ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h] \\<^sub>D + (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h)))" + using assms 1 preserves_inv by simp + also have "... = \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) \\<^sub>D \\<^sub>D\<^sup>-\<^sup>1[F f, F g, F h] \\<^sub>D + (F f \\<^sub>D D.inv (\ (g, h))) \\<^sub>D D.inv (\ (f, g \\<^sub>C h))" + proof - + have "\\ (f, g \\<^sub>C h) : F f \\<^sub>D F (g \\<^sub>C h) \\<^sub>D F (f \\<^sub>C g \\<^sub>C h)\ \ D.iso (\ (f, g \\<^sub>C h))" + using assms by auto + moreover have "\F f \\<^sub>D \ (g, h) : F f \\<^sub>D F g \\<^sub>D F h \\<^sub>D F f \\<^sub>D F (g \\<^sub>C h)\ \ + D.iso (F f \\<^sub>D \ (g, h))" + using assms + by (intro conjI D.hcomp_in_vhom, auto) + ultimately show ?thesis + using assms D.isos_compose D.comp_assoc + by (elim conjE D.in_homE) (auto simp add: D.inv_comp) + qed + finally show ?thesis by simp + qed qed lemma preserves_hcomp: assumes "C.hseq \ \" shows "F (\ \\<^sub>C \) = \ (C.cod \, C.cod \) \\<^sub>D (F \ \\<^sub>D F \) \\<^sub>D D.inv (\ (C.dom \, C.dom \))" proof - have "F (\ \\<^sub>C \) \\<^sub>D \ (C.dom \, C.dom \) = \ (C.cod \, C.cod \) \\<^sub>D (F \ \\<^sub>D F \)" using assms \.naturality C.VV.arr_char C.VV.cod_char FF_def C.VV.dom_simp by auto thus ?thesis by (metis (no_types) assms C.hcomp_simps(3) C.hseqE C.ide_dom C.src_dom C.trg_dom D.comp_arr_inv' D.comp_assoc cmp_components_are_iso cmp_simps(5) as_nat_trans.is_natural_1) qed lemma preserves_adjunction_data: assumes "adjunction_data_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \" shows "adjunction_data_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D unit (src\<^sub>C f)) (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g))" proof interpret adjunction_data_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \ using assms by auto show "D.ide (F f)" using preserves_ide by simp show "D.ide (F g)" using preserves_ide by simp show "\D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D unit (src\<^sub>C f) : src\<^sub>D (F f) \\<^sub>D F g \\<^sub>D F f\" using antipar C.VV.ide_char C.VV.arr_char cmp_in_hom(2) unit_in_hom FF_def by auto show "\D.inv (unit (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g) : F f \\<^sub>D F g \\<^sub>D src\<^sub>D (F g)\" using antipar C.VV.ide_char C.VV.arr_char FF_def cmp_in_hom(2) unit_in_hom(2) unit_char(2) by auto qed lemma preserves_equivalence: assumes "equivalence_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \" shows "equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D unit (src\<^sub>C f)) (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g))" proof - interpret equivalence_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \ using assms by auto show "equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D unit (src\<^sub>C f)) (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g))" using antipar unit_is_iso preserves_iso unit_char(2) C.VV.ide_char C.VV.arr_char FF_def D.isos_compose by (unfold_locales) auto qed lemma preserves_equivalence_maps: assumes "C.equivalence_map f" shows "D.equivalence_map (F f)" proof - obtain g \ \ where E: "equivalence_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \" using assms C.equivalence_map_def by auto interpret E: equivalence_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \ using E by auto show ?thesis using E preserves_equivalence C.equivalence_map_def D.equivalence_map_def map\<^sub>0_simps by blast qed lemma preserves_equivalent_objects: assumes "C.equivalent_objects a b" shows "D.equivalent_objects (map\<^sub>0 a) (map\<^sub>0 b)" using assms D.equivalent_objects_def preserves_equivalence_maps unfolding C.equivalent_objects_def by fast lemma preserves_isomorphic: assumes "C.isomorphic f g" shows "D.isomorphic (F f) (F g)" using assms C.isomorphic_def D.isomorphic_def preserves_iso by auto lemma preserves_quasi_inverses: assumes "C.quasi_inverses f g" shows "D.quasi_inverses (F f) (F g)" using assms C.quasi_inverses_def D.quasi_inverses_def preserves_equivalence by blast lemma preserves_quasi_inverse: assumes "C.equivalence_map f" shows "D.isomorphic (F (C.some_quasi_inverse f)) (D.some_quasi_inverse (F f))" using assms preserves_quasi_inverses C.quasi_inverses_some_quasi_inverse D.quasi_inverse_unique D.quasi_inverses_some_quasi_inverse preserves_equivalence_maps by blast end subsection "Identity Pseudofunctors" locale identity_pseudofunctor = B: bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B 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" begin text\ The underlying vertical functor is just the identity functor on the vertical category, which is already available as \B.map\. \ abbreviation map where "map \ B.map" interpretation I: weak_arrow_of_homs V\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B src\<^sub>B trg\<^sub>B map - proof - show "\\. B.arr \ \ B.isomorphic (map (src\<^sub>B \)) (src\<^sub>B (map \))" - by (simp add: B.isomorphic_reflexive) - show "\\. B.arr \ \ B.isomorphic (map (trg\<^sub>B \)) (trg\<^sub>B (map \))" - by (simp add: B.isomorphic_reflexive) - qed + using B.isomorphic_reflexive by unfold_locales auto interpretation II: "functor" B.VV.comp B.VV.comp I.FF using I.functor_FF by simp interpretation H\<^sub>BoII: composite_functor B.VV.comp B.VV.comp V\<^sub>B I.FF \\\\. fst \\ \\<^sub>B snd \\\ .. interpretation IoH\<^sub>B: composite_functor B.VV.comp V\<^sub>B V\<^sub>B \\\\. fst \\ \\<^sub>B snd \\\ map .. text\ The horizontal composition provides the compositor. \ abbreviation cmp where "cmp \ \\\. fst \\ \\<^sub>B snd \\" interpretation cmp: natural_transformation B.VV.comp V\<^sub>B H\<^sub>BoII.map IoH\<^sub>B.map cmp using B.VV.arr_char B.VV.dom_simp B.VV.cod_simp B.H.as_nat_trans.is_natural_1 B.H.as_nat_trans.is_natural_2 I.FF_def apply unfold_locales apply auto by (meson B.hseqE B.hseq_char')+ interpretation cmp: natural_isomorphism B.VV.comp V\<^sub>B H\<^sub>BoII.map IoH\<^sub>B.map cmp by unfold_locales simp sublocale pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B map cmp apply unfold_locales by (metis B.assoc_is_natural_2 B.assoc_naturality B.assoc_simps(1) B.comp_ide_self B.hcomp_simps(1) B.ide_char B.ide_hcomp B.map_simp fst_conv snd_conv) lemma is_pseudofunctor: shows "pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B map cmp" .. lemma unit_char': assumes "B.obj a" shows "unit a = a" proof - have "src\<^sub>B a = a \ B.ide a" using assms by auto hence "a = unit a" using assms B.comp_arr_dom B.comp_cod_arr I.map\<^sub>0_def map_def B.ide_in_hom(2) B.objE [of a] B.ide_is_iso [of a] by (intro unit_eqI) auto thus ?thesis by simp qed end lemma (in identity_pseudofunctor) map\<^sub>0_simp [simp]: assumes "B.obj a" shows "map\<^sub>0 a = a" using assms map\<^sub>0_def by auto (* TODO: Does not recognize map\<^sub>0_def unless the context is closed, then re-opened. *) subsection "Embedding Pseudofunctors" text \ In this section, we construct the embedding pseudofunctor of a sub-bicategory into the ambient bicategory. \ locale embedding_pseudofunctor = B: bicategory V H \\<^sub>B \ src\<^sub>B trg\<^sub>B + S: subbicategory begin no_notation B.in_hom ("\_ : _ \\<^sub>B _\") notation B.in_hhom ("\_ : _ \\<^sub>B _\") definition map where "map \ = (if S.arr \ then \ else B.null)" lemma map_in_hom [intro]: assumes "S.arr \" shows "\map \ : src\<^sub>B (map (S.src \)) \\<^sub>B src\<^sub>B (map (S.trg \))\" and "\map \ : map (S.dom \) \\<^sub>B map (S.cod \)\" proof - show 1: "\map \ : map (S.dom \) \\<^sub>B map (S.cod \)\" using assms map_def S.in_hom_char by fastforce show "\map \ : src\<^sub>B (map (S.src \)) \\<^sub>B src\<^sub>B (map (S.trg \))\" using assms 1 map_def S.arr_char S.src_def S.trg_def S.obj_char S.obj_src S.obj_trg by auto qed lemma map_simps [simp]: assumes "S.arr \" shows "B.arr (map \)" and "src\<^sub>B (map \) = src\<^sub>B (map (S.src \))" and "trg\<^sub>B (map \) = src\<^sub>B (map (S.trg \))" and "B.dom (map \) = map (S.dom \)" and "B.cod (map \) = map (S.cod \)" using assms map_in_hom by blast+ interpretation "functor" S.comp V map apply unfold_locales apply auto using map_def S.comp_char S.seq_char S.arr_char apply auto[1] using map_def S.comp_simp by auto interpretation weak_arrow_of_homs S.comp S.src S.trg V src\<^sub>B trg\<^sub>B map using S.arr_char map_def S.src_def S.trg_def S.src_closed S.trg_closed B.isomorphic_reflexive preserves_ide S.ide_src S.ide_trg apply unfold_locales by presburger+ interpretation HoFF: composite_functor S.VV.comp B.VV.comp V FF \\\\. fst \\ \\<^sub>B snd \\\ .. interpretation FoH: composite_functor S.VV.comp S.comp V \\\\. fst \\ \ snd \\\ map .. no_notation B.in_hom ("\_ : _ \\<^sub>B _\") definition cmp where "cmp \\ = (if S.VV.arr \\ then fst \\ \\<^sub>B snd \\ else B.null)" lemma cmp_in_hom [intro]: assumes "S.VV.arr \\" shows "\cmp \\ : src\<^sub>B (snd \\) \\<^sub>B trg\<^sub>B (fst \\)\" and "\cmp \\ : map (S.dom (fst \\)) \\<^sub>B map (S.dom (snd \\)) \\<^sub>B map (S.cod (fst \\) \ S.cod (snd \\))\" proof - show "\cmp \\ : map (S.dom (fst \\)) \\<^sub>B map (S.dom (snd \\)) \\<^sub>B map (S.cod (fst \\) \ S.cod (snd \\))\" proof show 1: "B.arr (cmp \\)" unfolding cmp_def using assms S.arr_char S.VV.arr_char S.inclusion S.src_def S.trg_def by auto show "B.dom (cmp \\) = map (S.dom (fst \\)) \\<^sub>B map (S.dom (snd \\))" unfolding cmp_def map_def using assms 1 cmp_def S.dom_simp S.cod_simp S.VV.arr_char S.arr_char S.hcomp_def S.inclusion S.dom_closed by auto show "B.cod (cmp \\) = map (S.cod (fst \\) \ S.cod (snd \\))" unfolding cmp_def map_def - using assms 1 cmp_def S.dom_simp S.cod_simp S.VV.arr_char S.arr_char S.hcomp_def - S.inclusion S.cod_closed - apply auto - using S.hcomp_closed apply auto[1] - by (metis (no_types, lifting) B.VV.arrE) + using assms 1 S.H.preserves_arr S.cod_simp S.hcomp_eqI S.hcomp_simps(4) S.hseq_char' + by auto qed thus "\cmp \\ : src\<^sub>B (snd \\) \\<^sub>B trg\<^sub>B (fst \\)\" using cmp_def by auto qed lemma cmp_simps [simp]: assumes "S.VV.arr \\" shows "B.arr (cmp \\)" and "src\<^sub>B (cmp \\) = S.src (snd \\)" and "trg\<^sub>B (cmp \\) = S.trg (fst \\)" and "B.dom (cmp \\) = map (S.dom (fst \\)) \\<^sub>B map (S.dom (snd \\))" and "B.cod (cmp \\) = map (S.cod (fst \\) \ S.cod (snd \\))" using assms cmp_in_hom S.src_def S.trg_def S.VV.arr_char by simp_all blast+ lemma iso_cmp: assumes "S.VV.ide \\" shows "B.iso (cmp \\)" using assms S.VV.ide_char S.VV.arr_char S.arr_char cmp_def S.ide_char S.src_def S.trg_def by auto interpretation \\<^sub>E: natural_isomorphism S.VV.comp V HoFF.map FoH.map cmp proof show "\\\. \ S.VV.arr \\ \ cmp \\ = B.null" using cmp_def by simp fix \\ assume \\: "S.VV.arr \\" have 1: "S.arr (fst \\) \ S.arr (snd \\) \ S.src (fst \\) = S.trg (snd \\)" using \\ S.VV.arr_char by simp show "B.dom (cmp \\) = HoFF.map (S.VV.dom \\)" using \\ FF_def S.VV.arr_char S.VV.dom_char S.arr_dom S.src_def S.trg_def S.dom_char S.src.preserves_dom S.trg.preserves_dom apply simp by (metis (no_types, lifting)) show "B.cod (cmp \\) = FoH.map (S.VV.cod \\)" using \\ 1 map_def S.hseq_char S.hcomp_def S.cod_char S.arr_cod S.VV.cod_simp by simp show "cmp (S.VV.cod \\) \\<^sub>B HoFF.map \\ = cmp \\" using \\ 1 cmp_def S.VV.arr_char S.VV.cod_char FF_def S.arr_cod S.cod_simp S.src_def S.trg_def map_def apply simp by (metis (no_types, lifting) B.comp_cod_arr B.hcomp_simps(4) cmp_simps(1) \\) show "FoH.map \\ \\<^sub>B cmp (S.VV.dom \\) = cmp \\" unfolding cmp_def map_def using \\ S.VV.arr_char B.VV.arr_char S.VV.dom_char S.VV.cod_char B.comp_arr_dom S.hcomp_def apply simp by (metis (no_types, lifting) B.hcomp_simps(3) cmp_def cmp_simps(1) S.arr_char S.dom_char S.hcomp_closed S.src_def S.trg_def) next show "\fg. S.VV.ide fg \ B.iso (cmp fg)" using iso_cmp by simp qed sublocale pseudofunctor S.comp S.hcomp S.\ \ S.src S.trg V H \\<^sub>B \ src\<^sub>B trg\<^sub>B map cmp proof fix f g h assume f: "S.ide f" and g: "S.ide g" and h: "S.ide h" and fg: "S.src f = S.trg g" and gh: "S.src g = S.trg h" have 1: "B.ide f \ B.ide g \ B.ide h \ src\<^sub>B f = trg\<^sub>B g \ src\<^sub>B g = trg\<^sub>B h" using f g h fg gh S.ide_char S.arr_char S.src_def S.trg_def by simp show "map (S.\ f g h) \\<^sub>B cmp (f \ g, h) \\<^sub>B cmp (f, g) \\<^sub>B map h = cmp (f, g \ h) \\<^sub>B (map f \\<^sub>B cmp (g, h)) \\<^sub>B \\<^sub>B[map f, map g, map h]" proof - have "map (S.\ f g h) \\<^sub>B cmp (f \ g, h) \\<^sub>B cmp (f, g) \\<^sub>B map h = \\<^sub>B[f, g, h] \\<^sub>B ((f \\<^sub>B g) \\<^sub>B h) \\<^sub>B ((f \\<^sub>B g) \\<^sub>B h)" unfolding map_def cmp_def using 1 f g h fg gh S.VVV.arr_char S.VV.arr_char B.VVV.arr_char B.VV.arr_char B.comp_arr_dom S.arr_char S.comp_char S.hcomp_closed S.hcomp_def S.ideD(1) S.src_def by (simp add: S.assoc_closed) also have "... = cmp (f, g \ h) \\<^sub>B (map f \\<^sub>B cmp (g, h)) \\<^sub>B \\<^sub>B[map f, map g, map h]" unfolding cmp_def map_def using 1 f g h fg gh S.VV.arr_char B.VVV.arr_char B.VV.arr_char B.comp_arr_dom B.comp_cod_arr S.hcomp_def S.comp_char S.arr_char S.assoc_closed S.hcomp_closed S.ideD(1) S.trg_def by auto finally show ?thesis by blast qed qed lemma is_pseudofunctor: shows "pseudofunctor S.comp S.hcomp S.\ \ S.src S.trg V H \\<^sub>B \ src\<^sub>B trg\<^sub>B map cmp" .. lemma map\<^sub>0_simp [simp]: assumes "S.obj a" shows "map\<^sub>0 a = a" using assms map\<^sub>0_def map_def S.obj_char by auto lemma unit_char': assumes "S.obj a" shows "unit a = a" proof - have a: "S.arr a" using assms by auto have 1: "B.ide a" using assms S.obj_char by auto have 2: "src\<^sub>B a = a" using assms S.obj_char by auto have "a = unit a" proof (intro unit_eqI) show "S.obj a" by fact show "\a : map\<^sub>0 a \\<^sub>B map a\" using assms a 2 map\<^sub>0_def map_def S.src_def S.dom_char S.cod_char S.obj_char by auto show "B.iso a" using assms 1 by simp show "a \\<^sub>B \[map\<^sub>0 a] = (map \[a] \\<^sub>B cmp (a, a)) \\<^sub>B (a \\<^sub>B a)" proof - have "a \\<^sub>B \[a] = \[a] \\<^sub>B cmp (a, a) \\<^sub>B (a \\<^sub>B a)" proof - have "a \\<^sub>B \[a] = \[a]" using assms 1 2 S.comp_cod_arr S.unitor_coincidence S.lunit_in_hom S.obj_char S.comp_simp by auto moreover have "(a \\<^sub>B a) \\<^sub>B (a \\<^sub>B a) = a \\<^sub>B a" using assms S.obj_char S.comp_ide_self by auto moreover have "B.dom \[a] = a \\<^sub>B a" using assms S.obj_char by simp moreover have "\[a] \\<^sub>B (a \\<^sub>B a) = \[a]" using assms S.obj_char B.comp_arr_dom by simp ultimately show ?thesis using assms cmp_def S.VV.arr_char by auto qed thus ?thesis using assms a 2 map\<^sub>0_def map_def S.src_def B.comp_assoc by simp qed qed thus ?thesis by simp qed end subsection "Composition of Pseudofunctors" text \ In this section, we show how pseudofunctors may be composed. The main work is to establish the coherence condition for associativity. \ locale composite_pseudofunctor = 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>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>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 :: "'c \ 'd" and \\<^sub>G :: "'c * 'c \ 'd" begin sublocale composite_functor V\<^sub>B V\<^sub>C V\<^sub>D F G .. sublocale weak_arrow_of_homs V\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D src\<^sub>D trg\<^sub>D \G o F\ proof show "\\. B.arr \ \ D.isomorphic ((G o F) (src\<^sub>B \)) (src\<^sub>D ((G o F) \))" proof - fix \ assume \: "B.arr \" show "D.isomorphic ((G o F) (src\<^sub>B \)) (src\<^sub>D ((G o F) \))" proof - have "(G o F) (src\<^sub>B \) = G (F (src\<^sub>B \))" using \ by simp also have "D.isomorphic ... (G (src\<^sub>C (F \)))" using \ F.weakly_preserves_src G.preserves_iso by (meson C.isomorphicE D.isomorphic_def G.preserves_hom) also have "D.isomorphic ... (src\<^sub>D (G (F \)))" using \ G.weakly_preserves_src by blast also have "... = src\<^sub>D ((G o F) \)" by simp finally show ?thesis by blast qed qed show "\\. B.arr \ \ D.isomorphic ((G o F) (trg\<^sub>B \)) (trg\<^sub>D ((G o F) \))" proof - fix \ assume \: "B.arr \" show "D.isomorphic ((G o F) (trg\<^sub>B \)) (trg\<^sub>D ((G o F) \))" proof - have "(G o F) (trg\<^sub>B \) = G (F (trg\<^sub>B \))" using \ by simp also have "D.isomorphic ... (G (trg\<^sub>C (F \)))" using \ F.weakly_preserves_trg G.preserves_iso by (meson C.isomorphicE D.isomorphic_def G.preserves_hom) also have "D.isomorphic ... (trg\<^sub>D (G (F \)))" using \ G.weakly_preserves_trg by blast also have "... = trg\<^sub>D ((G o F) \)" by simp finally show ?thesis by blast qed qed qed interpretation H\<^sub>DoGF_GF: composite_functor B.VV.comp D.VV.comp V\<^sub>D FF \\\\. fst \\ \\<^sub>D snd \\\ .. interpretation GFoH\<^sub>B: composite_functor B.VV.comp V\<^sub>B V\<^sub>D \\\\. fst \\ \\<^sub>B snd \\\ \G o F\ .. definition cmp where "cmp \\ = (if B.VV.arr \\ then G (F (H\<^sub>B (fst \\) (snd \\))) \\<^sub>D G (\\<^sub>F (B.VV.dom \\)) \\<^sub>D \\<^sub>G (F (B.dom (fst \\)), F (B.dom (snd \\))) - else D.null)" + else D.null)" lemma cmp_in_hom [intro]: assumes "B.VV.arr \\" shows "\cmp \\ : H\<^sub>DoGF_GF.map (B.VV.dom \\) \\<^sub>D GFoH\<^sub>B.map (B.VV.cod \\)\" proof - have "cmp \\ = G (F (H\<^sub>B (fst \\) (snd \\))) \\<^sub>D G (\\<^sub>F (B.VV.dom \\)) \\<^sub>D \\<^sub>G (F (B.dom (fst \\)), F (B.dom (snd \\)))" using assms cmp_def by simp moreover have "\ ... : H\<^sub>DoGF_GF.map (B.VV.dom \\) \\<^sub>D GFoH\<^sub>B.map (B.VV.cod \\)\" proof (intro D.comp_in_homI) show "\\\<^sub>G (F (B.dom (fst \\)), F (B.dom (snd \\))) : H\<^sub>DoGF_GF.map (B.VV.dom \\) \\<^sub>D G (F (B.dom (fst \\)) \\<^sub>C F (B.dom (snd \\)))\" using assms F.FF_def FF_def B.VV.arr_char B.VV.dom_simp by auto show "\G (\\<^sub>F (B.VV.dom \\)) : G (F (B.dom (fst \\)) \\<^sub>C F (B.dom (snd \\))) \\<^sub>D GFoH\<^sub>B.map (B.VV.dom \\)\" using assms B.VV.arr_char F.FF_def B.VV.dom_simp B.VV.cod_simp by auto show "\G (F (fst \\ \\<^sub>B snd \\)) : GFoH\<^sub>B.map (B.VV.dom \\) \\<^sub>D GFoH\<^sub>B.map (B.VV.cod \\)\" proof - have "B.VV.in_hom \\ (B.VV.dom \\) (B.VV.cod \\)" using assms by auto thus ?thesis by auto qed qed ultimately show ?thesis by argo qed lemma cmp_simps [simp]: assumes "B.VV.arr \\" shows "D.arr (cmp \\)" and "D.dom (cmp \\) = H\<^sub>DoGF_GF.map (B.VV.dom \\)" and "D.cod (cmp \\) = GFoH\<^sub>B.map (B.VV.cod \\)" using assms cmp_in_hom by blast+ interpretation \: natural_transformation B.VV.comp V\<^sub>D H\<^sub>DoGF_GF.map GFoH\<^sub>B.map cmp proof show "\\\. \ B.VV.arr \\ \ cmp \\ = D.null" unfolding cmp_def by simp fix \\ assume \\: "B.VV.arr \\" show "D.dom (cmp \\) = H\<^sub>DoGF_GF.map (B.VV.dom \\)" using \\ cmp_in_hom by blast show "D.cod (cmp \\) = GFoH\<^sub>B.map (B.VV.cod \\)" using \\ cmp_in_hom by blast show "GFoH\<^sub>B.map \\ \\<^sub>D cmp (B.VV.dom \\) = cmp \\" unfolding cmp_def using \\ B.VV.ide_char B.VV.arr_char D.comp_ide_arr B.VV.dom_char D.comp_assoc as_nat_trans.is_natural_1 apply simp by (metis (no_types, lifting) B.H.preserves_arr B.hcomp_simps(3)) show "cmp (B.VV.cod \\) \\<^sub>D H\<^sub>DoGF_GF.map \\ = cmp \\" proof - have "cmp (B.VV.cod \\) \\<^sub>D H\<^sub>DoGF_GF.map \\ = (G (F (B.cod (fst \\) \\<^sub>B B.cod (snd \\))) \\<^sub>D G (\\<^sub>F (B.cod (fst \\), B.cod (snd \\))) \\<^sub>D \\<^sub>G (F (B.cod (fst \\)), F (B.cod (snd \\)))) \\<^sub>D (fst (FF \\) \\<^sub>D snd (FF \\))" unfolding cmp_def using \\ B.VV.arr_char B.VV.dom_simp B.VV.cod_simp by simp also have "... = (G (\\<^sub>F (B.cod (fst \\), B.cod (snd \\))) \\<^sub>D \\<^sub>G (F (B.cod (fst \\)), F (B.cod (snd \\)))) \\<^sub>D (fst (FF \\) \\<^sub>D snd (FF \\))" proof - have "D.ide (G (F (B.cod (fst \\) \\<^sub>B B.cod (snd \\))))" using \\ B.VV.arr_char by simp moreover have "D.seq (G (F (B.cod (fst \\) \\<^sub>B B.cod (snd \\)))) (G (\\<^sub>F (B.cod (fst \\), B.cod (snd \\))) \\<^sub>D \\<^sub>G (F (B.cod (fst \\)), F (B.cod (snd \\))))" using \\ B.VV.arr_char B.VV.dom_char B.VV.cod_char F.FF_def apply (intro D.seqI) apply auto proof - show "G (F (B.cod (fst \\) \\<^sub>B B.cod (snd \\))) = D.cod (G (\\<^sub>F (B.cod (fst \\), B.cod (snd \\))) \\<^sub>D \\<^sub>G (F (B.cod (fst \\)), F (B.cod (snd \\))))" proof - have "D.seq (G (\\<^sub>F (B.cod (fst \\), B.cod (snd \\)))) (\\<^sub>G (F (B.cod (fst \\)), F (B.cod (snd \\))))" using \\ B.VV.arr_char F.FF_def B.VV.arr_char B.VV.dom_char B.VV.cod_char by (intro D.seqI) auto thus ?thesis using \\ B.VV.arr_char B.VV.cod_simp by simp qed qed ultimately show ?thesis using \\ D.comp_ide_arr [of "G (F (B.cod (fst \\) \\<^sub>B B.cod (snd \\)))" "G (\\<^sub>F (B.cod (fst \\), B.cod (snd \\))) \\<^sub>D \\<^sub>G (F (B.cod (fst \\)), F (B.cod (snd \\)))"] by simp qed also have "... = G (\\<^sub>F (B.cod (fst \\), B.cod (snd \\))) \\<^sub>D (\\<^sub>G (F (B.cod (fst \\)), F (B.cod (snd \\))) \\<^sub>D (fst (FF \\) \\<^sub>D snd (FF \\)))" using D.comp_assoc by simp also have "... = G (\\<^sub>F (B.cod (fst \\), B.cod (snd \\))) \\<^sub>D \\<^sub>G (C.VV.cod (F.FF \\)) \\<^sub>D G.H\<^sub>DoFF.map (F.FF \\)" using \\ B.VV.arr_char F.FF_def G.FF_def FF_def C.VV.cod_simp by auto also have "... = G (\\<^sub>F (B.cod (fst \\), B.cod (snd \\))) \\<^sub>D G.FoH\<^sub>C.map (F.FF \\) \\<^sub>D \\<^sub>G (C.VV.dom (F.FF \\))" using \\ B.VV.arr_char G.\.naturality C.VV.dom_simp C.VV.cod_simp by simp also have "... = (G (\\<^sub>F (B.cod (fst \\), B.cod (snd \\))) \\<^sub>D G.FoH\<^sub>C.map (F.FF \\)) \\<^sub>D \\<^sub>G (C.VV.dom (F.FF \\))" using D.comp_assoc by simp also have "... = (G (\\<^sub>F (B.VV.cod \\) \\<^sub>C F.H\<^sub>DoFF.map \\)) \\<^sub>D \\<^sub>G (C.VV.dom (F.FF \\))" proof - have "(B.cod (fst \\), B.cod (snd \\)) = B.VV.cod \\" using \\ B.VV.arr_char B.VV.cod_simp by simp moreover have "G.FoH\<^sub>C.map (F.FF \\) = G (F.H\<^sub>DoFF.map \\)" using \\ F.FF_def by simp moreover have "G (\\<^sub>F (B.cod (fst \\), B.cod (snd \\))) \\<^sub>D G (F.H\<^sub>DoFF.map \\) = G (\\<^sub>F (B.VV.cod \\) \\<^sub>C F.H\<^sub>DoFF.map \\)" using \\ B.VV.arr_char by (metis (no_types, lifting) F.\.is_natural_2 F.\.preserves_reflects_arr G.preserves_comp calculation(1)) ultimately show ?thesis by argo qed also have "... = G (F.FoH\<^sub>C.map \\ \\<^sub>C \\<^sub>F (B.VV.dom \\)) \\<^sub>D \\<^sub>G (C.VV.dom (F.FF \\))" using \\ F.\.naturality [of \\] F.FF_def by simp also have "... = G (F.FoH\<^sub>C.map \\) \\<^sub>D G (\\<^sub>F (B.VV.dom \\)) \\<^sub>D \\<^sub>G (C.VV.dom (F.FF \\))" proof - have "G (F.FoH\<^sub>C.map \\ \\<^sub>C \\<^sub>F (B.VV.dom \\)) = G (F.FoH\<^sub>C.map \\) \\<^sub>D G (\\<^sub>F (B.VV.dom \\))" using \\ by (metis (mono_tags, lifting) F.\.is_natural_1 F.\.preserves_reflects_arr G.preserves_comp) thus ?thesis using \\ D.comp_assoc by simp qed also have "... = cmp \\" using \\ B.VV.arr_char cmp_def F.FF_def F.FF.preserves_dom B.VV.dom_simp by auto finally show ?thesis by simp qed qed interpretation \: natural_isomorphism B.VV.comp V\<^sub>D H\<^sub>DoGF_GF.map GFoH\<^sub>B.map cmp proof show "\hk. B.VV.ide hk \ D.iso (cmp hk)" proof - fix hk assume hk: "B.VV.ide hk" have "D.iso (G (F (fst hk \\<^sub>B snd hk)) \\<^sub>D G (\\<^sub>F (B.VV.dom hk)) \\<^sub>D \\<^sub>G (F (B.dom (fst hk)), F (B.dom (snd hk))))" proof (intro D.isos_compose) show "D.iso (\\<^sub>G (F (B.dom (fst hk)), F (B.dom (snd hk))))" using hk G.\.components_are_iso [of "(F (B.dom (fst hk)), F (B.dom (snd hk)))"] C.VV.ide_char B.VV.arr_char B.VV.dom_char by (metis (no_types, lifting) B.VV.ideD(1) B.VV.ideD(2) B.VxV.dom_char F.FF_def F.FF.as_nat_iso.components_are_iso G.\.preserves_iso fst_conv snd_conv) show "D.iso (G (\\<^sub>F (B.VV.dom hk)))" using hk F.\.components_are_iso B.VV.arr_char B.VV.dom_char B.VV.ideD(2) by auto show "D.seq (G (\\<^sub>F (B.VV.dom hk))) (\\<^sub>G (F (B.dom (fst hk)), F (B.dom (snd hk))))" using hk B.VV.arr_char B.VV.ide_char B.VV.dom_char C.VV.arr_char F.FF_def C.VV.dom_simp C.VV.cod_simp by auto show "D.iso (G (F (fst hk \\<^sub>B snd hk)))" using hk F.\.components_are_iso B.VV.arr_char by simp show "D.seq (G (F (fst hk \\<^sub>B snd hk))) (G (\\<^sub>F (B.VV.dom hk)) \\<^sub>D \\<^sub>G (F (B.dom (fst hk)), F (B.dom (snd hk))))" using hk B.VV.arr_char B.VV.dom_char by (metis (no_types, lifting) B.VV.ideD(1) cmp_def cmp_simps(1)) qed thus "D.iso (cmp hk)" unfolding cmp_def using hk by simp qed qed sublocale pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \G o F\ cmp proof fix f g h assume f: "B.ide f" and g: "B.ide g" and h: "B.ide h" assume fg: "src\<^sub>B f = trg\<^sub>B g" and gh: "src\<^sub>B g = trg\<^sub>B h" show "map \\<^sub>B[f, g, h] \\<^sub>D cmp (f \\<^sub>B g, h) \\<^sub>D (cmp (f, g) \\<^sub>D map h) = cmp (f, g \\<^sub>B h) \\<^sub>D (map f \\<^sub>D cmp (g, h)) \\<^sub>D \\<^sub>D[map f, map g, map h]" proof - have "map \\<^sub>B[f, g, h] \\<^sub>D cmp (f \\<^sub>B g, h) \\<^sub>D (cmp (f, g) \\<^sub>D map h) = G (F \\<^sub>B[f, g, h]) \\<^sub>D (G (F ((f \\<^sub>B g) \\<^sub>B h)) \\<^sub>D G (\\<^sub>F (f \\<^sub>B g, h)) \\<^sub>D \\<^sub>G (F (f \\<^sub>B g), F h)) \\<^sub>D (G (F (f \\<^sub>B g)) \\<^sub>D G (\\<^sub>F (f, g)) \\<^sub>D \\<^sub>G (F f, F g) \\<^sub>D G (F h))" unfolding cmp_def using f g h fg gh B.VV.arr_char B.VV.dom_simp by simp also have "... = G (F \\<^sub>B[f, g, h]) \\<^sub>D (G (\\<^sub>F (f \\<^sub>B g, h)) \\<^sub>D \\<^sub>G (F (f \\<^sub>B g), F h)) \\<^sub>D (G (F (f \\<^sub>B g)) \\<^sub>D G (\\<^sub>F (f, g)) \\<^sub>D \\<^sub>G (F f, F g) \\<^sub>D G (F h))" using f g h fg gh D.comp_ide_arr D.comp_assoc by (metis B.ideD(1) B.ide_hcomp B.src_hcomp F.cmp_simps(1) F.cmp_simps(5) G.as_nat_trans.is_natural_2) also have "... = G (F \\<^sub>B[f, g, h]) \\<^sub>D (G (\\<^sub>F (f \\<^sub>B g, h)) \\<^sub>D \\<^sub>G (F (f \\<^sub>B g), F h)) \\<^sub>D (G (\\<^sub>F (f, g)) \\<^sub>D \\<^sub>G (F f, F g) \\<^sub>D G (F h))" using f g fg by (metis (no_types) D.comp_assoc F.cmp_simps(1) F.cmp_simps(5) G.as_nat_trans.is_natural_2) also have "... = (G (F \\<^sub>B[f, g, h]) \\<^sub>D G (\\<^sub>F (f \\<^sub>B g, h))) \\<^sub>D \\<^sub>G (F (f \\<^sub>B g), F h) \\<^sub>D (G (\\<^sub>F (f, g)) \\<^sub>D \\<^sub>G (F f, F g) \\<^sub>D G (F h))" using D.comp_assoc by simp also have "... = G (F \\<^sub>B[f, g, h] \\<^sub>C \\<^sub>F (f \\<^sub>B g, h)) \\<^sub>D \\<^sub>G (F (f \\<^sub>B g), F h) \\<^sub>D (G (\\<^sub>F (f, g)) \\<^sub>D \\<^sub>G (F f, F g) \\<^sub>D G (F h))" using f g h fg gh B.VV.arr_char B.VV.cod_simp by simp also have "... = G (F \\<^sub>B[f, g, h] \\<^sub>C \\<^sub>F (f \\<^sub>B g, h)) \\<^sub>D \\<^sub>G (F (f \\<^sub>B g), F h) \\<^sub>D (G (\\<^sub>F (f, g)) \\<^sub>D G (F h)) \\<^sub>D (\\<^sub>G (F f, F g) \\<^sub>D G (F h))" using f g h fg gh B.VV.arr_char C.VV.arr_char F.FF_def D.whisker_right B.VV.dom_simp C.VV.cod_simp by auto also have "... = G (F \\<^sub>B[f, g, h] \\<^sub>C \\<^sub>F (f \\<^sub>B g, h)) \\<^sub>D (\\<^sub>G (F (f \\<^sub>B g), F h) \\<^sub>D (G (\\<^sub>F (f, g)) \\<^sub>D G (F h))) \\<^sub>D (\\<^sub>G (F f, F g) \\<^sub>D G (F h))" using D.comp_assoc by simp also have "... = G (F \\<^sub>B[f, g, h] \\<^sub>C \\<^sub>F (f \\<^sub>B g, h)) \\<^sub>D (G (\\<^sub>F (f, g) \\<^sub>C F h) \\<^sub>D \\<^sub>G (F f \\<^sub>C F g, F h)) \\<^sub>D (\\<^sub>G (F f, F g) \\<^sub>D G (F h))" proof - have "\\<^sub>G (F (f \\<^sub>B g), F h) = \\<^sub>G (C.VV.cod (\\<^sub>F (f, g), F h))" using f g h fg gh B.VV.arr_char C.VV.arr_char B.VV.cod_simp C.VV.cod_simp by simp moreover have "G (\\<^sub>F (f, g)) \\<^sub>D G (F h) = G.H\<^sub>DoFF.map (\\<^sub>F (f, g), F h)" using f g h fg gh B.VV.arr_char C.VV.arr_char G.FF_def by simp moreover have "G.FoH\<^sub>C.map (\\<^sub>F (f, g), F h) = G (\\<^sub>F (f, g) \\<^sub>C F h)" using f g h fg gh B.VV.arr_char by simp moreover have "\\<^sub>G (C.VV.dom (\\<^sub>F (f, g), F h)) = \\<^sub>G (F f \\<^sub>C F g, F h)" using f g h fg gh C.VV.arr_char F.cmp_in_hom [of f g] C.VV.dom_simp by auto ultimately show ?thesis using f g h fg gh B.VV.arr_char G.\.naturality by (metis (mono_tags, lifting) C.VV.arr_cod_iff_arr C.VV.arr_dom_iff_arr G.FoH\<^sub>C.is_extensional G.H\<^sub>DoFF.is_extensional G.\.is_extensional) qed also have "... = (G (F \\<^sub>B[f, g, h] \\<^sub>C \\<^sub>F (f \\<^sub>B g, h)) \\<^sub>D (G (\\<^sub>F (f, g) \\<^sub>C F h))) \\<^sub>D \\<^sub>G (F f \\<^sub>C F g, F h) \\<^sub>D (\\<^sub>G (F f, F g) \\<^sub>D G (F h))" using D.comp_assoc by simp also have "... = G ((F \\<^sub>B[f, g, h] \\<^sub>C \\<^sub>F (f \\<^sub>B g, h)) \\<^sub>C (\\<^sub>F (f, g) \\<^sub>C F h)) \\<^sub>D \\<^sub>G (F f \\<^sub>C F g, F h) \\<^sub>D (\\<^sub>G (F f, F g) \\<^sub>D G (F h))" using f g h fg gh B.VV.arr_char F.FF_def B.VV.dom_simp B.VV.cod_simp by auto also have "... = G (F \\<^sub>B[f, g, h] \\<^sub>C \\<^sub>F (f \\<^sub>B g, h) \\<^sub>C (\\<^sub>F (f, g) \\<^sub>C F h)) \\<^sub>D \\<^sub>G (F f \\<^sub>C F g, F h) \\<^sub>D (\\<^sub>G (F f, F g) \\<^sub>D G (F h))" using C.comp_assoc by simp also have "... = G (\\<^sub>F (f, g \\<^sub>B h) \\<^sub>C (F f \\<^sub>C \\<^sub>F (g, h)) \\<^sub>C \\<^sub>C[F f, F g, F h]) \\<^sub>D \\<^sub>G (F f \\<^sub>C F g, F h) \\<^sub>D (\\<^sub>G (F f, F g) \\<^sub>D G (F h))" using f g h fg gh F.assoc_coherence [of f g h] by simp also have "... = G ((\\<^sub>F (f, g \\<^sub>B h) \\<^sub>C (F f \\<^sub>C \\<^sub>F (g, h))) \\<^sub>C \\<^sub>C[F f, F g, F h]) \\<^sub>D \\<^sub>G (F f \\<^sub>C F g, F h) \\<^sub>D (\\<^sub>G (F f, F g) \\<^sub>D G (F h))" using C.comp_assoc by simp also have "... = (G (\\<^sub>F (f, g \\<^sub>B h) \\<^sub>C (F f \\<^sub>C \\<^sub>F (g, h))) \\<^sub>D G \\<^sub>C[F f, F g, F h]) \\<^sub>D \\<^sub>G (F f \\<^sub>C F g, F h) \\<^sub>D (\\<^sub>G (F f, F g) \\<^sub>D G (F h))" using f g h fg gh B.VV.arr_char F.FF_def B.VV.dom_simp B.VV.cod_simp by auto also have "... = G (\\<^sub>F (f, g \\<^sub>B h) \\<^sub>C (F f \\<^sub>C \\<^sub>F (g, h))) \\<^sub>D G \\<^sub>C[F f, F g, F h] \\<^sub>D \\<^sub>G (F f \\<^sub>C F g, F h) \\<^sub>D (\\<^sub>G (F f, F g) \\<^sub>D G (F h))" using D.comp_assoc by simp also have "... = G (\\<^sub>F (f, g \\<^sub>B h) \\<^sub>C (F f \\<^sub>C \\<^sub>F (g, h))) \\<^sub>D \\<^sub>G (F f, F g \\<^sub>C F h) \\<^sub>D (G (F f) \\<^sub>D \\<^sub>G (F g, F h)) \\<^sub>D \\<^sub>D[G (F f), G (F g), G (F h)]" using f g h fg gh G.assoc_coherence [of "F f" "F g" "F h"] by simp also have "... = (G (\\<^sub>F (f, g \\<^sub>B h) \\<^sub>C (F f \\<^sub>C \\<^sub>F (g, h))) \\<^sub>D \\<^sub>G (F f, F g \\<^sub>C F h) \\<^sub>D (G (F f) \\<^sub>D \\<^sub>G (F g, F h))) \\<^sub>D \\<^sub>D[G (F f), G (F g), G (F h)]" using D.comp_assoc by simp also have "... = (cmp (f, g \\<^sub>B h) \\<^sub>D (G (F f) \\<^sub>D cmp (g, h))) \\<^sub>D \\<^sub>D[G (F f), G (F g), G (F h)]" proof - have "G (\\<^sub>F (f, g \\<^sub>B h) \\<^sub>C (F f \\<^sub>C \\<^sub>F (g, h))) \\<^sub>D \\<^sub>G (F f, F g \\<^sub>C F h) \\<^sub>D (G (F f) \\<^sub>D \\<^sub>G (F g, F h)) = cmp (f, g \\<^sub>B h) \\<^sub>D (G (F f) \\<^sub>D cmp (g, h))" proof - have "cmp (f, g \\<^sub>B h) \\<^sub>D (G (F f) \\<^sub>D cmp (g, h)) = (G (F (f \\<^sub>B g \\<^sub>B h)) \\<^sub>D G (\\<^sub>F (f, g \\<^sub>B h)) \\<^sub>D \\<^sub>G (F f, F (g \\<^sub>B h))) \\<^sub>D (G (F f) \\<^sub>D G (F (g \\<^sub>B h)) \\<^sub>D G (\\<^sub>F (g, h)) \\<^sub>D \\<^sub>G (F g, F h))" unfolding cmp_def using f g h fg gh B.VV.arr_char B.VV.dom_simp B.VV.cod_simp by simp also have "... = (G (\\<^sub>F (f, g \\<^sub>B h)) \\<^sub>D \\<^sub>G (F f, F (g \\<^sub>B h))) \\<^sub>D (G (F f) \\<^sub>D G (F (g \\<^sub>B h)) \\<^sub>D G (\\<^sub>F (g, h)) \\<^sub>D \\<^sub>G (F g, F h))" proof - have "G (F (f \\<^sub>B g \\<^sub>B h)) \\<^sub>D G (\\<^sub>F (f, g \\<^sub>B h)) = G (\\<^sub>F (f, g \\<^sub>B h))" using f g h fg gh B.VV.arr_char D.comp_ide_arr B.VV.dom_simp B.VV.cod_simp by simp thus ?thesis using D.comp_assoc by metis qed also have "... = G (\\<^sub>F (f, g \\<^sub>B h)) \\<^sub>D \\<^sub>G (F f, F (g \\<^sub>B h)) \\<^sub>D (G (F f) \\<^sub>D G (F (g \\<^sub>B h)) \\<^sub>D G (\\<^sub>F (g, h)) \\<^sub>D \\<^sub>G (F g, F h))" using D.comp_assoc by simp also have "... = G (\\<^sub>F (f, g \\<^sub>B h)) \\<^sub>D \\<^sub>G (F f, F (g \\<^sub>B h)) \\<^sub>D (G (F f) \\<^sub>D G (\\<^sub>F (g, h)) \\<^sub>D \\<^sub>G (F g, F h))" proof - have "G (F (g \\<^sub>B h)) \\<^sub>D G (\\<^sub>F (g, h)) = G (\\<^sub>F (g, h))" using f g h fg gh B.VV.arr_char D.comp_ide_arr B.VV.dom_simp B.VV.cod_simp by simp thus ?thesis using D.comp_assoc by metis qed also have "... = G (\\<^sub>F (f, g \\<^sub>B h)) \\<^sub>D \\<^sub>G (F f, F (g \\<^sub>B h)) \\<^sub>D (G (F f) \\<^sub>D G (\\<^sub>F (g, h))) \\<^sub>D (G (F f) \\<^sub>D \\<^sub>G (F g, F h))" using f g h fg gh D.whisker_left [of "G (F f)" "G (\\<^sub>F (g, h))" "\\<^sub>G (F g, F h)"] by fastforce also have "... = G (\\<^sub>F (f, g \\<^sub>B h)) \\<^sub>D (\\<^sub>G (F f, F (g \\<^sub>B h)) \\<^sub>D (G (F f) \\<^sub>D G (\\<^sub>F (g, h)))) \\<^sub>D (G (F f) \\<^sub>D \\<^sub>G (F g, F h))" using D.comp_assoc by simp also have "... = G (\\<^sub>F (f, g \\<^sub>B h)) \\<^sub>D (G (F f \\<^sub>C \\<^sub>F (g, h)) \\<^sub>D \\<^sub>G (F f, F g \\<^sub>C F h)) \\<^sub>D (G (F f) \\<^sub>D \\<^sub>G (F g, F h))" proof - have "\\<^sub>G (C.VV.cod (F f, \\<^sub>F (g, h))) = \\<^sub>G (F f, F (g \\<^sub>B h))" using f g h fg gh B.VV.arr_char C.VV.cod_char B.VV.dom_simp B.VV.cod_simp by auto moreover have "G.H\<^sub>DoFF.map (F f, \\<^sub>F (g, h)) = G (F f) \\<^sub>D G (\\<^sub>F (g, h))" using f g h fg gh B.VV.arr_char G.FF_def by auto moreover have "G.FoH\<^sub>C.map (F f, \\<^sub>F (g, h)) = G (F f \\<^sub>C \\<^sub>F (g, h))" using f g h fg gh B.VV.arr_char C.VV.arr_char by simp moreover have "C.VV.dom (F f, \\<^sub>F (g, h)) = (F f, F g \\<^sub>C F h)" using f g h fg gh B.VV.arr_char C.VV.arr_char C.VV.dom_char F.cmp_in_hom [of g h] by auto ultimately show ?thesis using f g h fg gh B.VV.arr_char C.VV.arr_char G.\.naturality [of "(F f, \\<^sub>F (g, h))"] by simp qed also have "... = (G (\\<^sub>F (f, g \\<^sub>B h)) \\<^sub>D (G (F f \\<^sub>C \\<^sub>F (g, h)))) \\<^sub>D \\<^sub>G (F f, F g \\<^sub>C F h) \\<^sub>D (G (F f) \\<^sub>D \\<^sub>G (F g, F h))" using D.comp_assoc by simp also have "... = G (\\<^sub>F (f, g \\<^sub>B h) \\<^sub>C (F f \\<^sub>C \\<^sub>F (g, h))) \\<^sub>D \\<^sub>G (F f, F g \\<^sub>C F h) \\<^sub>D (G (F f) \\<^sub>D \\<^sub>G (F g, F h))" using f g h fg gh B.VV.arr_char by (metis (no_types, lifting) B.assoc_simps(1) C.comp_assoc C.seqE F.preserves_assoc(1) F.preserves_reflects_arr G.preserves_comp) finally show ?thesis by simp qed thus ?thesis by simp qed also have "... = cmp (f, g \\<^sub>B h) \\<^sub>D (G (F f) \\<^sub>D cmp (g, h)) \\<^sub>D \\<^sub>D[G (F f), G (F g), G (F h)]" using D.comp_assoc by simp finally show ?thesis by simp qed qed lemma is_pseudofunctor: shows "pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (G o F) cmp" .. lemma map\<^sub>0_simp [simp]: assumes "B.obj a" shows "map\<^sub>0 a = G.map\<^sub>0 (F.map\<^sub>0 a)" using assms map\<^sub>0_def by auto lemma unit_char': assumes "B.obj a" shows "unit a = G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a)" proof - have "G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a) = unit a" proof (intro unit_eqI [of a "G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a)"]) show "B.obj a" by fact show "\G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a) : map\<^sub>0 a \\<^sub>D map a\" using assms by auto show "D.iso (G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a))" by (simp add: D.isos_compose F.unit_char(2) G.unit_char(2) assms) show "(G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a)) \\<^sub>D \\<^sub>D[map\<^sub>0 a] = (map \\<^sub>B[a] \\<^sub>D cmp (a, a)) \\<^sub>D (G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a) \\<^sub>D G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a))" proof - have 1: "cmp (a, a) = G (\\<^sub>F (a, a)) \\<^sub>D \\<^sub>G (F a, F a)" - proof - + proof - have "cmp (a, a) = (G (F (a \\<^sub>B a)) \\<^sub>D G (\\<^sub>F (a, a))) \\<^sub>D \\<^sub>G (F a, F a)" using assms cmp_def B.VV.ide_char B.VV.arr_char B.VV.dom_char B.VV.cod_char B.VxV.dom_char B.objE D.comp_assoc B.obj_simps by simp also have "... = G (\\<^sub>F (a, a)) \\<^sub>D \\<^sub>G (F a, F a)" using assms D.comp_cod_arr B.VV.arr_char B.VV.ide_char by auto finally show ?thesis by blast qed have "(map \\<^sub>B[a] \\<^sub>D cmp (a, a)) \\<^sub>D (G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a) \\<^sub>D G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a)) = map \\<^sub>B[a] \\<^sub>D G (\\<^sub>F (a, a)) \\<^sub>D (\\<^sub>G (F a, F a) \\<^sub>D (G (F.unit a) \\<^sub>D G (F.unit a))) \\<^sub>D (G.unit (F.map\<^sub>0 a) \\<^sub>D G.unit (F.map\<^sub>0 a))" using assms 1 D.comp_assoc D.interchange by simp also have "... = (map \\<^sub>B[a] \\<^sub>D G (\\<^sub>F (a, a)) \\<^sub>D G (F.unit a \\<^sub>C F.unit a)) \\<^sub>D \\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a) \\<^sub>D (G.unit (F.map\<^sub>0 a) \\<^sub>D G.unit (F.map\<^sub>0 a))" using assms G.\.naturality [of "(F.unit a, F.unit a)"] C.VV.arr_char C.VV.dom_char C.VV.cod_char G.FF_def D.comp_assoc by simp also have "... = (G (F \\<^sub>B[a] \\<^sub>C \\<^sub>F (a, a) \\<^sub>C (F.unit a \\<^sub>C F.unit a))) \\<^sub>D \\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a) \\<^sub>D (G.unit (F.map\<^sub>0 a) \\<^sub>D G.unit (F.map\<^sub>0 a))" proof - have "C.arr (F \\<^sub>B[a] \\<^sub>C \\<^sub>F (a, a) \\<^sub>C (F.unit a \\<^sub>C F.unit a))" using assms B.VV.ide_char B.VV.arr_char F.cmp_in_hom(2) by (intro C.seqI' C.comp_in_homI) auto hence "map \\<^sub>B[a] \\<^sub>D G (\\<^sub>F (a, a)) \\<^sub>D G (F.unit a \\<^sub>C F.unit a) = G (F \\<^sub>B[a] \\<^sub>C \\<^sub>F (a, a) \\<^sub>C (F.unit a \\<^sub>C F.unit a))" by auto thus ?thesis by argo qed also have "... = G (F.unit a \\<^sub>C \\<^sub>C[F.map\<^sub>0 a]) \\<^sub>D \\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a) \\<^sub>D (G.unit (F.map\<^sub>0 a) \\<^sub>D G.unit (F.map\<^sub>0 a))" using assms F.unit_char C.comp_assoc by simp also have "... = G (F.unit a) \\<^sub>D (G \\<^sub>C[F.map\<^sub>0 a] \\<^sub>D \\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a)) \\<^sub>D (G.unit (F.map\<^sub>0 a) \\<^sub>D G.unit (F.map\<^sub>0 a))" using assms D.comp_assoc by simp also have "... = (G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a)) \\<^sub>D \\<^sub>D[G.map\<^sub>0 (F.map\<^sub>0 a)]" using assms G.unit_char D.comp_assoc by simp also have "... = (G (F.unit a) \\<^sub>D G.unit (F.map\<^sub>0 a)) \\<^sub>D \\<^sub>D[map\<^sub>0 a]" using assms map\<^sub>0_def by auto finally show ?thesis by simp qed qed thus ?thesis by simp qed end subsection "Restriction of Pseudofunctors" text \ In this section, we construct the restriction and corestriction of a pseudofunctor to a subbicategory of its domain and codomain, respectively. \ locale restricted_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: pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D F \ + C': subbicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C Arr 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 \ :: "'c * 'c \ 'd" and Arr :: "'c \ bool" begin abbreviation map where "map \ \\. if C'.arr \ then F \ else D.null" abbreviation cmp where "cmp \ \\\. if C'.VV.arr \\ then \ \\ else D.null" interpretation "functor" C'.comp V\<^sub>D map using C'.inclusion C'.arr_char C'.dom_char C'.cod_char C'.seq_char C'.comp_char C'.arr_dom C'.arr_cod apply (unfold_locales) apply auto by presburger interpretation weak_arrow_of_homs C'.comp C'.src C'.trg V\<^sub>D src\<^sub>D trg\<^sub>D map using C'.arrE C'.ide_src C'.ide_trg C'.inclusion C'.src_def C'.trg_def F.weakly_preserves_src F.weakly_preserves_trg by unfold_locales auto interpretation H\<^sub>D\<^sub>'oFF: composite_functor C'.VV.comp D.VV.comp V\<^sub>D FF \\\\. fst \\ \\<^sub>D snd \\\ .. interpretation FoH\<^sub>C\<^sub>': composite_functor C'.VV.comp C'.comp V\<^sub>D \\\\. C'.hcomp (fst \\) (snd \\)\ map .. interpretation \: natural_transformation C'.VV.comp V\<^sub>D H\<^sub>D\<^sub>'oFF.map FoH\<^sub>C\<^sub>'.map cmp using C'.arr_char C'.dom_char C'.cod_char C'.VV.arr_char C'.VV.dom_char C'.VV.cod_char FF_def C'.inclusion C'.dom_closed C'.cod_closed C'.src_def C'.trg_def - F.cmp_simps'(4) C'.hcomp_def C'.hcomp_closed F.\.is_natural_1 F.\.is_natural_2 - C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def F.cmp_simps'(5) + C'.hcomp_def C'.hcomp_closed F.\.is_natural_1 F.\.is_natural_2 + C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def by unfold_locales auto interpretation \: natural_isomorphism C'.VV.comp V\<^sub>D H\<^sub>D\<^sub>'oFF.map FoH\<^sub>C\<^sub>'.map cmp using C.VV.ide_char C.VV.arr_char C'.VV.ide_char C'.VV.arr_char C'.arr_char C'.src_def C'.trg_def C'.ide_char F.\.components_are_iso by unfold_locales auto sublocale pseudofunctor C'.comp C'.hcomp C'.\ \\<^sub>C C'.src C'.trg V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D map cmp using F.assoc_coherence C'.VVV.arr_char C'.VV.arr_char C'.arr_char C'.hcomp_def C'.src_def C'.trg_def C'.assoc_closed C'.hcomp_closed C'.ide_char by unfold_locales (simp add: C'.ide_char C'.src_def C'.trg_def) lemma is_pseudofunctor: shows "pseudofunctor C'.comp C'.hcomp C'.\ \\<^sub>C C'.src C'.trg V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D map cmp" .. lemma map\<^sub>0_simp [simp]: assumes "C'.obj a" shows "map\<^sub>0 a = F.map\<^sub>0 a" using assms map\<^sub>0_def C'.obj_char by auto lemma unit_char': assumes "C'.obj a" shows "F.unit a = unit a" - proof (intro unit_eqI) - show "C'.obj a" by fact - show "\F.unit a : map\<^sub>0 a \\<^sub>D map a\" - using assms map\<^sub>0_def F.unit_in_hom(2) [of a] C'.obj_char by auto - show "D.iso (F.unit a)" - using assms by (simp add: C'.obj_char F.unit_char(2)) - show "F.unit a \\<^sub>D \\<^sub>D[map\<^sub>0 a] = (map \\<^sub>C[a] \\<^sub>D cmp (a, a)) \\<^sub>D (F.unit a \\<^sub>D F.unit a)" - proof - - have "F.unit a \\<^sub>D \\<^sub>D[map\<^sub>0 a] = F.unit a \\<^sub>D \\<^sub>D[F.map\<^sub>0 a]" - using assms map\<^sub>0_def F.map\<^sub>0_def by auto - also have "... = (F \\<^sub>C[a] \\<^sub>D \ (a, a)) \\<^sub>D (F.unit a \\<^sub>D F.unit a)" - using assms C'.obj_char F.unit_char(3) [of a] by simp - also have "... = (map \\<^sub>C[a] \\<^sub>D cmp (a, a)) \\<^sub>D (F.unit a \\<^sub>D F.unit a)" - using assms C.VV.arr_char D.iso_is_arr iso_\ by auto - finally show ?thesis by simp - qed - qed + using assms map\<^sub>0_simp C'.obj_char F.unit_in_hom(2) [of a] F.unit_char(2-3) \_simps(1) + apply (intro unit_eqI) + apply auto + by blast end text \ We define the corestriction construction only for the case of sub-bicategories determined by a set of objects of the ambient bicategory. There are undoubtedly more general constructions, but this one is adequate for our present needs. \ locale corestricted_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: 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 \ + D': subbicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \\\. D.arr \ \ Obj (src\<^sub>D \) \ Obj (trg\<^sub>D \)\ 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 \ :: "'c * 'c \ 'd" and Obj :: "'d \ bool" + assumes preserves_arr: "\\. C.arr \ \ D'.arr (F \)" begin abbreviation map where "map \ F" abbreviation cmp where "cmp \ \" interpretation "functor" V\<^sub>C D'.comp F using preserves_arr F.is_extensional D'.arr_char D'.dom_char D'.cod_char D'.comp_char by (unfold_locales) auto interpretation weak_arrow_of_homs V\<^sub>C src\<^sub>C trg\<^sub>C D'.comp D'.src D'.trg F proof fix \ assume \: "C.arr \" obtain \ where \: "\\ : F (src\<^sub>C \) \\<^sub>D src\<^sub>D (F \)\ \ D.iso \" using \ F.weakly_preserves_src by auto have 2: "D'.in_hom \ (F (src\<^sub>C \)) (D'.src (F \))" using \ \ D'.arr_char D'.dom_char D'.cod_char D'.src_def D.vconn_implies_hpar(1-2) preserves_arr - by (intro D'.in_homI) auto + by (metis (no_types, lifting) C.src.preserves_arr D'.in_hom_char D'.src.preserves_arr + D.arrI) moreover have "D'.iso \" using 2 \ D'.iso_char D'.arr_char by fastforce ultimately show "D'.isomorphic (F (src\<^sub>C \)) (D'.src (F \))" using D'.isomorphic_def by auto obtain \ where \: "\\ : F (trg\<^sub>C \) \\<^sub>D trg\<^sub>D (F \)\ \ D.iso \" using \ F.weakly_preserves_trg by auto have 2: "D'.in_hom \ (F (trg\<^sub>C \)) (D'.trg (F \))" using \ \ D'.arr_char D'.dom_char D'.cod_char D'.trg_def D.vconn_implies_hpar(1-2) preserves_arr - by (intro D'.in_homI) auto + by (metis (no_types, lifting) C.trg.preserves_arr D'.in_hom_char D'.trg.preserves_arr + D.arrI) moreover have "D'.iso \" using 2 \ D'.iso_char D'.arr_char by fastforce ultimately show "D'.isomorphic (F (trg\<^sub>C \)) (D'.trg (F \))" using D'.isomorphic_def by auto qed interpretation H\<^sub>D\<^sub>'oFF: composite_functor C.VV.comp D'.VV.comp D'.comp FF \\\\. D'.hcomp (fst \\) (snd \\)\ .. interpretation FoH\<^sub>C: composite_functor C.VV.comp V\<^sub>C D'.comp \\\\. fst \\ \\<^sub>C snd \\\ F .. interpretation natural_transformation C.VV.comp D'.comp H\<^sub>D\<^sub>'oFF.map FoH\<^sub>C.map \ proof show "\\\. \ C.VV.arr \\ \ \ \\ = D'.null" by (simp add: F.\.is_extensional) fix \\ assume \\: "C.VV.arr \\" have 1: "D'.arr (\ \\)" using \\ D'.arr_char F.\.is_natural_1 F.\.components_are_iso by (metis (no_types, lifting) D.src_vcomp D.trg_vcomp FoH\<^sub>C.preserves_arr F.\.preserves_reflects_arr) show "D'.dom (\ \\) = H\<^sub>D\<^sub>'oFF.map (C.VV.dom \\)" using 1 \\ D'.dom_char C.VV.arr_char C.VV.dom_char F.FF_def FF_def D'.hcomp_def by simp show "D'.cod (\ \\) = FoH\<^sub>C.map (C.VV.cod \\)" using 1 \\ D'.cod_char C.VV.arr_char F.FF_def FF_def D'.hcomp_def by simp show "D'.comp (FoH\<^sub>C.map \\) (\ (C.VV.dom \\)) = \ \\" using 1 \\ D'.arr_char D'.comp_char C.VV.dom_char F.\.is_natural_1 C.VV.arr_dom D.src_vcomp D.trg_vcomp FoH\<^sub>C.preserves_arr F.\.preserves_reflects_arr by (metis (mono_tags, lifting)) show "D'.comp (\ (C.VV.cod \\)) (H\<^sub>D\<^sub>'oFF.map \\) = \ \\" proof - have "Obj (F.map\<^sub>0 (trg\<^sub>C (fst \\))) \ Obj (F.map\<^sub>0 (trg\<^sub>C (snd \\)))" using \\ C.VV.arr_char by (metis (no_types, lifting) C.src_trg C.trg.preserves_reflects_arr D'.arr_char F.map\<^sub>0_def preserves_hseq) moreover have "Obj (F.map\<^sub>0 (src\<^sub>C (snd \\)))" using \\ C.VV.arr_char by (metis (no_types, lifting) C.src.preserves_reflects_arr C.trg_src D'.arr_char F.map\<^sub>0_def preserves_hseq) ultimately show ?thesis using \\ 1 D'.arr_char D'.comp_char D'.hseq_char C.VV.arr_char C.VV.cod_char C.VxV.cod_char FF_def F.FF_def D'.hcomp_char preserves_hseq apply simp using F.\.is_natural_2 by force qed qed interpretation natural_isomorphism C.VV.comp D'.comp H\<^sub>D\<^sub>'oFF.map FoH\<^sub>C.map \ apply unfold_locales using D'.iso_char D'.arr_char by fastforce sublocale pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C D'.comp D'.hcomp D'.\ \\<^sub>D D'.src D'.trg F \ proof fix f g h assume f: "C.ide f" and g: "C.ide g" and h: "C.ide h" and fg: "src\<^sub>C f = trg\<^sub>C g" and gh: "src\<^sub>C g = trg\<^sub>C h" have "D'.comp (F \\<^sub>C[f, g, h]) (D'.comp (\ (f \\<^sub>C g, h)) (D'.hcomp (\ (f, g)) (F h))) = F \\<^sub>C[f, g, h] \\<^sub>D \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h)" proof - have 1: "D'.arr (cmp (f, g) \\<^sub>D map h)" - using f g h fg gh D'.arr_char C.VV.arr_char - by (metis (no_types, lifting) C.ideD(1) D'.hcomp_closed F.\.preserves_reflects_arr - F.cmp_simps(1) F.cmp_simps(2) F.preserves_hseq preserves_reflects_arr) + by (metis (mono_tags, lifting) C.ideD(1) D'.arr_char D'.hcomp_closed + F.\.preserves_reflects_arr F.cmp_simps(1-2) F.preserves_hseq + f fg g gh h preserves_reflects_arr) moreover have 2: "D.seq (cmp (f \\<^sub>C g, h)) (cmp (f, g) \\<^sub>D map h)" using 1 f g h fg gh D'.arr_char C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def by (intro D.seqI) auto moreover have "D'.arr (cmp (f \\<^sub>C g, h) \\<^sub>D (cmp (f, g) \\<^sub>D map h))" using 1 2 D'.arr_char by (metis (no_types, lifting) D'.comp_char D'.seq_char D.seqE F.\.preserves_reflects_arr preserves_reflects_arr) ultimately show ?thesis using f g h fg gh D'.dom_char D'.cod_char D'.comp_char D'.hcomp_def C.VV.arr_char apply simp by force qed also have "... = \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h]" using f g h fg gh F.assoc_coherence [of f g h] by blast also have "... = D'.comp (\ (f, g \\<^sub>C h)) (D'.comp (D'.hcomp (F f) (\ (g, h))) (D'.\ (F f) (F g) (F h)))" proof - have "D.seq (map f \\<^sub>D cmp (g, h)) \\<^sub>D[map f, map g, map h]" using f g h fg gh C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def by (intro D.seqI) auto moreover have "D'.arr \\<^sub>D[map f, map g, map h]" using f g h fg gh D'.arr_char preserves_arr by auto moreover have "D'.arr (map f \\<^sub>D cmp (g, h))" using f g h fg gh by (metis (no_types, lifting) D'.arr_char D.seqE D.vseq_implies_hpar(1) D.vseq_implies_hpar(2) calculation(1-2)) moreover have "D'.arr ((map f \\<^sub>D cmp (g, h)) \\<^sub>D \\<^sub>D[map f, map g, map h])" using f g h fg gh by (metis (no_types, lifting) D'.arr_char D'.comp_closed D.seqE calculation(1-3)) moreover have "D.seq (cmp (f, g \\<^sub>C h)) ((map f \\<^sub>D cmp (g, h)) \\<^sub>D \\<^sub>D[map f, map g, map h])" using f g h fg gh F.cmp_simps'(1) F.cmp_simps(4) F.cmp_simps(5) by auto ultimately show ?thesis using f g h fg gh C.VV.arr_char D'.VVV.arr_char D'.VV.arr_char D'.comp_char D'.hcomp_def by simp qed finally show "D'.comp (F \\<^sub>C[f, g, h]) (D'.comp (\ (f \\<^sub>C g, h)) (D'.hcomp (\ (f, g)) (F h))) = D'.comp (\ (f, g \\<^sub>C h)) (D'.comp (D'.hcomp (F f) (\ (g, h))) (D'.\ (F f) (F g) (F h)))" by blast qed lemma is_pseudofunctor: shows "pseudofunctor V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C D'.comp D'.hcomp D'.\ \\<^sub>D D'.src D'.trg F \" .. lemma map\<^sub>0_simp [simp]: assumes "C.obj a" shows "map\<^sub>0 a = F.map\<^sub>0 a" using assms map\<^sub>0_def D'.src_def by auto lemma unit_char': assumes "C.obj a" shows "F.unit a = unit a" proof (intro unit_eqI) show "C.obj a" by fact show 1: "D'.in_hom (F.unit a) (map\<^sub>0 a) (map a)" - proof - - have "D.in_hom (F.unit a) (F.map\<^sub>0 a) (map a)" - using assms by auto - moreover have "map\<^sub>0 a = map\<^sub>0 a" - using assms map\<^sub>0_def D'.src_def by auto - ultimately show ?thesis - using assms D'.in_hom_char preserves_arr D'.arr_char - by (metis (no_types, lifting) C.obj_def' D'.obj_def D'.src_def D.arrI - D.vconn_implies_hpar(1,3) F.unit_simps(2-3) map\<^sub>0_def map\<^sub>0_simps(1)) - qed + using D'.arr_char D'.in_hom_char assms unit_in_hom(2) by force show "D'.iso (F.unit a)" using assms D'.iso_char D'.arr_char F.unit_char(2) \D'.in_hom (F.unit a) (map\<^sub>0 a) (map a)\ by auto show "D'.comp (F.unit a) \\<^sub>D[map\<^sub>0 a] = D'.comp (D'.comp (map \\<^sub>C[a]) (cmp (a, a))) (D'.hcomp (F.unit a) (F.unit a))" proof - have "D'.comp (F.unit a) \\<^sub>D[map\<^sub>0 a] = F.unit a \\<^sub>D \\<^sub>D[src\<^sub>D (map a)]" using assms D'.comp_char D'.arr_char apply simp by (metis (no_types, lifting) C.obj_simps(1-2) F.preserves_src preserves_arr) also have "... = (map \\<^sub>C[a] \\<^sub>D cmp (a, a)) \\<^sub>D (F.unit a \\<^sub>D F.unit a)" using assms F.unit_char(3) [of a] by auto also have "... = D'.comp (D'.comp (map \\<^sub>C[a]) (cmp (a, a))) (D'.hcomp (F.unit a) (F.unit a))" proof - have "D'.arr (map \\<^sub>C[a] \\<^sub>D cmp (a, a))" using assms D'.comp_simp by auto moreover have "D.seq (map \\<^sub>C[a] \\<^sub>D cmp (a, a)) (F.unit a \\<^sub>D F.unit a)" using assms C.VV.arr_char F.cmp_simps(4-5) by (intro D.seqI) auto ultimately show ?thesis - using assms 1 D'.comp_char D'.hcomp_def C.VV.arr_char D'.hseq_char - D'.arr_char F.\_simps(2) - by auto + by (metis (no_types, lifting) "1" D'.comp_eqI' D'.hcomp_eqI' D'.hseqI' + D'.iso_is_arr D'.seq_char D'.vconn_implies_hpar(1-2) + \_simps(1) \D'.iso (F.unit a)\ assms map\<^sub>0_simps(2-3)) qed finally show ?thesis by blast qed qed end subsection "Equivalence Pseudofunctors" text \ In this section, we define ``equivalence pseudofunctors'', which are pseudofunctors that are locally fully faithful, locally essentially surjective, and biessentially surjective on objects. In a later section, we will show that a pseudofunctor is an equivalence pseudofunctor if and only if it can be extended to an equivalence of bicategories. The definition below requires that an equivalence pseudofunctor be (globally) faithful with respect to vertical composition. Traditional formulations do not consider a pseudofunctor as a single global functor, so we have to consider whether this condition is too strong. In fact, a pseudofunctor (as defined here) is locally faithful if and only if it is globally faithful. \ context pseudofunctor begin definition locally_faithful where "locally_faithful \ \f g \ \'. \\ : f \\<^sub>C g\ \ \\' : f \\<^sub>C g\ \ F \ = F \' \ \ = \'" lemma locally_faithful_iff_faithful: shows "locally_faithful \ faithful_functor V\<^sub>C V\<^sub>D F" proof show "faithful_functor V\<^sub>C V\<^sub>D F \ locally_faithful" - proof - - assume 1: "faithful_functor V\<^sub>C V\<^sub>D F" - interpret faithful_functor V\<^sub>C V\<^sub>D F - using 1 by blast - show "locally_faithful" - unfolding locally_faithful_def using is_faithful by blast - qed + by (metis category.in_homE faithful_functor.is_faithful functor_axioms + functor_def locally_faithful_def) show "locally_faithful \ faithful_functor V\<^sub>C V\<^sub>D F" proof - assume 1: "locally_faithful" show "faithful_functor V\<^sub>C V\<^sub>D F" proof fix \ \' assume par: "C.par \ \'" and eq: "F \ = F \'" obtain f g where fg: "\\ : f \\<^sub>C g\ \ \\' : f \\<^sub>C g\" using par by auto show "\ = \'" using 1 fg locally_faithful_def eq by simp qed qed qed end text \ In contrast, it is not true that a pseudofunctor that is locally full is also globally full, because we can have \\\ : F h \\<^sub>D F k\\ even if \h\ and \k\ are not in the same hom-category. So it would be a mistake to require that an equivalence functor be globally full. \ locale equivalence_pseudofunctor = pseudofunctor + faithful_functor V\<^sub>C V\<^sub>D F + assumes biessentially_surjective_on_objects: "D.obj a' \ \a. C.obj a \ D.equivalent_objects (map\<^sub>0 a) a'" and locally_essentially_surjective: "\ C.obj a; C.obj b; \g : map\<^sub>0 a \\<^sub>D map\<^sub>0 b\; D.ide g \ \ \f. \f : a \\<^sub>C b\ \ C.ide f \ D.isomorphic (F f) g" and locally_full: "\ C.ide f; C.ide f'; src\<^sub>C f = src\<^sub>C f'; trg\<^sub>C f = trg\<^sub>C f'; \\ : F f \\<^sub>D F f'\ \ \ \\. \\ : f \\<^sub>C f'\ \ F \ = \" begin lemma reflects_ide: assumes "C.endo \" and "D.ide (F \)" shows "C.ide \" using assms is_faithful [of "C.dom \" \] C.ide_char' by (metis C.arr_dom C.cod_dom C.dom_dom C.seqE D.ide_char preserves_dom) lemma reflects_iso: assumes "C.arr \" and "D.iso (F \)" shows "C.iso \" proof - obtain \' where \': "\\' : C.cod \ \\<^sub>C C.dom \\ \ F \' = D.inv (F \)" using assms locally_full [of "C.cod \" "C.dom \" "D.inv (F \)"] D.inv_in_hom C.in_homE preserves_hom C.in_homI by auto have "C.inverse_arrows \ \'" using assms \' reflects_ide apply (intro C.inverse_arrowsI) apply (metis C.cod_comp C.dom_comp C.ide_dom C.in_homE C.seqI D.comp_inv_arr' faithful_functor_axioms faithful_functor_def functor.preserves_ide as_nat_trans.preserves_comp_2 preserves_dom) by (metis C.cod_comp C.dom_comp C.ide_cod C.in_homE C.seqI D.comp_arr_inv' faithful_functor_axioms faithful_functor_def functor.preserves_ide preserves_cod as_nat_trans.preserves_comp_2) thus ?thesis by auto qed lemma reflects_isomorphic: assumes "C.ide f" and "C.ide f'" and "src\<^sub>C f = src\<^sub>C f'" and "trg\<^sub>C f = trg\<^sub>C f'" and "D.isomorphic (F f) (F f')" shows "C.isomorphic f f'" using assms C.isomorphic_def D.isomorphic_def locally_full reflects_iso C.arrI by metis lemma reflects_equivalence: assumes "C.ide f" and "C.ide g" and "\\ : src\<^sub>C f \\<^sub>C g \\<^sub>C f\" and "\\ : f \\<^sub>C g \\<^sub>C src\<^sub>C g\" and "equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D unit (src\<^sub>C f)) (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g))" shows "equivalence_in_bicategory V\<^sub>C H\<^sub>C \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \" proof - interpret E': equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \F f\ \F g\ \D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D unit (src\<^sub>C f)\ \D.inv (unit (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g)\ using assms by auto show ?thesis proof show "C.ide f" using assms(1) by simp show "C.ide g" using assms(2) by simp show "\\ : src\<^sub>C f \\<^sub>C g \\<^sub>C f\" using assms(3) by simp show "\\ : f \\<^sub>C g \\<^sub>C src\<^sub>C g\" using assms(4) by simp have 0: "src\<^sub>C f = trg\<^sub>C g \ src\<^sub>C g = trg\<^sub>C f" using \\\ : src\<^sub>C f \\<^sub>C g \\<^sub>C f\\ by (metis C.hseqE C.ideD(1) C.ide_cod C.ide_dom C.in_homE assms(4)) show "C.iso \" proof - have "D.iso (F \)" proof - have 1: "\D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D unit (src\<^sub>C f) : map\<^sub>0 (src\<^sub>C f) \\<^sub>D F g \\<^sub>D F f\" using \C.ide f\ \C.ide g\ \\\ : src\<^sub>C f \\<^sub>C g \\<^sub>C f\\ unit_char cmp_in_hom cmp_components_are_iso by (metis (mono_tags, lifting) C.ideD(1) E'.unit_in_vhom preserves_src) have 2: "D.iso (\ (g, f)) \ \\ (g, f) : F g \\<^sub>D F f \\<^sub>D F (g \\<^sub>C f)\" using 0 \C.ide f\ \C.ide g\ cmp_in_hom by simp have 3: "D.iso (D.inv (unit (src\<^sub>C f))) \ \D.inv (unit (src\<^sub>C f)) : F (src\<^sub>C f) \\<^sub>D map\<^sub>0 (src\<^sub>C f)\" using \C.ide f\ unit_char by simp have "D.iso (\ (g, f) \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D unit (src\<^sub>C f)) \\<^sub>D D.inv (unit (src\<^sub>C f)))" using 1 2 3 E'.unit_is_iso D.isos_compose by blast moreover have "\ (g, f) \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D unit (src\<^sub>C f)) \\<^sub>D D.inv (unit (src\<^sub>C f)) = F \" proof - have "\ (g, f) \\<^sub>D (D.inv (\ (g, f)) \\<^sub>D F \ \\<^sub>D unit (src\<^sub>C f)) \\<^sub>D D.inv (unit (src\<^sub>C f)) = (\ (g, f) \\<^sub>D (D.inv (\ (g, f))) \\<^sub>D F \ \\<^sub>D (unit (src\<^sub>C f)) \\<^sub>D D.inv (unit (src\<^sub>C f)))" using D.comp_assoc by simp also have "... = F (g \\<^sub>C f) \\<^sub>D F \ \\<^sub>D F (src\<^sub>C f)" using 2 3 D.comp_arr_inv D.comp_inv_arr D.inv_is_inverse by (metis C.ideD(1) C.obj_src D.comp_assoc D.dom_inv D.in_homE unit_char(2) assms(1)) also have "... = F \" using \\\ : src\<^sub>C f \\<^sub>C g \\<^sub>C f\\ D.comp_arr_dom D.comp_cod_arr by auto finally show ?thesis by simp qed ultimately show ?thesis by simp qed thus ?thesis using assms reflects_iso by auto qed show "C.iso \" proof - have "D.iso (F \)" proof - have 1: "\D.inv (unit (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g) : F f \\<^sub>D F g \\<^sub>D map\<^sub>0 (src\<^sub>C g)\" using \C.ide f\ \C.ide g\ \\\ : f \\<^sub>C g \\<^sub>C src\<^sub>C g\\ unit_char cmp_in_hom cmp_components_are_iso by (metis (mono_tags, lifting) C.ideD(1) E'.counit_in_vhom preserves_src) have 2: "D.iso (D.inv (\ (f, g))) \ \D.inv (\ (f, g)) : F (f \\<^sub>C g) \\<^sub>D F f \\<^sub>D F g\" using 0 \C.ide f\ \C.ide g\ \\\ : f \\<^sub>C g \\<^sub>C src\<^sub>C g\\ cmp_in_hom(2) D.inv_in_hom by simp have 3: "D.iso (unit (trg\<^sub>C f)) \ \unit (trg\<^sub>C f) : map\<^sub>0 (trg\<^sub>C f) \\<^sub>D F (trg\<^sub>C f)\" using \C.ide f\ unit_char by simp have "D.iso (unit (trg\<^sub>C f) \\<^sub>D (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g)) \\<^sub>D D.inv (\ (f, g)))" using 0 1 2 3 E'.counit_is_iso D.isos_compose by (metis D.arrI D.cod_comp D.cod_inv D.seqI D.seqI') moreover have "unit (trg\<^sub>C f) \\<^sub>D (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g)) \\<^sub>D D.inv (\ (f, g)) = F \" proof - have "unit (trg\<^sub>C f) \\<^sub>D (D.inv (unit (trg\<^sub>C f)) \\<^sub>D F \ \\<^sub>D \ (f, g)) \\<^sub>D D.inv (\ (f, g)) = (unit (trg\<^sub>C f) \\<^sub>D D.inv (unit (trg\<^sub>C f))) \\<^sub>D F \ \\<^sub>D (\ (f, g) \\<^sub>D D.inv (\ (f, g)))" using D.comp_assoc by simp also have "... = F (trg\<^sub>C f) \\<^sub>D F \ \\<^sub>D F (f \\<^sub>C g)" using 0 3 D.comp_arr_inv' D.comp_inv_arr' by (simp add: C.VV.arr_char C.VV.ide_char assms(1-2)) also have "... = F \" using 0 \\\ : f \\<^sub>C g \\<^sub>C src\<^sub>C g\\ D.comp_arr_dom D.comp_cod_arr by auto finally show ?thesis by simp qed ultimately show ?thesis by simp qed thus ?thesis using assms reflects_iso by auto qed qed qed lemma reflects_equivalence_map: assumes "C.ide f" and "D.equivalence_map (F f)" shows "C.equivalence_map f" proof - obtain g' \ \ where E': "equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (F f) g' \ \" using assms D.equivalence_map_def by auto interpret E': equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \F f\ g' \ \ using E' by auto obtain g where g: "\g : trg\<^sub>C f \\<^sub>C src\<^sub>C f\ \ C.ide g \ D.isomorphic (F g) g'" using assms E'.antipar locally_essentially_surjective [of "trg\<^sub>C f" "src\<^sub>C f" g'] by auto obtain \ where \: "\\ : g' \\<^sub>D F g\ \ D.iso \" using g D.isomorphic_def D.isomorphic_symmetric by blast - have E'': "equivalence_in_bicategory (\\<^sub>D) (\\<^sub>D) \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) - ((F g \\<^sub>D F f) \\<^sub>D (\ \\<^sub>D F f) \\<^sub>D \) - (\ \\<^sub>D (D.inv (F f) \\<^sub>D g') \\<^sub>D (F f \\<^sub>D D.inv \))" + interpret E'': equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \F f\ \F g\ + \(F g \\<^sub>D F f) \\<^sub>D (\ \\<^sub>D F f) \\<^sub>D \\ + \\ \\<^sub>D (D.inv (F f) \\<^sub>D g') \\<^sub>D (F f \\<^sub>D D.inv \)\ using assms \ E'.equivalence_in_bicategory_axioms D.ide_is_iso D.equivalence_respects_iso [of "F f" g' \ \ "F f" "F f" \ "F g"] by auto - interpret E'': equivalence_in_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \F f\ \F g\ - \(F g \\<^sub>D F f) \\<^sub>D (\ \\<^sub>D F f) \\<^sub>D \\ - \\ \\<^sub>D (D.inv (F f) \\<^sub>D g') \\<^sub>D (F f \\<^sub>D D.inv \)\ - using E'' by auto let ?\' = "\ (g, f) \\<^sub>D (F g \\<^sub>D F f) \\<^sub>D (\ \\<^sub>D F f) \\<^sub>D \ \\<^sub>D D.inv (unit (src\<^sub>C f))" have \': "\?\' : F (src\<^sub>C f) \\<^sub>D F (g \\<^sub>C f)\" using assms \ g unit_char E'.unit_in_hom(2) E'.antipar E''.antipar cmp_in_hom apply (intro D.comp_in_homI) apply auto using E'.antipar(2) by blast have iso_\': "D.iso ?\'" - using assms g \ \' E'.antipar E''.antipar cmp_in_hom unit_char - E'.unit_in_hom E'.unit_is_iso - by (intro D.isos_compose) auto + using assms g \ \' E'.antipar unit_char + by (metis C.in_hhomE D.arrI D.inv_comp_left(2) D.inv_comp_right(2) D.iso_hcomp + D.iso_inv_iso D.isos_compose D.seqE E''.antipar(2) E''.unit_is_iso + E'.unit_is_iso as_nat_iso.components_are_iso cmp_components_are_iso) let ?\' = "unit (src\<^sub>C g) \\<^sub>D \ \\<^sub>D (D.inv (F f) \\<^sub>D g') \\<^sub>D (F f \\<^sub>D D.inv \) \\<^sub>D D.inv (\ (f, g))" have \': "\?\' : F (f \\<^sub>C g) \\<^sub>D F (trg\<^sub>C f)\" proof (intro D.comp_in_homI) show "\D.inv (\ (f, g)) : F (f \\<^sub>C g) \\<^sub>D F f \\<^sub>D F g\" using assms g cmp_in_hom C.VV.ide_char C.VV.arr_char by auto show "\F f \\<^sub>D D.inv \ : F f \\<^sub>D F g \\<^sub>D F f \\<^sub>D g'\" - using assms g \ E'.antipar E''.antipar D.ide_in_hom(2) by auto + using assms g \ E''.antipar D.ide_in_hom(2) by auto show "\D.inv (F f) \\<^sub>D g' : F f \\<^sub>D g' \\<^sub>D F f \\<^sub>D g'\" - using assms E'.antipar E''.antipar D.ide_is_iso - by (simp add: D.ide_in_hom(2)) + using assms E'.antipar D.ide_is_iso by auto show "\\ : F f \\<^sub>D g' \\<^sub>D trg\<^sub>D (F f)\" using E'.counit_in_hom by simp show "\unit (src\<^sub>C g) : trg\<^sub>D (F f) \\<^sub>D F (trg\<^sub>C f)\" - using assms g unit_char E'.antipar by auto + using assms g unit_char by auto qed have iso_\': "D.iso ?\'" proof - have "D.iso (\ (f, g))" using assms g C.VV.ide_char C.VV.arr_char by auto thus ?thesis - using assms g \ E'.antipar E''.antipar cmp_in_hom unit_char - E'.counit_in_hom D.iso_inv_iso E'.counit_is_iso \' - by (metis C.ideD(1) C.obj_src D.arrI D.iso_hcomp D.hseq_char D.ide_is_iso - D.isos_compose D.seqE E'.ide_right as_nat_iso.components_are_iso) + by (metis C.in_hhomE D.arrI D.hseq_char' D.ide_is_iso D.inv_comp_left(2) + D.inv_comp_right(2) D.iso_hcomp D.iso_inv_iso D.isos_compose D.seqE + D.seq_if_composable E''.counit_is_iso E'.counit_is_iso E'.ide_left + \' \ g unit_char(2)) qed obtain \ where \: "\\ : src\<^sub>C f \\<^sub>C g \\<^sub>C f\ \ F \ = ?\'" using assms g E'.antipar \' locally_full [of "src\<^sub>C f" "g \\<^sub>C f" ?\'] by (metis C.ide_hcomp C.ideD(1) C.in_hhomE C.src.preserves_ide C.hcomp_simps(1-2) C.src_trg C.trg_trg) have iso_\: "C.iso \" using \ \' iso_\' reflects_iso by auto have 1: "\\. \\ : f \\<^sub>C g \\<^sub>C src\<^sub>C g\ \ F \ = ?\'" using assms g \' locally_full [of "f \\<^sub>C g" "src\<^sub>C g" ?\'] by force obtain \ where \: "\\ : f \\<^sub>C g \\<^sub>C src\<^sub>C g\ \ F \ = ?\'" using 1 by blast have iso_\: "C.iso \" using \ \' iso_\' reflects_iso by auto have "equivalence_in_bicategory (\\<^sub>C) (\\<^sub>C) \\<^sub>C \\<^sub>C src\<^sub>C trg\<^sub>C f g \ \" using assms g \ \ iso_\ iso_\ by (unfold_locales, auto) thus ?thesis using C.equivalence_map_def by auto qed lemma reflects_equivalent_objects: assumes "C.obj a" and "C.obj b" and "D.equivalent_objects (map\<^sub>0 a) (map\<^sub>0 b)" shows "C.equivalent_objects a b" proof - obtain f' where f': "\f' : map\<^sub>0 a \\<^sub>D map\<^sub>0 b\ \ D.equivalence_map f'" using assms D.equivalent_objects_def D.equivalence_map_def by auto obtain f where f: "\f : a \\<^sub>C b\ \ C.ide f \ D.isomorphic (F f) f'" using assms f' locally_essentially_surjective [of a b f'] D.equivalence_map_is_ide by auto have "D.equivalence_map (F f)" using f f' D.equivalence_map_preserved_by_iso [of f' "F f"] D.isomorphic_symmetric by simp hence "C.equivalence_map f" using f f' reflects_equivalence_map [of f] by simp thus ?thesis using f C.equivalent_objects_def by auto qed end text\ For each pair of objects \a\, \b\ of \C\, an equivalence pseudofunctor restricts to an equivalence of categories between \C.hhom a b\ and \D.hhom (map\<^sub>0 a) (map\<^sub>0 b)\. \ (* TODO: Change the "perspective" of this locale to be the defined functor. *) locale equivalence_pseudofunctor_at_hom = equivalence_pseudofunctor + fixes a :: 'a and a' :: 'a assumes obj_a: "C.obj a" and obj_a': "C.obj a'" begin sublocale hhom\<^sub>C: subcategory V\<^sub>C \\\. \\ : a \\<^sub>C a'\\ using C.hhom_is_subcategory by simp sublocale hhom\<^sub>D: subcategory V\<^sub>D \\\. \\ : map\<^sub>0 a \\<^sub>D map\<^sub>0 a'\\ using D.hhom_is_subcategory by simp definition F\<^sub>1 where "F\<^sub>1 = (\\. if hhom\<^sub>C.arr \ then F \ else D.null)" interpretation F\<^sub>1: "functor" hhom\<^sub>C.comp hhom\<^sub>D.comp F\<^sub>1 unfolding F\<^sub>1_def using hhom\<^sub>C.arr_char hhom\<^sub>D.arr_char hhom\<^sub>C.dom_char hhom\<^sub>D.dom_char hhom\<^sub>C.cod_char hhom\<^sub>D.cod_char hhom\<^sub>C.seq_char hhom\<^sub>C.comp_char hhom\<^sub>D.comp_char by unfold_locales auto interpretation F\<^sub>1: fully_faithful_and_essentially_surjective_functor hhom\<^sub>C.comp hhom\<^sub>D.comp F\<^sub>1 proof show "\\ \'. \hhom\<^sub>C.par \ \'; F\<^sub>1 \ = F\<^sub>1 \'\ \ \ = \'" unfolding F\<^sub>1_def using is_faithful hhom\<^sub>C.dom_char hhom\<^sub>D.dom_char hhom\<^sub>C.cod_char hhom\<^sub>D.cod_char by (metis C.in_hhom_def hhom\<^sub>C.arrE) show "\f f' \. \hhom\<^sub>C.ide f; hhom\<^sub>C.ide f'; hhom\<^sub>D.in_hom \ (F\<^sub>1 f') (F\<^sub>1 f)\ \ \\. hhom\<^sub>C.in_hom \ f' f \ F\<^sub>1 \ = \" proof (unfold F\<^sub>1_def) fix f f' \ assume f: "hhom\<^sub>C.ide f" and f': "hhom\<^sub>C.ide f'" assume "hhom\<^sub>D.in_hom \ (if hhom\<^sub>C.arr f' then F f' else D.null) (if hhom\<^sub>C.arr f then F f else D.null)" hence \: "hhom\<^sub>D.in_hom \ (F f') (F f)" using f f' by simp have "\\. hhom\<^sub>C.in_hom \ f' f \ F \ = \" proof - have 1: "src\<^sub>C f' = src\<^sub>C f \ trg\<^sub>C f' = trg\<^sub>C f" using f f' hhom\<^sub>C.ide_char by (metis C.in_hhomE hhom\<^sub>C.arrE) hence ex: "\\. C.in_hom \ f' f \ F \ = \" - using f f' \ hhom\<^sub>C.in_hom_char hhom\<^sub>D.in_hom_char hhom\<^sub>C.ide_char locally_full - by simp + by (meson \ f f' hhom\<^sub>D.in_hom_char horizontal_homs.hhom_is_subcategory + locally_full subcategory.ide_char weak_arrow_of_homs_axioms + weak_arrow_of_homs_def) obtain \ where \: "C.in_hom \ f' f \ F \ = \" using ex by blast have "hhom\<^sub>C.in_hom \ f' f" - proof - - have 2: "hhom\<^sub>C.arr f' \ hhom\<^sub>C.arr f" - using f f' hhom\<^sub>C.arr_char hhom\<^sub>C.ide_char by simp - moreover have "hhom\<^sub>C.arr \" - using \ 1 2 hhom\<^sub>C.arr_char C.arrI C.vconn_implies_hpar(1-2) by auto - ultimately show ?thesis - using f f' \ hhom\<^sub>C.arr_char hhom\<^sub>C.ide_char hhom\<^sub>C.in_hom_char by simp - qed + by (metis C.arrI C.in_hhom_def C.vconn_implies_hpar(1-2) \ f f' + hhom\<^sub>C.arr_char hhom\<^sub>C.ide_char hhom\<^sub>C.in_hom_char) thus ?thesis using \ by auto qed thus "\\. hhom\<^sub>C.in_hom \ f' f \ (if hhom\<^sub>C.arr \ then F \ else D.null) = \" by auto qed show "\g. hhom\<^sub>D.ide g \ \f. hhom\<^sub>C.ide f \ hhom\<^sub>D.isomorphic (F\<^sub>1 f) g" proof (unfold F\<^sub>1_def) fix g assume g: "hhom\<^sub>D.ide g" show "\f. hhom\<^sub>C.ide f \ hhom\<^sub>D.isomorphic (if hhom\<^sub>C.arr f then F f else D.null) g" proof - have "C.obj a \ C.obj a'" using obj_a obj_a' by simp moreover have 1: "D.ide g \ \g : map\<^sub>0 a \\<^sub>D map\<^sub>0 a'\" using g obj_a obj_a' hhom\<^sub>D.ide_char by auto ultimately have 2: "\f. C.in_hhom f a a' \ C.ide f \ D.isomorphic (F f) g" using locally_essentially_surjective [of a a' g] by simp obtain f \ where f: "C.in_hhom f a a' \ C.ide f \ D.in_hom \ (F f) g \ D.iso \" using 2 by auto have "hhom\<^sub>C.ide f" using f hhom\<^sub>C.ide_char hhom\<^sub>C.arr_char by simp moreover have "hhom\<^sub>D.isomorphic (F f) g" proof - have "hhom\<^sub>D.arr \ \ hhom\<^sub>D.arr (D.inv \)" - using f 1 hhom\<^sub>D.arr_char D.in_hhom_def by auto + by (metis 1 D.arrI D.in_hhom_def D.vconn_implies_hpar(1-4) D.inv_in_homI + f hhom\<^sub>D.arrI) hence "hhom\<^sub>D.in_hom \ (F f) g \ hhom\<^sub>D.iso \" - using f g hhom\<^sub>D.in_hom_char hhom\<^sub>D.iso_char hhom\<^sub>C.arr_char hhom\<^sub>D.arr_char - hhom\<^sub>D.ideD(1) preserves_arr - by (simp add: C.in_hhom_def) + by (metis D.in_homE f hhom\<^sub>D.cod_simp hhom\<^sub>D.dom_simp hhom\<^sub>D.in_homI hhom\<^sub>D.iso_char) thus ?thesis unfolding hhom\<^sub>D.isomorphic_def by blast qed ultimately show "\f. hhom\<^sub>C.ide f \ hhom\<^sub>D.isomorphic (if hhom\<^sub>C.arr f then F f else D.null) g" by force qed qed qed lemma equivalence_functor_F\<^sub>1: shows "fully_faithful_and_essentially_surjective_functor hhom\<^sub>C.comp hhom\<^sub>D.comp F\<^sub>1" and "equivalence_functor hhom\<^sub>C.comp hhom\<^sub>D.comp F\<^sub>1" .. definition G\<^sub>1 where "G\<^sub>1 = (SOME G. \\\. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G F\<^sub>1 (fst \\) (snd \\))" lemma G\<^sub>1_props: assumes "C.obj a" and "C.obj a'" shows "\\ \. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \ \" proof - have "\G. \\\. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G F\<^sub>1 (fst \\) (snd \\)" using F\<^sub>1.extends_to_adjoint_equivalence by simp hence "\\\. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 (fst \\) (snd \\)" unfolding G\<^sub>1_def using someI_ex [of "\G. \\\. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G F\<^sub>1 (fst \\) (snd \\)"] by blast thus ?thesis by simp qed definition \ where "\ = (SOME \. \\. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \ \)" definition \ where "\ = (SOME \. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \ \)" lemma \\_props: shows "adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \ \" using obj_a obj_a' \_def \_def G\<^sub>1_props someI_ex [of "\\. \\. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \ \"] someI_ex [of "\\. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \ \"] by simp sublocale \\: adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \ \ using \\_props by simp sublocale \\: meta_adjunction hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \\.\ \\.\ using \\.induces_meta_adjunction by simp end context identity_pseudofunctor begin sublocale equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B map cmp - proof - show "\a'. B.obj a' \ \a. B.obj a \ B.equivalent_objects (map\<^sub>0 a) a'" - by (auto simp add: B.equivalent_objects_reflexive map\<^sub>0_def B.obj_simps) - show "\a b g. \B.obj a; B.obj b; B.in_hhom g (map\<^sub>0 a) (map\<^sub>0 b); B.ide g\ - \ \f. B.in_hhom f a b \ B.ide f \ B.isomorphic (map f) g" - using B.isomorphic_reflexive map\<^sub>0_def B.obj_simps by auto - show "\f f' \. \B.ide f; B.ide f'; src\<^sub>B f = src\<^sub>B f'; trg\<^sub>B f = trg\<^sub>B f'; - \\ : map f \\<^sub>B map f'\\ - \ \\. \\ : f \\<^sub>B f'\ \ map \ = \" - using B.arrI by auto - qed + using B.isomorphic_reflexive B.arrI + apply unfold_locales + by (auto simp add: B.equivalent_objects_reflexive map\<^sub>0_def B.obj_simps) lemma is_equivalence_pseudofunctor: shows "equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B map cmp" .. end locale composite_equivalence_pseudofunctor = composite_pseudofunctor + 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 + 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 begin interpretation faithful_functor V\<^sub>B V\<^sub>D \G o F\ using F.faithful_functor_axioms G.faithful_functor_axioms faithful_functors_compose by blast interpretation equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \G o F\ cmp proof show "\c. D.obj c \ \a. B.obj a \ D.equivalent_objects (map\<^sub>0 a) c" proof - fix c assume c: "D.obj c" obtain b where b: "C.obj b \ D.equivalent_objects (G.map\<^sub>0 b) c" using c G.biessentially_surjective_on_objects by auto obtain a where a: "B.obj a \ C.equivalent_objects (F.map\<^sub>0 a) b" using b F.biessentially_surjective_on_objects by auto have "D.equivalent_objects (map\<^sub>0 a) c" using a b map\<^sub>0_def G.preserves_equivalent_objects D.equivalent_objects_transitive by fastforce thus "\a. B.obj a \ D.equivalent_objects (map\<^sub>0 a) c" using a by auto qed show "\a a' h. \B.obj a; B.obj a'; \h : map\<^sub>0 a \\<^sub>D map\<^sub>0 a'\; D.ide h\ \ \f. B.in_hhom f a a' \ B.ide f \ D.isomorphic ((G \ F) f) h" proof - fix a a' h assume a: "B.obj a" and a': "B.obj a'" and h_in_hom: "\h : map\<^sub>0 a \\<^sub>D map\<^sub>0 a'\" and ide_h: "D.ide h" obtain g where g: "C.in_hhom g (F.map\<^sub>0 a) (F.map\<^sub>0 a') \ C.ide g \ D.isomorphic (G g) h" using a a' h_in_hom ide_h map\<^sub>0_def B.obj_simps G.locally_essentially_surjective [of "F.map\<^sub>0 a" "F.map\<^sub>0 a'" h] by auto obtain f where f: "B.in_hhom f a a' \ B.ide f \ C.isomorphic (F f) g" using a a' g F.locally_essentially_surjective by blast have "D.isomorphic ((G o F) f) h" - proof - - have "(G o F) f = G (F f)" - by simp - also have "D.isomorphic ... (G g)" - using f G.preserves_iso D.isomorphic_def by blast - also have "D.isomorphic ... h" - using g by simp - finally show "D.isomorphic ((G \ F) f) h" by simp - qed + by (metis D.isomorphic_transitive G.preserves_isomorphic comp_apply f g) thus "\f. B.in_hhom f a a' \ B.ide f \ D.isomorphic ((G \ F) f) h" using f by auto qed show "\f f' \. \B.ide f; B.ide f'; src\<^sub>B f = src\<^sub>B f'; trg\<^sub>B f = trg\<^sub>B f'; \\ : (G \ F) f \\<^sub>D (G \ F) f'\\ \ \\. \\ : f \\<^sub>B f'\ \ (G \ F) \ = \" proof - fix f f' \ assume f: "B.ide f" and f': "B.ide f'" and src: "src\<^sub>B f = src\<^sub>B f'" and trg: "trg\<^sub>B f = trg\<^sub>B f'" and \: "\\ : (G \ F) f \\<^sub>D (G \ F) f'\" have \: "\\ : G (F f) \\<^sub>D G (F f')\" using \ by simp have 1: "src\<^sub>C (F f) = src\<^sub>C (F f') \ trg\<^sub>C (F f) = trg\<^sub>C (F f')" using f f' src trg by simp have 2: "\\. \\ : F f \\<^sub>C F f'\ \ G \ = \" using f f' 1 \ G.locally_full F.preserves_ide by simp obtain \ where \: "\\ : F f \\<^sub>C F f'\ \ G \ = \" using 2 by auto obtain \ where \: "\\ : f \\<^sub>B f'\ \ F \ = \" using f f' src trg 2 \ F.locally_full by blast show "\\. \\ : f \\<^sub>B f'\ \ (G \ F) \ = \" using \ \ by auto qed qed sublocale equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D \G o F\ cmp .. lemma is_equivalence_pseudofunctor: shows "equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D (G o F) cmp" .. end end diff --git a/thys/Bicategory/Strictness.thy b/thys/Bicategory/Strictness.thy --- a/thys/Bicategory/Strictness.thy +++ b/thys/Bicategory/Strictness.thy @@ -1,4253 +1,4210 @@ (* Title: Strictness Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) section "Strictness" theory Strictness imports Category3.ConcreteCategory Pseudofunctor CanonicalIsos begin text \ In this section we consider bicategories in which some or all of the canonical isomorphisms are assumed to be identities. A \emph{normal} bicategory is one in which the unit isomorphisms are identities, so that unit laws for horizontal composition are satisfied ``on the nose''. A \emph{strict} bicategory (also known as a \emph{2-category}) is a bicategory in which both the unit and associativity isomoprhisms are identities, so that horizontal composition is strictly associative as well as strictly unital. From any given bicategory \B\ we may construct a related strict bicategory \S\, its \emph{strictification}, together with a pseudofunctor that embeds \B\ in \S\. The Strictness Theorem states that this pseudofunctor is an equivalence pseudofunctor, so that bicategory \B\ is biequivalent to its strictification. The Strictness Theorem is often used informally to justify suppressing canonical isomorphisms; which amounts to proving a theorem about 2-categories and asserting that it holds for all bicategories. Here we are working formally, so we can't just wave our hands and mutter something about the Strictness Theorem when we want to avoid dealing with units and associativities. However, in cases where we can establish that the property we would like to prove is reflected by the embedding of a bicategory in its strictification, then we can formally apply the Strictness Theorem to generalize to all bicategories a result proved for 2-categories. We will apply this approach here to simplify the proof of some facts about internal equivalences in a bicategory. \ subsection "Normal and Strict Bicategories" text \ A \emph{normal} bicategory is one in which the unit isomorphisms are identities, so that unit laws for horizontal composition are satisfied ``on the nose''. \ locale normal_bicategory = bicategory + assumes strict_lunit: "\f. ide f \ \[f] = f" and strict_runit: "\f. ide f \ \[f] = f" begin lemma strict_unit: assumes "obj a" shows "ide \[a]" using assms strict_runit unitor_coincidence(2) [of a] by auto lemma strict_lunit': assumes "ide f" shows "\\<^sup>-\<^sup>1[f] = f" using assms strict_lunit by simp lemma strict_runit': assumes "ide f" shows "\\<^sup>-\<^sup>1[f] = f" using assms strict_runit by simp lemma hcomp_obj_arr: assumes "obj b" and "arr f" and "b = trg f" shows "b \ f = f" using assms strict_lunit by (metis comp_arr_dom comp_ide_arr ide_cod ide_dom lunit_naturality) lemma hcomp_arr_obj: assumes "arr f" and "obj a" and "src f = a" shows "f \ a = f" using assms strict_runit by (metis comp_arr_dom comp_ide_arr ide_cod ide_dom runit_naturality) end text \ A \emph{strict} bicategory is a normal bicategory in which the associativities are also identities, so that associativity of horizontal composition holds ``on the nose''. \ locale strict_bicategory = normal_bicategory + assumes strict_assoc: "\f g h. \ide f; ide g; ide h; src f = trg g; src g = trg h\ \ ide \[f, g, h]" begin lemma strict_assoc': assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "ide \\<^sup>-\<^sup>1[f, g, h]" using assms strict_assoc by simp lemma hcomp_assoc: shows "(\ \ \) \ \ = \ \ \ \ \" proof (cases "hseq \ \ \ hseq \ \") show "\ (hseq \ \ \ hseq \ \) \ ?thesis" by (metis hseqE hseq_char' match_1 match_2) show "hseq \ \ \ hseq \ \ \ ?thesis" proof - assume 1: "hseq \ \ \ hseq \ \" have 2: "arr \ \ arr \ \ arr \ \ src \ = trg \ \ src \ = trg \" using 1 by blast have "(\ \ \) \ \ = \[cod \, cod \, cod \] \ ((\ \ \) \ \)" using 1 assoc_in_hom strict_assoc comp_cod_arr assoc_simps(4) hseq_char by simp also have "... = (\ \ \ \ \) \ \[dom \, dom \, dom \]" using 1 assoc_naturality by auto also have "... = \ \ \ \ \" using 2 assoc_in_hom [of "dom \" "dom \" "dom \"] strict_assoc comp_arr_dom by auto finally show ?thesis by simp qed qed text \ In a strict bicategory, every canonical isomorphism is an identity. \ interpretation bicategorical_language .. interpretation E: self_evaluation_map V H \ \ src trg .. notation E.eval ("\_\") lemma ide_eval_Can: assumes "Can t" shows "ide \t\" proof - have 1: "\u1 u2. \ ide \u1\; ide \u2\; Arr u1; Arr u2; Dom u1 = Cod u2 \ \ ide (\u1\ \ \u2\)" by (metis (no_types, lifting) E.eval_simps'(4-5) comp_ide_self ide_char) have "\u. Can u \ ide \u\" proof - fix u show "Can u \ ide \u\" (* TODO: Rename \_ide_simp \_ide_simp to \_ide_eq \_ide_eq *) using 1 \_def \'_def strict_lunit strict_runit strict_assoc strict_assoc' \_ide_simp \_ide_simp Can_implies_Arr comp_ide_arr E.eval_simps'(2-3) by (induct u) auto qed thus ?thesis using assms by simp qed lemma ide_can: assumes "Ide f" and "Ide g" and "\<^bold>\f\<^bold>\ = \<^bold>\g\<^bold>\" shows "ide (can g f)" using assms Can_red Can_Inv red_in_Hom Inv_in_Hom can_def ide_eval_Can Can.simps(4) Dom_Inv red_simps(4) by presburger end context bicategory begin text \ The following result gives conditions for strictness of a bicategory that are typically somewhat easier to verify than those used for the definition. \ lemma is_strict_if: assumes "\f. ide f \ f \ src f = f" and "\f. ide f \ trg f \ f = f" and "\a. obj a \ ide \[a]" and "\f g h. \ide f; ide g; ide h; src f = trg g; src g = trg h\ \ ide \[f, g, h]" shows "strict_bicategory V H \ \ src trg" proof show "\f g h. \ide f; ide g; ide h; src f = trg g; src g = trg h\ \ ide \[f, g, h]" by fact fix f assume f: "ide f" show "\[f] = f" proof - have "f = \[f]" using assms f unit_simps(5) by (intro lunit_eqI) (auto simp add: comp_arr_ide) thus ?thesis by simp qed show "\[f] = f" proof - have "f = \[f]" proof (intro runit_eqI) show "ide f" by fact show "\f : f \ src f \ f\" using f assms(1) by auto show "f \ src f = (f \ \[src f]) \ \[f, src f, src f]" proof - have "(f \ \[src f]) \ \[f, src f, src f] = (f \ src f) \ \[f, src f, src f]" using f assms(2-3) unit_simps(5) by simp also have "... = (f \ src f \ src f) \ \[f, src f, src f]" using f assms(1-2) ideD(1) trg_src src.preserves_ide by metis also have "... = f \ src f" using f comp_arr_ide assms(1,4) assoc_in_hom [of f "src f" "src f"] by auto finally show ?thesis by simp qed qed thus ?thesis by simp qed qed end subsection "Strictification" (* * TODO: Perhaps change the typeface used for a symbol that stands for a bicategory; * for example, to avoid the clashes here between B used as the name of a bicategory * and B used to denote a syntactic identity term. *) text \ The Strictness Theorem asserts that every bicategory is biequivalent to a strict bicategory. More specifically, it shows how to construct, given an arbitrary bicategory, a strict bicategory (its \emph{strictification}) that is biequivalent to it. Consequently, given a property \P\ of bicategories that is ``bicategorical'' (\emph{i.e.}~respects biequivalence), if we want to show that \P\ holds for a bicategory \B\ then it suffices to show that \P\ holds for the strictification of \B\, and if we want to show that \P\ holds for all bicategories, it is sufficient to show that it holds for all strict bicategories. This is very useful, because it becomes quite tedious, even with the aid of a proof assistant, to do ``diagram chases'' with all the units and associativities fully spelled out. Given a bicategory \B\, the strictification \S\ of \B\ may be constructed as the bicategory whose arrows are triples \(A, B, \)\, where \X\ and \Y\ are ``normal identity terms'' (essentially, nonempty horizontally composable lists of 1-cells of \B\) having the same syntactic source and target, and \\\ : \X\ \ \Y\\\ in \B\. Vertical composition in \S\ is given by composition of the underlying arrows in \B\. Horizontal composition in \S\ is given by \(A, B, \) \ (A', B', \') = (AA', BB', \)\, where \AA'\ and \BB'\ denote concatenations of lists and where \\\ is defined as the composition \can BB' (B \<^bold>\ B') \ (\ \ \') \ can (A \<^bold>\ A') AA'\, where \can (A \<^bold>\ A') AA'\ and \can BB' (B \<^bold>\ B')\ are canonical isomorphisms in \B\. The canonical isomorphism \can (A \<^bold>\ A') AA'\ corresponds to taking a pair of lists \A \<^bold>\ A'\ and ``shifting the parentheses to the right'' to obtain a single list \AA'\. The canonical isomorphism can \BB' (B \<^bold>\ B')\ corresponds to the inverse rearrangement. The bicategory \B\ embeds into its strictification \S\ via the functor \UP\ that takes each arrow \\\ of \B\ to \(\<^bold>\dom \\<^bold>\, \<^bold>\cod \\<^bold>\, \)\, where \\<^bold>\dom \\<^bold>\\ and \\<^bold>\cod \\<^bold>\\ denote one-element lists. This mapping extends to a pseudofunctor. There is also a pseudofunctor \DN\, which maps \(A, B, \)\ in \S\ to \\\ in \B\; this is such that \DN o UP\ is the identity on \B\ and \UP o DN\ is equivalent to the identity on \S\, so we obtain a biequivalence between \B\ and \S\. It seems difficult to find references that explicitly describe a strictification construction in elementary terms like this (in retrospect, it ought to have been relatively easy to rediscover such a construction, but my thinking got off on the wrong track). One reference that I did find useful was \cite{unapologetic-strictification}, which discusses strictification for monoidal categories. \ locale strictified_bicategory = B: bicategory V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B for V\<^sub>B :: "'a comp" (infixr "\\<^sub>B" 55) and H\<^sub>B :: "'a \ 'a \ 'a" (infixr "\\<^sub>B" 53) and \\<^sub>B :: "'a \ 'a \ 'a \ 'a" ("\\<^sub>B[_, _, _]") and \\<^sub>B :: "'a \ 'a" ("\\<^sub>B[_]") and src\<^sub>B :: "'a \ 'a" and trg\<^sub>B :: "'a \ 'a" begin sublocale E: self_evaluation_map V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B .. notation B.in_hhom ("\_ : _ \\<^sub>B _\") notation B.in_hom ("\_ : _ \\<^sub>B _\") notation E.eval ("\_\") notation E.Nmlize ("\<^bold>\_\<^bold>\") text \ The following gives the construction of a bicategory whose arrows are triples \(A, B, \)\, where \Nml A \ Ide A\, \Nml B \ Ide B\, \Src A = Src B\, \Trg A = Trg B\, and \\ : \A\ \ \B\\. We use @{locale concrete_category} to construct the vertical composition, so formally the arrows of the bicategory will be of the form \MkArr A B \\. \ text \ The 1-cells of the bicategory correspond to normal, identity terms \A\ in the bicategorical language associated with \B\. \ abbreviation IDE where "IDE \ {A. E.Nml A \ E.Ide A}" text \ If terms \A\ and \B\ determine 1-cells of the strictification and have a common source and target, then the 2-cells between these 1-cells correspond to arrows \\\ of the underlying bicategory such that \\\ : \A\ \\<^sub>B \B\\\. \ abbreviation HOM where "HOM A B \ {\. E.Src A = E.Src B \ E.Trg A = E.Trg B \ \\ : \A\ \\<^sub>B \B\\}" text \ The map taking term \A \ OBJ\ to its evaluation \\A\ \ HOM A A\ defines the embedding of 1-cells as identity 2-cells. \ abbreviation EVAL where "EVAL \ E.eval" sublocale concrete_category IDE HOM EVAL \\_ _ _ \ \. \ \\<^sub>B \\ using E.ide_eval_Ide B.comp_arr_dom B.comp_cod_arr B.comp_assoc by (unfold_locales, auto) lemma is_concrete_category: shows "concrete_category IDE HOM EVAL (\_ _ _ \ \. \ \\<^sub>B \)" .. abbreviation vcomp (infixr "\" 55) where "vcomp \ COMP" lemma arr_char: shows "arr F \ E.Nml (Dom F) \ E.Ide (Dom F) \ E.Nml (Cod F) \ E.Ide (Cod F) \ E.Src (Dom F) = E.Src (Cod F) \ E.Trg (Dom F) = E.Trg (Cod F) \ \Map F : \Dom F\ \\<^sub>B \Cod F\\ \ F \ Null" using arr_char by auto lemma arrI (* [intro] *): assumes "E.Nml (Dom F)" and "E.Ide (Dom F)" and "E.Nml (Cod F)" and "E.Ide (Cod F)" and "E.Src (Dom F) = E.Src (Cod F)" and "E.Trg (Dom F) = E.Trg (Cod F)" and "\Map F : \Dom F\ \\<^sub>B \Cod F\\" and "F \ Null" shows "arr F" using assms arr_char by blast lemma arrE [elim]: assumes "arr F" shows "(\ E.Nml (Dom F); E.Ide (Dom F); E.Nml (Cod F); E.Ide (Cod F); E.Src (Dom F) = E.Src (Cod F); E.Trg (Dom F) = E.Trg (Cod F); \Map F : \Dom F\ \\<^sub>B \Cod F\\; F \ Null \ \ T) \ T" using assms arr_char by simp lemma ide_char: shows "ide F \ endo F \ B.ide (Map F)" proof show "ide F \ endo F \ B.ide (Map F)" using ide_char by (simp add: E.ide_eval_Ide) show "endo F \ B.ide (Map F) \ ide F" by (metis (no_types, lifting) B.ide_char B.in_homE arr_char ide_char mem_Collect_eq seq_char) qed lemma ideI [intro]: assumes "arr F" and "Dom F = Cod F" and "B.ide (Map F)" shows "ide F" using assms ide_char dom_char cod_char seq_char by presburger lemma ideE [elim]: assumes "ide F" shows "(\ arr F; Dom F = Cod F; B.ide (Map F); Map F = \Dom F\; Map F = \Cod F\ \ \ T) \ T" - proof - - assume 1: "\ arr F; Dom F = Cod F; B.ide (Map F); Map F = \Dom F\; - Map F = \Cod F\ \ \ T" - show T - proof - - have "arr F" - using assms by auto - moreover have "Dom F = Cod F" - using assms ide_char dom_char cod_char - by (metis (no_types, lifting) Dom_cod calculation ideD(3)) - moreover have "B.ide (Map F)" - using assms ide_char by blast - moreover have "Map F = \Dom F\" - using assms ide_char dom_char Map_ide(1) by blast - ultimately show T - using 1 by simp - qed - qed + using assms + by (metis (no_types, lifting) Map_ide(2) ide_char seq_char) text \ Source and target are defined by the corresponding syntactic operations on terms. \ definition src where "src F \ if arr F then MkIde (E.Src (Dom F)) else null" definition trg where "trg F \ if arr F then MkIde (E.Trg (Dom F)) else null" lemma src_simps [simp]: assumes "arr F" shows "Dom (src F) = E.Src (Dom F)" and "Cod (src F) = E.Src (Dom F)" and "Map (src F) = \E.Src (Dom F)\" using assms src_def arr_char by auto lemma trg_simps [simp]: assumes "arr F" shows "Dom (trg F) = E.Trg (Dom F)" and "Cod (trg F) = E.Trg (Dom F)" and "Map (trg F) = \E.Trg (Dom F)\" using assms trg_def arr_char by auto interpretation src: endofunctor vcomp src using src_def comp_char E.Obj_implies_Ide apply (unfold_locales) apply auto[4] proof - show "\g f. seq g f \ src (g \ f) = src g \ src f" proof - fix g f assume gf: "seq g f" have "src (g \ f) = MkIde (E.Src (Dom (g \ f)))" using gf src_def comp_char by simp also have "... = MkIde (E.Src (Dom f))" using gf by (simp add: seq_char) also have "... = MkIde (E.Src (Dom g)) \ MkIde (E.Src (Dom f))" using gf seq_char E.Obj_implies_Ide by auto also have "... = src g \ src f" using gf src_def comp_char by auto finally show "src (g \ f) = src g \ src f" by blast qed qed interpretation trg: endofunctor vcomp trg using trg_def comp_char E.Obj_implies_Ide apply (unfold_locales) apply auto[4] proof - show "\g f. seq g f \ trg (g \ f) = trg g \ trg f" proof - fix g f assume gf: "seq g f" have "trg (g \ f) = MkIde (E.Trg (Dom (g \ f)))" using gf trg_def comp_char by simp also have "... = MkIde (E.Trg (Dom f))" using gf by (simp add: seq_char) also have "... = MkIde (E.Trg (Dom g)) \ MkIde (E.Trg (Dom f))" using gf seq_char E.Obj_implies_Ide by auto also have "... = trg g \ trg f" using gf trg_def comp_char by auto finally show "trg (g \ f) = trg g \ trg f" by blast qed qed interpretation horizontal_homs vcomp src trg using src_def trg_def Cod_in_Obj Map_in_Hom E.Obj_implies_Ide by unfold_locales auto notation in_hhom ("\_ : _ \ _\") definition hcomp (infixr "\" 53) where "\ \ \ \ if arr \ \ arr \ \ src \ = trg \ then MkArr (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)) else null" lemma arr_hcomp: assumes "arr \" and "arr \" and "src \ = trg \" shows "arr (\ \ \)" proof - have 1: "E.Ide (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \ E.Nml (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \ E.Ide (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \ E.Nml (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \)" using assms arr_char src_def trg_def E.Ide_HcompNml E.Nml_HcompNml(1) by auto moreover have "\B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) : \Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \\ \\<^sub>B \Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \\\" proof - have "\Map \ \\<^sub>B Map \ : \Dom \ \<^bold>\ Dom \\ \\<^sub>B \Cod \ \<^bold>\ Cod \\\" using assms arr_char dom_char cod_char src_def trg_def E.eval_simps'(2-3) by auto moreover have "\B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) : \Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \\ \\<^sub>B \Dom \ \<^bold>\ Dom \\\ \ \B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) : \Cod \ \<^bold>\ Cod \\ \\<^sub>B \Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \\\" using assms 1 arr_char src_def trg_def apply (intro conjI B.in_homI) by auto ultimately show ?thesis by auto qed moreover have "E.Src (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = E.Src (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \ E.Trg (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = E.Trg (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \)" using assms arr_char src_def trg_def by (simp add: E.Src_HcompNml E.Trg_HcompNml) ultimately show ?thesis unfolding hcomp_def using assms by (intro arrI, auto) qed lemma src_hcomp [simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "src (\ \ \) = src \" using assms arr_char hcomp_def src_def trg_def arr_hcomp E.Src_HcompNml by simp lemma trg_hcomp [simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "trg (hcomp \ \) = trg \" using assms arr_char hcomp_def src_def trg_def arr_hcomp E.Trg_HcompNml by simp lemma hseq_char: shows "arr (\ \ \) \ arr \ \ arr \ \ src \ = trg \" using arr_hcomp hcomp_def by simp lemma Dom_hcomp [simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "Dom (\ \ \) = Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \" using assms hcomp_def [of \ \] by simp lemma Cod_hcomp [simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "Cod (\ \ \) = Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \" using assms hcomp_def [of \ \] by simp lemma Map_hcomp [simp]: assumes "arr \" and "arr \" and "src \ = trg \" shows "Map (\ \ \) = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using assms hcomp_def [of \ \] by simp interpretation "functor" VV.comp vcomp \\\\. hcomp (fst \\) (snd \\)\ proof show "\f. \ VV.arr f \ fst f \ snd f = null" using hcomp_def by auto show A: "\f. VV.arr f \ arr (fst f \ snd f)" using VV.arrE hseq_char by blast show "\f. VV.arr f \ dom (fst f \ snd f) = fst (VV.dom f) \ snd (VV.dom f)" proof - fix f assume f: "VV.arr f" have "dom (fst f \ snd f) = MkIde (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" using f VV.arrE [of f] dom_char arr_hcomp hcomp_def by simp also have "... = fst (VV.dom f) \ snd (VV.dom f)" proof - have "hcomp (fst (VV.dom f)) (snd (VV.dom f)) = MkArr (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) (B.can (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\ Dom (snd f)) \\<^sub>B (\Dom (fst f)\ \\<^sub>B \Dom (snd f)\) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)))" using f VV.arrE [of f] arr_hcomp hcomp_def VV.dom_simp by simp moreover have "B.can (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\ Dom (snd f)) \\<^sub>B (\Dom (fst f)\ \\<^sub>B \Dom (snd f)\) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) = \Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)\" proof - have 1: "E.Ide (Dom (fst f) \<^bold>\ Dom (snd f))" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) src_simps(1) trg_simps(1)) have 2: "E.Ide (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) E.Ide_HcompNml src_simps(1) trg_simps(1)) have 3: "\<^bold>\Dom (fst f) \<^bold>\ Dom (snd f)\<^bold>\ = \<^bold>\Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)\<^bold>\" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize_Nml src_simps(1) trg_simps(1)) have "(\Dom (fst f)\ \\<^sub>B \Dom (snd f)\) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) = B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" using "1" "2" "3" B.comp_cod_arr by force thus ?thesis using 1 2 3 f VV.arr_char B.can_Ide_self B.vcomp_can by simp qed ultimately show ?thesis by simp qed finally show "dom (fst f \ snd f) = fst (VV.dom f) \ snd (VV.dom f)" by simp qed show "\f. VV.arr f \ cod (fst f \ snd f) = fst (VV.cod f) \ snd (VV.cod f)" proof - fix f assume f: "VV.arr f" have "cod (fst f \ snd f) = MkIde (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f))" using f VV.arrE [of f] cod_char arr_hcomp hcomp_def by simp also have "... = fst (VV.cod f) \ snd (VV.cod f)" proof - have "hcomp (fst (VV.cod f)) (snd (VV.cod f)) = MkArr (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f)) \\<^sub>B (\Cod (fst f)\ \\<^sub>B \Cod (snd f)\) \\<^sub>B B.can (Cod (fst f) \<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)))" using f VV.arrE [of f] arr_hcomp hcomp_def VV.cod_simp by simp moreover have "B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f)) \\<^sub>B (\Cod (fst f)\ \\<^sub>B \Cod (snd f)\) \\<^sub>B B.can (Cod (fst f) \<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) = \Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)\" proof - have 1: "E.Ide (Cod (fst f) \<^bold>\ Cod (snd f))" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) src_simps(1) trg_simps(1)) have 2: "E.Ide (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f))" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) E.Ide_HcompNml src_simps(1) trg_simps(1)) have 3: "\<^bold>\Cod (fst f) \<^bold>\ Cod (snd f)\<^bold>\ = \<^bold>\Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)\<^bold>\" using f VV.arr_char arr_char dom_char apply simp by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize_Nml src_simps(1) trg_simps(1)) have "(\Cod (fst f)\ \\<^sub>B \Cod (snd f)\) \\<^sub>B B.can (Cod (fst f) \<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) = B.can (Cod (fst f) \<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f))" using "1" "2" "3" B.comp_cod_arr by force thus ?thesis using 1 2 3 f VV.arr_char B.can_Ide_self B.vcomp_can by simp qed ultimately show ?thesis by simp qed finally show "cod (fst f \ snd f) = fst (VV.cod f) \ snd (VV.cod f)" by simp qed show "\g f. VV.seq g f \ fst (VV.comp g f) \ snd (VV.comp g f) = (fst g \ snd g) \ (fst f \ snd f)" proof - fix f g assume fg: "VV.seq g f" have f: "arr (fst f) \ arr (snd f) \ src (fst f) = trg (snd f)" using fg VV.seq_char VV.arr_char by simp have g: "arr (fst g) \ arr (snd g) \ src (fst g) = trg (snd g)" using fg VV.seq_char VV.arr_char by simp have 1: "arr (fst (VV.comp g f)) \ arr (snd (VV.comp g f)) \ src (fst (VV.comp g f)) = trg (snd (VV.comp g f))" using fg VV.arrE by blast have 0: "VV.comp g f = (fst g \ fst f, snd g \ snd f)" using fg 1 VV.comp_char VxV.comp_char by (metis (no_types, lifting) VV.seq_char VxV.seqE) let ?X = "MkArr (Dom (fst (VV.comp g f)) \<^bold>\\<^bold>\\<^bold>\ Dom (snd (VV.comp g f))) (Cod (fst (VV.comp g f)) \<^bold>\\<^bold>\\<^bold>\ Cod (snd (VV.comp g f))) (B.can (Cod (fst (VV.comp g f)) \<^bold>\\<^bold>\\<^bold>\ Cod (snd (VV.comp g f))) (Cod (fst (VV.comp g f)) \<^bold>\ Cod (snd (VV.comp g f))) \\<^sub>B (Map (fst (VV.comp g f)) \\<^sub>B Map (snd (VV.comp g f))) \\<^sub>B B.can (Dom (fst (VV.comp g f)) \<^bold>\ Dom (snd (VV.comp g f))) (Dom (fst (VV.comp g f)) \<^bold>\\<^bold>\\<^bold>\ Dom (snd (VV.comp g f))))" have 2: "fst (VV.comp g f) \ snd (VV.comp g f) = ?X" unfolding hcomp_def using 1 by simp also have "... = (fst g \ snd g) \ (fst f \ snd f)" proof - let ?GG = "MkArr (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)) (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B (Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)))" let ?FF = "MkArr (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)) (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f)) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f)) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)))" have 4: "arr ?FF \ arr ?GG \ Dom ?GG = Cod ?FF" proof - have "arr ?FF \ arr ?GG" using f g fg VV.arr_char VV.seqE hcomp_def A by presburger thus ?thesis using 0 1 by (simp add: fg seq_char) qed have "(fst g \ snd g) \ (fst f \ snd f) = ?GG \ ?FF" unfolding hcomp_def using 1 f g fg VV.arr_char VV.seqE by simp also have "... = ?X" proof (intro arr_eqI) show "seq ?GG ?FF" using fg 4 seq_char by blast show "arr ?X" using fg 1 arr_hcomp hcomp_def by simp show "Dom (?GG \ ?FF) = Dom ?X" using fg 0 1 4 seq_char by simp show "Cod (?GG \ ?FF) = Cod ?X" using fg 0 1 4 seq_char by simp show "Map (?GG \ ?FF) = Map ?X" proof - have "Map (?GG \ ?FF) = Map ?GG \\<^sub>B Map ?FF" using 4 by auto also have "... = (B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B (Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g))) \\<^sub>B (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f)) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f)) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f)))" using fg by simp also have "... = B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B ((Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f))) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" proof - have "(B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B (Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g))) \\<^sub>B (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f)) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f)) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))) = B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B ((Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B (B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)) \\<^sub>B B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f))) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f))) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" using B.comp_assoc by simp also have "... = B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B ((Map (fst g) \\<^sub>B Map (snd g)) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f))) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" proof - have "(B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g))) \\<^sub>B (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f))) = \Cod (fst f) \<^bold>\ Cod (snd f)\" proof - have "(B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g))) \\<^sub>B (B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f))) = B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Cod (fst f) \<^bold>\ Cod (snd f))" proof - have "E.Ide (Dom (fst g) \<^bold>\ Dom (snd g))" using g arr_char - apply simp - by (metis (no_types, lifting) src_simps(1) trg_simps(1)) + by (metis (no_types, lifting) E.Ide.simps(3) src_simps(2) trg_simps(2)) moreover have "E.Ide (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g))" - using g arr_char - apply simp - by (metis (no_types, lifting) E.Ide_HcompNml src_simps(1) trg_simps(1)) + using 4 by auto moreover have "E.Ide (Cod (fst f) \<^bold>\ Cod (snd f))" using f arr_char - apply simp - by (metis (no_types, lifting) src_simps(1) trg_simps(1)) + by (metis (no_types, lifting) E.Ide.simps(3) src_simps(2) trg_simps(2)) moreover have "\<^bold>\Dom (fst g) \<^bold>\ Dom (snd g)\<^bold>\ = \<^bold>\Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)\<^bold>\" - using g - apply simp - by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize_Nml - arrE src_simps(1) trg_simps(1)) + using g E.Nml_HcompNml(1) calculation(1) by fastforce moreover have "\<^bold>\Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)\<^bold>\ = \<^bold>\Cod (fst f) \<^bold>\ Cod (snd f)\<^bold>\" - using g - apply simp - by (metis (no_types, lifting) "0" "1" E.Nmlize.simps(3) - calculation(4) fst_conv seq_char snd_conv) + using g fg seq_char + by (metis (no_types, lifting) VV.seq_char VxV.seqE calculation(4)) moreover have "Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g) = Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)" using 0 1 by (simp add: seq_char) ultimately show ?thesis using B.vcomp_can by simp qed also have "... = \Cod (fst f) \<^bold>\ Cod (snd f)\" proof - have "Dom (fst g) \<^bold>\ Dom (snd g) = Cod (fst f) \<^bold>\ Cod (snd f)" using 0 f g fg seq_char VV.seq_char VV.arr_char by simp thus ?thesis using f B.can_Ide_self [of "Dom (fst f) \<^bold>\ Dom (snd f)"] - apply simp - by (metis (no_types, lifting) B.can_Ide_self E.eval.simps(3) - E.Ide.simps(3) arr_char src_simps(2) trg_simps(2)) + by (metis (no_types, lifting) B.can_Ide_self E.Ide.simps(3) + arrE src_simps(2) trg_simps(2)) qed finally show ?thesis by simp qed hence "(B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)) \\<^sub>B B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f))) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f)) = \Cod (fst f) \<^bold>\ Cod (snd f)\ \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f))" by simp also have "... = Map (fst f) \\<^sub>B Map (snd f)" proof - have 1: "\p. arr p \ map (cod p) \ map p = map p" by blast have 3: "\Cod (fst f)\ \\<^sub>B Map (fst f) = Map (map (cod (fst f)) \ map (fst f))" by (simp add: f) have 4: "map (cod (fst f)) \ map (fst f) = fst f" using 1 f map_simp by simp show ?thesis proof - have 2: "\Cod (snd f)\ \\<^sub>B Map (snd f) = Map (snd f)" - proof - - have "\Cod (snd f)\ \\<^sub>B Map (snd f) = - Map (map (cod (snd f)) \ map (snd f))" - by (simp add: f) - moreover have "map (cod (snd f)) \ map (snd f) = snd f" - using 1 f map_simp by simp - ultimately show ?thesis by presburger - qed + using 1 f map_simp + by (metis (no_types, lifting) Dom_cod Map_cod Map_comp arr_cod) have "B.seq \Cod (snd f)\ (Map (snd f))" using f 2 by auto moreover have "B.seq \Cod (fst f)\ (Map (fst f))" using 4 f 3 by auto moreover have "\Cod (fst f)\ \\<^sub>B Map (fst f) \\<^sub>B \Cod (snd f)\ \\<^sub>B Map (snd f) = Map (fst f) \\<^sub>B Map (snd f)" using 2 3 4 by presburger ultimately show ?thesis by (simp add: B.interchange) qed qed finally have "(B.can (Dom (fst g) \<^bold>\ Dom (snd g)) (Dom (fst g) \<^bold>\\<^bold>\\<^bold>\ Dom (snd g)) \\<^sub>B B.can (Cod (fst f) \<^bold>\\<^bold>\\<^bold>\ Cod (snd f)) (Cod (fst f) \<^bold>\ Cod (snd f))) \\<^sub>B (Map (fst f) \\<^sub>B Map (snd f)) = Map (fst f) \\<^sub>B Map (snd f)" by simp thus ?thesis using fg B.comp_cod_arr by simp qed finally show ?thesis by simp qed also have "... = B.can (Cod (fst g) \<^bold>\\<^bold>\\<^bold>\ Cod (snd g)) (Cod (fst g) \<^bold>\ Cod (snd g)) \\<^sub>B (Map (fst g \ fst f) \\<^sub>B Map (snd g \ snd f)) \\<^sub>B B.can (Dom (fst f) \<^bold>\ Dom (snd f)) (Dom (fst f) \<^bold>\\<^bold>\\<^bold>\ Dom (snd f))" proof - have 2: "Dom (fst g) = Cod (fst f)" using 0 f g fg VV.seq_char [of g f] VV.arr_char arr_char seq_char by (metis (no_types, lifting) fst_conv) hence "Map (fst g \ fst f) = Map (fst g) \\<^sub>B Map (fst f)" using f g Map_comp [of "fst f" "fst g"] by simp moreover have "B.seq (Map (fst g)) (Map (fst f)) \ B.seq (Map (snd g)) (Map (snd f))" using f g 0 1 2 arr_char by (metis (no_types, lifting) B.seqI' prod.sel(2) seq_char) ultimately show ?thesis using 0 1 seq_char Map_comp B.interchange by auto qed also have "... = Map ?X" using fg 0 1 by (simp add: seq_char) finally show ?thesis by simp qed qed finally show ?thesis by simp qed finally show "fst (VV.comp g f) \ snd (VV.comp g f) = (fst g \ snd g) \ (fst f \ snd f)" by simp qed qed interpretation horizontal_composition vcomp hcomp src trg using hseq_char by (unfold_locales, auto) lemma hcomp_assoc: assumes "arr \" and "arr \" and "arr \" and "src \ = trg \" and "src \ = trg \" shows "(\ \ \) \ \ = \ \ \ \ \" proof (intro arr_eqI) have \\: "\Map \ \\<^sub>B Map \ : \Dom \ \<^bold>\ Dom \\ \\<^sub>B \Cod \ \<^bold>\ Cod \\\" using assms src_def trg_def arr_char by (auto simp add: E.eval_simps'(2-3) Pair_inject) have \\: "\Map \ \\<^sub>B Map \ : \Dom \ \<^bold>\ Dom \\ \\<^sub>B \Cod \ \<^bold>\ Cod \\\" using assms src_def trg_def arr_char by (auto simp add: E.eval_simps'(2-3) Pair_inject) show "hseq (\ \ \) \" using assms \\ \\ by auto show "hseq \ (\ \ \)" using assms \\ \\ by auto show "Dom ((\ \ \) \ \) = Dom (\ \ \ \ \)" proof - have "E.Nml (Dom \) \ E.Nml (Dom \) \ E.Nml (Dom \)" using assms by blast moreover have "E.Src (Dom \) = E.Trg (Dom \) \ E.Src (Dom \) = E.Trg (Dom \)" using assms \\ \\ by (metis (no_types, lifting) src_simps(2) trg_simps(2)) ultimately show ?thesis using assms \\ \\ E.HcompNml_assoc by simp qed show "Cod ((\ \ \) \ \) = Cod (\ \ \ \ \)" proof - have "E.Nml (Cod \) \ E.Nml (Cod \) \ E.Nml (Cod \)" using assms by blast moreover have "E.Src (Cod \) = E.Trg (Cod \) \ E.Src (Cod \) = E.Trg (Cod \)" using assms \\ \\ by (metis (no_types, lifting) arrE src_simps(2) trg_simps(2)) ultimately show ?thesis using assms \\ \\ E.HcompNml_assoc by simp qed show "Map ((\ \ \) \ \) = Map (\ \ \ \ \)" proof - have "Map ((\ \ \) \ \) = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" proof - have 1: "Map ((\ \ \) \ \) = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) \\<^sub>B (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B Map \) \\<^sub>B B.can ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using assms \\ \\ hcomp_def E.HcompNml_assoc src_def trg_def arr_char E.Nml_HcompNml E.Ide_HcompNml by auto (* 5 sec *) also have "... = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) \\<^sub>B (B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \)) \\<^sub>B B.can ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" proof - have "B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B Map \ = B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \)" proof - have "B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B Map \ = (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B B.can (Cod \) (Cod \)) \\<^sub>B ((Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B Map \)" proof - have "B.seq (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \)) ((Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" by (metis (no_types, lifting) B.arrI Map_hcomp arrE arr_hcomp assms(1) assms(2) assms(4)) moreover have "B.seq (B.can (Cod \) (Cod \)) (Map \)" using B.can_in_hom assms(3) by blast moreover have "B.ide (B.can (Cod \) (Cod \))" using B.can_Ide_self E.ide_eval_Ide arr_char assms(3) by presburger ultimately show ?thesis by (metis (no_types) B.comp_ide_arr B.interchange) qed also have "... = (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B B.can (Cod \) (Cod \)) \\<^sub>B ((Map \ \\<^sub>B Map \) \\<^sub>B Map \) \\<^sub>B (B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B B.can (Dom \) (Dom \))" proof - have "B.seq (Map \ \\<^sub>B Map \) (B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" by (metis (no_types, lifting) B.arrI B.comp_null(2) B.ext Map_hcomp arrE arr_hcomp assms(1) assms(2) assms(4)) moreover have "B.seq (Map \) (B.can (Dom \) (Dom \))" using assms(3) by fastforce ultimately show ?thesis using B.interchange by (metis (no_types, lifting) B.can_Ide_self B.comp_arr_ide E.ide_eval_Ide arrE assms(3)) qed also have "... = (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B B.can (Cod \) (Cod \)) \\<^sub>B (B.can ((Cod \ \<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \)) \\<^sub>B (B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B B.can (Dom \) (Dom \))" proof - have "(Map \ \\<^sub>B Map \) \\<^sub>B Map \ = B.\' \Cod \\ \Cod \\ \Cod \\ \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B \\<^sub>B \Dom \\ \Dom \\ \Dom \\" using B.hcomp_reassoc(1) by (metis (no_types, lifting) B.hcomp_in_vhomE B.in_homE \\ \\ arrE assms(1) assms(2) assms(3)) also have "... = B.can ((Cod \ \<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \)" using assms arr_char src_def trg_def arr_char B.canE_associator by simp finally show ?thesis by simp qed also have "... = ((B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B B.can (Cod \) (Cod \)) \\<^sub>B (B.can ((Cod \ \<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \))) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B (B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \) \\<^sub>B (B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B B.can (Dom \) (Dom \)))" using B.comp_assoc by simp also have "... = B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \)" proof - have "(B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B B.can (Cod \) (Cod \)) \\<^sub>B (B.can ((Cod \ \<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \)) = B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \)" proof - have "E.Ide (Cod \ \<^bold>\ Cod \)" by (metis (no_types, lifting) E.Ide.simps(3) arrE assms(1-2,4) src_simps(1) trg_simps(1)) moreover have "E.Ide (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \)" using E.Ide_HcompNml assms(1) assms(2) calculation by auto moreover have "\<^bold>\Cod \ \<^bold>\ Cod \\<^bold>\ = \<^bold>\Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \\<^bold>\" using E.Nml_HcompNml(1) assms(1) assms(2) calculation(1) by fastforce moreover have "E.Src (Cod \ \<^bold>\ Cod \) = E.Trg (Cod \)" by (metis (no_types, lifting) E.Src.simps(3) arrE assms(2-3,5) src_simps(2) trg_simps(2)) moreover have "E.Src (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) = E.Trg (Cod \)" using E.Src_HcompNml assms(1) assms(2) calculation(1) calculation(4) by fastforce moreover have "\<^bold>\Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \\<^bold>\ = \<^bold>\(Cod \ \<^bold>\ Cod \) \<^bold>\ Cod \\<^bold>\" by (metis (no_types, lifting) E.Arr.simps(3) E.Nmlize_Hcomp_Hcomp E.Nmlize_Hcomp_Hcomp' E.Ide_implies_Arr E.Src.simps(3) arrE assms(3) calculation(1) calculation(4)) ultimately show ?thesis using assms(3) B.hcomp_can B.vcomp_can by auto qed moreover have "B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \) \\<^sub>B (B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \\<^sub>B B.can (Dom \) (Dom \)) = B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \)" proof - have "E.Ide (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" by (metis (no_types, lifting) E.Ide_HcompNml arrE assms(1-2,4) src_simps(2) trg_simps(2)) moreover have "E.Ide (Dom \ \<^bold>\ Dom \)" by (metis (no_types, lifting) E.Ide.simps(3) arrE assms(1-2,4) src_simps(1) trg_simps(1)) moreover have "\<^bold>\Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \\<^bold>\ = \<^bold>\Dom \ \<^bold>\ Dom \\<^bold>\" using E.Nml_HcompNml(1) assms(1-2) calculation(2) by fastforce moreover have "E.Src (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = E.Trg (Dom \)" by (metis (no_types, lifting) E.Ide.simps(3) E.Src_HcompNml arrE assms(1-3,5) calculation(2) src_simps(2) trg_simps(2)) moreover have "E.Src (Dom \ \<^bold>\ Dom \) = E.Trg (Dom \)" using E.Src_HcompNml assms(1-2) calculation(2) calculation(4) by fastforce moreover have "E.Ide ((Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \)" using E.Ide.simps(3) assms(3) calculation(2) calculation(5) by blast moreover have "\<^bold>\(Dom \ \<^bold>\ Dom \) \<^bold>\ Dom \\<^bold>\ = \<^bold>\Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \\<^bold>\" using E.Nmlize_Hcomp_Hcomp calculation(6) by auto ultimately show ?thesis using assms(3) B.hcomp_can B.vcomp_can by auto qed ultimately show ?thesis by simp qed finally show ?thesis by simp qed thus ?thesis by simp qed also have "... = (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) \\<^sub>B B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \)) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) \\<^sub>B B.can ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using B.comp_assoc by simp also have "... = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" proof - have "B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) \\<^sub>B B.can ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \)" proof - have "E.Ide (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \)" using assms src_def trg_def by fastforce moreover have "E.Ide ((Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \)" using assms arr_char src_def trg_def E.Ide_HcompNml E.Src_HcompNml by auto moreover have "E.Ide (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \)" using assms arr_char src_def trg_def by (simp add: E.Nml_HcompNml(1) E.Ide_HcompNml E.Trg_HcompNml) moreover have "\<^bold>\Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \\<^bold>\ = \<^bold>\(Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \\<^bold>\" using assms arr_char src_def trg_def E.Nml_HcompNml E.HcompNml_assoc by simp moreover have "\<^bold>\(Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \<^bold>\ Cod \\<^bold>\ = \<^bold>\Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \\<^bold>\" using assms arr_char src_def trg_def E.Nml_HcompNml E.HcompNml_assoc by simp ultimately show ?thesis by simp qed moreover have "B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) \\<^sub>B B.can ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" proof - have "E.Ide (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \)" using assms src_def trg_def by fastforce moreover have "E.Ide ((Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \)" using assms arr_char src_def trg_def E.Ide_HcompNml E.Src_HcompNml by auto moreover have "E.Ide (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using assms arr_char src_def trg_def by (simp add: E.Nml_HcompNml(1) E.Ide_HcompNml E.Trg_HcompNml) moreover have "\<^bold>\Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \\<^bold>\ = \<^bold>\(Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \\<^bold>\" using assms arr_char src_def trg_def E.Nml_HcompNml E.HcompNml_assoc by simp moreover have "\<^bold>\(Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) \<^bold>\ Dom \\<^bold>\ = \<^bold>\Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \\<^bold>\" using assms arr_char src_def trg_def E.Nml_HcompNml E.HcompNml_assoc by simp ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed finally show ?thesis by simp qed also have "... = Map (\ \ \ \ \)" proof - have 1: "Map (\ \ \ \ \) = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using assms Map_hcomp [of \ "\ \ \"] Map_hcomp [of \ \] by simp also have "... = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \\<^sub>B ((B.can (Cod \) (Cod \) \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \)) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B (B.can (Dom \) (Dom \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" proof - have "Map \ \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = (B.can (Cod \) (Cod \) \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \)) \\<^sub>B (Map \ \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" using assms B.interchange B.comp_cod_arr by (metis (no_types, lifting) B.can_Ide_self B.in_homE Map_hcomp arrE hseq_char) also have "... = (B.can (Cod \) (Cod \) \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \)) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B (B.can (Dom \) (Dom \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" using assms B.interchange B.comp_arr_dom [of "Map \" "B.can (Dom \) (Dom \)"] by (metis (no_types, lifting) B.can_Ide_self B.comp_null(2) B.ext B.in_homE Map_hcomp arrE hseq_char) finally have "Map \ \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) = (B.can (Cod \) (Cod \) \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \)) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B (B.can (Dom \) (Dom \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" by simp thus ?thesis by simp qed also have "... = (B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) \\<^sub>B (B.can (Cod \) (Cod \) \\<^sub>B B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \))) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B ((B.can (Dom \) (Dom \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \))" using B.comp_assoc by simp also have "... = B.can (Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \ \<^bold>\\<^bold>\\<^bold>\ Cod \) (Cod \ \<^bold>\ Cod \ \<^bold>\ Cod \) \\<^sub>B (Map \ \\<^sub>B Map \ \\<^sub>B Map \) \\<^sub>B B.can (Dom \ \<^bold>\ Dom \ \<^bold>\ Dom \) (Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \ \<^bold>\\<^bold>\\<^bold>\ Dom \)" using assms \\ \\ E.HcompNml_assoc src_def trg_def arr_char E.Src_HcompNml E.Trg_HcompNml E.Nml_HcompNml E.Ide_HcompNml by simp finally show ?thesis by simp qed ultimately show ?thesis by metis qed qed lemma obj_char: shows "obj a \ endo a \ E.Obj (Dom a) \ Map a = \Dom a\" proof assume a: "obj a" show "endo a \ E.Obj (Dom a) \ Map a = \Dom a\" proof (intro conjI) show "endo a" using a ide_char by blast show "E.Obj (Dom a)" using a ide_char src_def by (metis (no_types, lifting) E.Ide_implies_Arr E.Obj_Trg arrE obj_def trg_simps(1) trg_src) show "Map a = \Dom a\" using a ide_char src_def by blast qed next assume a: "endo a \ E.Obj (Dom a) \ Map a = \Dom a\" show "obj a" proof - have "arr a" using a by auto moreover have "src a = a" using a E.Obj_in_Hom(1) seq_char by (intro arr_eqI, auto) ultimately show ?thesis using obj_def by simp qed qed lemma hcomp_obj_self: assumes "obj a" shows "a \ a = a" proof (intro arr_eqI) show "hseq a a" using assms by auto show "arr a" using assms by auto show 1: "Dom (a \ a) = Dom a" unfolding hcomp_def using assms arr_char E.HcompNml_Trg_Nml apply simp by (metis (no_types, lifting) objE obj_def trg_simps(1)) show 2: "Cod (a \ a) = Cod a" unfolding hcomp_def using assms 1 arr_char E.HcompNml_Trg_Nml apply simp by (metis (no_types, lifting) Dom_hcomp ideE objE) show "Map (a \ a) = Map a" using "1" Map_ide(1) assms by fastforce qed lemma hcomp_ide_src: assumes "ide f" shows "f \ src f = f" proof (intro arr_eqI) show "hseq f (src f)" using assms by simp show "arr f" using assms by simp show 1: "Dom (f \ src f) = Dom f" unfolding hcomp_def using assms apply simp using assms ide_char arr_char E.HcompNml_Nml_Src by (metis (no_types, lifting) ideD(1)) show "Cod (f \ src f) = Cod f" unfolding hcomp_def using assms apply simp using assms ide_char arr_char E.HcompNml_Nml_Src by (metis (no_types, lifting) ideD(1)) show "Map (f \ src f) = Map f" by (simp add: "1" Map_ide(1) assms) qed lemma hcomp_trg_ide: assumes "ide f" shows "trg f \ f = f" proof (intro arr_eqI) show "hseq (trg f) f" using assms by auto show "arr f" using assms by auto show 1: "Dom (trg f \ f) = Dom f" unfolding hcomp_def using assms apply simp using assms ide_char arr_char E.HcompNml_Trg_Nml by (metis (no_types, lifting) ideD(1)) show "Cod (trg f \ f) = Cod f" unfolding hcomp_def using assms apply simp using assms ide_char arr_char E.HcompNml_Trg_Nml by (metis (no_types, lifting) ideD(1)) show "Map (trg f \ f) = Map f" by (simp add: "1" Map_ide(1) assms) qed interpretation L: full_functor vcomp vcomp L proof fix a a' g assume a: "ide a" and a': "ide a'" assume g: "in_hom g (L a') (L a)" have a_eq: "a = MkIde (Dom a)" using a dom_char [of a] by simp have a'_eq: "a' = MkIde (Dom a')" using a' dom_char [of a'] by simp have 1: "Cod g = Dom a" proof - have "Dom (L a) = Dom a" by (simp add: a hcomp_trg_ide) thus ?thesis using g cod_char [of g] by (metis (no_types, lifting) Dom_cod in_homE) qed have 2: "Dom g = Dom a'" using a' g hcomp_trg_ide in_hom_char by auto let ?f = "MkArr (Dom a') (Cod a) (Map g)" have f: "in_hom ?f a' a" by (metis (no_types, lifting) "1" "2" MkArr_Map a a' g ideE in_hom_char) moreover have "L ?f = g" proof - have "L ?f = trg (MkArr (Dom a') (Cod a) (Map g)) \ MkArr (Dom a') (Cod a) (Map g)" using f by auto also have "... = MkIde (E.Trg (Cod a)) \ MkArr (Dom a') (Cod a) (Map g)" using a a' f trg_def [of a] vconn_implies_hpar by auto also have "... = MkArr (E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Dom a') (E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Cod a) (B.can (E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Cod a) (E.Trg (Cod a) \<^bold>\ Cod a) \\<^sub>B (\E.Trg (Cod a)\ \\<^sub>B Map g) \\<^sub>B B.can (E.Trg (Cod a) \<^bold>\ Dom a') (E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Dom a'))" using hcomp_def apply simp by (metis (no_types, lifting) Cod.simps(1) arrE f in_homE src_trg trg.preserves_arr trg_def) also have "... = MkArr (Dom a') (Cod a) (B.can (Cod a) (E.Trg (Cod a) \<^bold>\ Cod a) \\<^sub>B (trg\<^sub>B \Cod a\ \\<^sub>B Map g) \\<^sub>B B.can (E.Trg (Cod a) \<^bold>\ Dom a') (Dom a'))" proof - have "E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Dom a' = Dom a'" using a a' arr_char E.HcompNml_Trg_Nml by (metis (no_types, lifting) f ideE trg_simps(1) vconn_implies_hpar(4)) moreover have "E.Trg (Cod a) \<^bold>\\<^bold>\\<^bold>\ Cod a = Cod a" using a a' arr_char E.HcompNml_Trg_Nml by blast moreover have "\E.Trg (Cod a)\ = trg\<^sub>B \Cod a\" using a a' arr_char E.eval_simps'(3) by fastforce ultimately show ?thesis by simp qed also have "... = MkArr (Dom a') (Cod a) (B.lunit \Cod a\ \\<^sub>B (trg\<^sub>B \Cod a\ \\<^sub>B Map g) \\<^sub>B B.lunit' \Dom a'\)" proof - have "E.Trg (Cod a) = E.Trg (Dom a')" using a a' a_eq g ide_char arr_char src_def trg_def trg_hcomp \Cod g = Dom a\ \Dom g = Dom a'\ by (metis (no_types, lifting) Cod.simps(1) in_homE) moreover have "B.can (Cod a) (E.Trg (Cod a) \<^bold>\ Cod a) = B.lunit \Cod a\" using a ide_char arr_char B.canE_unitor(2) by blast moreover have "B.can (E.Trg (Dom a') \<^bold>\ Dom a') (Dom a') = B.lunit' \Dom a'\" using a' ide_char arr_char B.canE_unitor(4) by blast ultimately show ?thesis by simp qed also have "... = MkArr (Dom g) (Cod g) (Map g)" proof - have "src\<^sub>B \Cod a\ = src\<^sub>B (Map g)" using a f g ide_char arr_char src_def B.comp_cod_arr by (metis (no_types, lifting) B.vconn_implies_hpar(1) B.vconn_implies_hpar(3) Cod.simps(1) Map.simps(1) in_homE) moreover have "B.lunit \Cod g\ \\<^sub>B (trg\<^sub>B (Map g) \\<^sub>B Map g) \\<^sub>B B.lunit' \Dom g\ = Map g" proof - have "B.lunit \Cod g\ \\<^sub>B (trg\<^sub>B (Map g) \\<^sub>B Map g) \\<^sub>B B.lunit' \Dom g\ = B.lunit \Cod g\ \\<^sub>B B.lunit' \Cod g\ \\<^sub>B Map g" using g ide_char arr_char B.lunit'_naturality by (metis (no_types, lifting) partial_magma_axioms B.in_homE partial_magma.arrI) also have "... = (B.lunit \Cod g\ \\<^sub>B B.lunit' \Cod g\) \\<^sub>B Map g" using B.comp_assoc by simp also have "... = \Cod g\ \\<^sub>B Map g" using g E.ide_eval_Ide B.comp_arr_inv' by fastforce also have "... = Map g" using g E.ide_eval_Ide B.comp_cod_arr by fastforce finally show ?thesis by simp qed ultimately have "B.lunit \Cod a\ \\<^sub>B (trg\<^sub>B \Cod a\ \\<^sub>B Map g) \\<^sub>B B.lunit' \Dom a'\ = Map g" using a a' 1 2 f g hcomp_def dom_char cod_char by (metis (no_types, lifting) B.comp_null(2) B.ext B.lunit_simps(2) B.lunit_simps(3) B.src.preserves_reflects_arr B.trg_vcomp B.vseq_implies_hpar(1) ideE) thus ?thesis using a 1 2 by auto qed also have "... = g" using g MkArr_Map by blast finally show ?thesis by simp qed ultimately show "\f. in_hom f a' a \ L f = g" by blast qed interpretation R: full_functor vcomp vcomp R proof fix a a' g assume a: "ide a" and a': "ide a'" assume g: "in_hom g (R a') (R a)" have a_eq: "a = MkIde (Dom a)" using a dom_char [of a] by simp have a'_eq: "a' = MkIde (Dom a')" using a' dom_char [of a'] by simp have 1: "Cod g = Dom a" using a g hcomp_ide_src in_hom_char by force have 2: "Dom g = Dom a'" using a' g hcomp_ide_src by auto let ?f = "MkArr (Dom a') (Cod a) (Map g)" have f: "in_hom ?f a' a" proof (intro in_homI) show 3: "arr (MkArr (Dom a') (Cod a) (Map g))" by (metis (no_types, lifting) "1" "2" Cod.simps(1) MkArr_Map a_eq g in_homE) show "dom (MkArr (Dom a') (Cod a) (Map g)) = a'" using a a' 3 dom_char by auto show "cod (MkArr (Dom a') (Cod a) (Map g)) = a" using a a' 3 cod_char by auto qed moreover have "R ?f = g" proof - have "R ?f = MkArr (Dom a') (Cod a) (Map g) \ src (MkArr (Dom a') (Cod a) (Map g))" using f by auto also have "... = MkArr (Dom a') (Cod a) (Map g) \ MkIde (E.Src (Cod a))" using a a' f src_def [of a] vconn_implies_hpar by auto also have "... = MkArr (Dom a' \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a)) (Cod a \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a)) (B.can (Cod a \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a)) (Cod a \<^bold>\ E.Src (Cod a)) \\<^sub>B (Map g \\<^sub>B \E.Src (Cod a)\) \\<^sub>B B.can (Dom a' \<^bold>\ E.Src (Cod a)) (Dom a' \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a)))" using hcomp_def apply simp by (metis (no_types, lifting) Cod_cod arrE f in_homE trg_src src.preserves_arr src_def) also have "... = MkArr (Dom a') (Cod a) (B.can (Cod a) (Cod a \<^bold>\ E.Src (Cod a)) \\<^sub>B (Map g \\<^sub>B src\<^sub>B \Cod a\) \\<^sub>B B.can (Dom a' \<^bold>\ E.Src (Cod a)) (Dom a'))" proof - have "Dom a' \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a) = Dom a'" using a a' arr_char E.HcompNml_Nml_Src by (metis (no_types, lifting) f ideE src_simps(1) vconn_implies_hpar(3)) moreover have "Cod a \<^bold>\\<^bold>\\<^bold>\ E.Src (Cod a) = Cod a" using a a' arr_char E.HcompNml_Nml_Src by blast moreover have "\E.Src (Cod a)\ = src\<^sub>B \Cod a\" using a a' arr_char E.eval_simps'(2) by fastforce ultimately show ?thesis by simp qed also have "... = MkArr (Dom a') (Cod a) (B.runit \Cod a\ \\<^sub>B (Map g \\<^sub>B src\<^sub>B \Cod a\) \\<^sub>B B.runit' \Dom a'\)" by (metis (no_types, lifting) B.canE_unitor(1) B.canE_unitor(3) a a' arrE f ideE src_simps(1) vconn_implies_hpar(3)) also have "... = MkArr (Dom g) (Cod g) (Map g)" proof - have "src\<^sub>B \Cod a\ = src\<^sub>B (Map g)" using a f g ide_char arr_char src_def B.comp_cod_arr by (metis (no_types, lifting) B.vconn_implies_hpar(1) B.vconn_implies_hpar(3) Cod.simps(1) Map.simps(1) in_homE) moreover have "B.runit \Cod g\ \\<^sub>B (Map g \\<^sub>B src\<^sub>B (Map g)) \\<^sub>B B.runit' \Dom g\ = Map g" proof - have "B.runit \Cod g\ \\<^sub>B (Map g \\<^sub>B src\<^sub>B (Map g)) \\<^sub>B B.runit' \Dom g\ = B.runit \Cod g\ \\<^sub>B B.runit'\Cod g\ \\<^sub>B Map g" using g ide_char arr_char B.runit'_naturality [of "Map g"] by (metis (no_types, lifting) partial_magma_axioms B.in_homE partial_magma.arrI) also have "... = (B.runit \Cod g\ \\<^sub>B B.runit' \Cod g\) \\<^sub>B Map g" using B.comp_assoc by simp also have "... = \Cod g\ \\<^sub>B Map g" using g E.ide_eval_Ide B.comp_arr_inv' by fastforce also have "... = Map g" using g E.ide_eval_Ide B.comp_cod_arr by fastforce finally show ?thesis by simp qed ultimately have "B.runit \Cod a\ \\<^sub>B (Map g \\<^sub>B src\<^sub>B \Cod a\) \\<^sub>B B.runit' \Dom a'\ = Map g" using a a' 1 2 f g hcomp_def dom_char cod_char by (metis (no_types, lifting) ideE) thus ?thesis using a 1 2 by auto qed also have "... = g" using g MkArr_Map by blast finally show ?thesis by simp qed ultimately show "\f. in_hom f a' a \ R f = g" by blast qed interpretation L: faithful_functor vcomp vcomp L proof fix f f' assume par: "par f f'" and eq: "L f = L f'" show "f = f'" proof (intro arr_eqI) have 1: "Dom f = Dom f' \ Cod f = Cod f'" using par dom_char cod_char by auto show "arr f" using par by simp show "arr f'" using par by simp show 2: "Dom f = Dom f'" and 3: "Cod f = Cod f'" using 1 by auto show "Map f = Map f'" proof - have "B.L (Map f) = trg\<^sub>B (Map f) \\<^sub>B Map f" using par by auto also have "... = trg\<^sub>B (Map f') \\<^sub>B Map f'" proof - have "\E.Trg (Dom f)\ \\<^sub>B Map f = \E.Trg (Dom f')\ \\<^sub>B Map f'" proof - have A: "\B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f) : \E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f\ \\<^sub>B \E.Trg (Dom f)\ \\<^sub>B \Dom f\\" using par arr_char B.can_in_hom E.Ide_HcompNml E.Ide_Nmlize_Ide E.Nml_Trg E.Nmlize_Nml E.HcompNml_Trg_Nml src_def trg_def by (metis (no_types, lifting) E.eval_simps(3) E.ide_eval_Ide E.Ide_implies_Arr B.canE_unitor(4) B.lunit'_in_vhom) have B: "\B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f) : \E.Trg (Dom f)\ \\<^sub>B \Cod f\ \\<^sub>B \E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f\\" using par arr_char B.can_in_hom E.Ide_HcompNml E.Ide_Nmlize_Ide E.Nml_Trg E.Nmlize_Nml E.HcompNml_Trg_Nml src_def trg_def by (metis (no_types, lifting) E.Nmlize.simps(3) E.eval.simps(3) E.Ide.simps(3) E.Ide_implies_Arr E.Src_Trg trg.preserves_arr trg_simps(2)) have C: "\\E.Trg (Dom f)\ \\<^sub>B Map f : \E.Trg (Dom f)\ \\<^sub>B \Dom f\ \\<^sub>B \E.Trg (Dom f)\ \\<^sub>B \Cod f\\" using par arr_char by (metis (no_types, lifting) E.eval_simps'(1) E.eval_simps(3) E.ide_eval_Ide E.Ide_implies_Arr E.Obj_Trg E.Obj_implies_Ide B.hcomp_in_vhom B.ide_in_hom(2) B.src_trg) have 3: "(\E.Trg (Dom f)\ \\<^sub>B Map f) \\<^sub>B B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f) = (\E.Trg (Dom f')\ \\<^sub>B Map f') \\<^sub>B B.can (E.Trg (Dom f') \<^bold>\ Dom f') (E.Trg (Dom f') \<^bold>\\<^bold>\\<^bold>\ Dom f')" proof - have 2: "B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f) \\<^sub>B (\E.Trg (Dom f)\ \\<^sub>B Map f) \\<^sub>B B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f) = B.can (E.Trg (Dom f') \<^bold>\\<^bold>\\<^bold>\ Cod f') (E.Trg (Dom f') \<^bold>\ Cod f') \\<^sub>B (\E.Trg (Dom f')\ \\<^sub>B Map f') \\<^sub>B B.can (E.Trg (Dom f') \<^bold>\ Dom f') (E.Trg (Dom f') \<^bold>\\<^bold>\\<^bold>\ Dom f')" using par eq hcomp_def trg_def src_trg trg.preserves_arr Map_hcomp trg_simps(1) trg_simps(2) trg_simps(3) by auto have "B.mono (B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f))" using par arr_char B.inverse_arrows_can B.iso_is_section B.section_is_mono src_def trg_def E.Nmlize_Nml E.HcompNml_Trg_Nml E.Ide_implies_Arr trg.preserves_arr trg_simps(1) by auto moreover have "B.seq (B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f)) ((\E.Trg (Dom f)\ \\<^sub>B Map f) \\<^sub>B B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f))" using A B C by auto moreover have "B.seq (B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f)) ((\E.Trg (Dom f')\ \\<^sub>B Map f') \\<^sub>B B.can (E.Trg (Dom f') \<^bold>\ Dom f') (E.Trg (Dom f') \<^bold>\\<^bold>\\<^bold>\ Dom f'))" using par 1 2 arr_char calculation(2) by auto moreover have "B.can (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Cod f) (E.Trg (Dom f) \<^bold>\ Cod f) = B.can (E.Trg (Dom f') \<^bold>\\<^bold>\\<^bold>\ Cod f') (E.Trg (Dom f') \<^bold>\ Cod f')" using par 1 arr_char by simp ultimately show ?thesis using 2 B.monoE cod_char by auto qed show ?thesis proof - have "B.epi (B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f))" using par arr_char B.inverse_arrows_can B.iso_is_retraction B.retraction_is_epi E.Nmlize_Nml E.HcompNml_Trg_Nml src_def trg_def E.Ide_implies_Arr by (metis (no_types, lifting) E.Nmlize.simps(3) E.Ide.simps(3) E.Src_Trg trg.preserves_arr trg_simps(1)) moreover have "B.seq (\E.Trg (Dom f)\ \\<^sub>B Map f) (B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f))" using A C by auto moreover have "B.seq (\E.Trg (Dom f')\ \\<^sub>B Map f') (B.can (E.Trg (Dom f) \<^bold>\ Dom f) (E.Trg (Dom f) \<^bold>\\<^bold>\\<^bold>\ Dom f))" using 1 3 calculation(2) by auto ultimately show ?thesis using par 1 3 arr_char B.epiE by simp qed qed moreover have "trg\<^sub>B (Map f) = \E.Trg (Dom f)\ \ trg\<^sub>B (Map f') = \E.Trg (Dom f')\" using par arr_char trg_def E.Ide_implies_Arr B.comp_arr_dom B.vseq_implies_hpar(2) E.eval_simps(3) by (metis (no_types, lifting) B.vconn_implies_hpar(2)) ultimately show ?thesis by simp qed also have "... = B.L (Map f')" using par B.hseqE B.hseq_char' by auto finally have "B.L (Map f) = B.L (Map f')" by simp thus ?thesis using 2 3 par arr_char B.L.is_faithful by (metis (no_types, lifting) B.in_homE) qed qed qed interpretation R: faithful_functor vcomp vcomp R proof fix f f' assume par: "par f f'" and eq: "R f = R f'" show "f = f'" proof (intro arr_eqI) have 1: "Dom f = Dom f' \ Cod f = Cod f'" using par dom_char cod_char by auto show "arr f" using par by simp show "arr f'" using par by simp show 2: "Dom f = Dom f'" and 3: "Cod f = Cod f'" using 1 by auto show "Map f = Map f'" proof - have "B.R (Map f) = Map f \\<^sub>B src\<^sub>B (Map f)" using par apply simp by (metis B.hseqE B.hseq_char') also have "... = Map f' \\<^sub>B src\<^sub>B (Map f')" proof - have "Map f \\<^sub>B \E.Src (Dom f)\ = Map f' \\<^sub>B \E.Src (Dom f')\" proof - have 2: "E.Ide (Cod f \<^bold>\ E.Src (Dom f))" using par arr_char src.preserves_arr by auto hence 3: "E.Ide (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f))" using par arr_char E.Nml_Src E.Ide_HcompNml calculation by auto have 4: "\<^bold>\Cod f \<^bold>\ E.Src (Dom f)\<^bold>\ = \<^bold>\Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)\<^bold>\" using par arr_char by (simp add: E.Nml_HcompNml(1)) have A: "\B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) : \Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)\ \\<^sub>B \Dom f\ \\<^sub>B \E.Src (Dom f)\\" using par arr_char B.can_in_hom E.Ide_HcompNml E.Ide_Nmlize_Ide E.Nml_Src E.Nmlize_Nml E.HcompNml_Nml_Src src_def trg_def by (metis (no_types, lifting) E.eval_simps(2) E.ide_eval_Ide E.Ide_implies_Arr B.canE_unitor(3) B.runit'_in_vhom) have B: "\B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f)) : \Cod f\ \\<^sub>B \E.Src (Dom f)\ \\<^sub>B \Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)\\" using 2 3 4 B.can_in_hom [of "Cod f \<^bold>\ E.Src (Dom f)" "Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)"] by simp have C: "\Map f \\<^sub>B \E.Src (Dom f)\ : \Dom f\ \\<^sub>B \E.Src (Dom f)\ \\<^sub>B \Cod f\ \\<^sub>B \E.Src (Dom f)\\" using par arr_char E.Ide_Nmlize_Ide E.Nml_Trg E.Nmlize_Nml E.HcompNml_Trg_Nml src_def trg_def E.ide_eval_Ide E.Ide_implies_Arr E.Obj_implies_Ide apply (intro B.hcomp_in_vhom) apply (simp add: B.ide_in_hom(2)) apply simp by (metis (no_types, lifting) A B.ideD(1) B.not_arr_null B.seq_if_composable B.src.preserves_reflects_arr B.vconn_implies_hpar(3) E.HcompNml_Nml_Src) have 5: "(Map f \\<^sub>B \E.Src (Dom f)\) \\<^sub>B B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) = (Map f' \\<^sub>B \E.Src (Dom f')\) \\<^sub>B B.can (Dom f' \<^bold>\ E.Src (Dom f')) (Dom f' \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f'))" proof - have 6: "B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f)) \\<^sub>B (Map f \\<^sub>B \E.Src (Dom f)\) \\<^sub>B B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) = B.can (Cod f' \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f')) (Cod f' \<^bold>\ E.Src (Dom f')) \\<^sub>B (Map f' \\<^sub>B \E.Src (Dom f')\) \\<^sub>B B.can (Dom f' \<^bold>\ E.Src (Dom f')) (Dom f' \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f'))" using par eq hcomp_def src_def trg_src src.preserves_arr Map_hcomp src_simps(1) src_simps(2) src_simps(3) by auto have "B.mono (B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f)))" using 2 3 4 B.inverse_arrows_can(1) B.iso_is_section B.section_is_mono by simp moreover have "B.seq (B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f))) ((Map f \\<^sub>B \E.Src (Dom f)\) \\<^sub>B B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)))" using A B C by auto moreover have "B.seq (B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f))) ((Map f' \\<^sub>B \E.Src (Dom f')\) \\<^sub>B B.can (Dom f' \<^bold>\ E.Src (Dom f')) (Dom f' \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f')))" using par 1 6 arr_char calculation(2) by auto moreover have "B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)) (Cod f \<^bold>\ E.Src (Dom f)) = B.can (Cod f' \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f')) (Cod f' \<^bold>\ E.Src (Dom f'))" using par 1 arr_char by simp ultimately show ?thesis using 6 B.monoE cod_char by auto qed show ?thesis proof - have "B.epi (B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)))" using 2 3 4 B.inverse_arrows_can(1) B.iso_is_retraction B.retraction_is_epi by (metis (no_types, lifting) E.Nml_Src E.Nmlize.simps(3) E.Nmlize_Nml E.HcompNml_Nml_Src E.Ide.simps(3) par arrE) moreover have "B.seq (Map f \\<^sub>B \E.Src (Dom f)\) (B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)))" using A C by auto moreover have "B.seq (Map f' \\<^sub>B \E.Src (Dom f')\) (B.can (Dom f \<^bold>\ E.Src (Dom f)) (Dom f \<^bold>\\<^bold>\\<^bold>\ E.Src (Dom f)))" using 1 5 calculation(2) by auto ultimately show ?thesis using par 1 5 arr_char B.epiE by simp qed qed moreover have "src\<^sub>B (Map f) = \E.Src (Dom f)\ \ src\<^sub>B (Map f') = \E.Src (Dom f')\" using par arr_char src_def by (metis (no_types, lifting) B.vconn_implies_hpar(1) E.Nml_implies_Arr E.eval_simps(2)) ultimately show ?thesis by simp qed also have "... = B.R (Map f')" using par B.hseqE B.hseq_char' by auto finally have "B.R (Map f) = B.R (Map f')" by simp thus ?thesis using 2 3 par arr_char B.R.is_faithful by (metis (no_types, lifting) B.in_homE) qed qed qed definition \ where "\ \ \ \ \ if VVV.arr (\, \, \) then hcomp \ (hcomp \ \) else null" interpretation natural_isomorphism VVV.comp vcomp HoHV HoVH \\\\\. \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))\ proof show "\\\\. \ VVV.arr \\\ \ \ (fst \\\) (fst (snd \\\)) (snd (snd \\\)) = null" using \_def by simp show "\\\\. VVV.arr \\\ \ dom (\ (fst \\\) (fst (snd \\\)) (snd (snd \\\))) = HoHV (VVV.dom \\\)" using VVV.arr_char VV.arr_char \_def hcomp_assoc HoHV_def VVV.dom_simp VV.dom_simp by force show 1: "\\\\. VVV.arr \\\ \ cod (\ (fst \\\) (fst (snd \\\)) (snd (snd \\\))) = HoVH (VVV.cod \\\)" using VVV.arr_char VV.arr_char \_def HoVH_def VVV.cod_simp VV.cod_simp by force show "\\\\. VVV.arr \\\ \ HoVH \\\ \ \ (fst (VVV.dom \\\)) (fst (snd (VVV.dom \\\))) (snd (snd (VVV.dom \\\))) = \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))" using \_def HoVH.as_nat_trans.is_natural_1 HoVH_def by auto show "\\\\. VVV.arr \\\ \ \ (fst (VVV.cod \\\)) (fst (snd (VVV.cod \\\))) (snd (snd (VVV.cod \\\))) \ HoHV \\\ = \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))" proof - fix \\\ assume \\\: "VVV.arr \\\" have "HoHV \\\ = \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))" unfolding \_def HoHV_def using \\\ HoHV.preserves_cod hcomp_assoc VVV.arr_char VV.arr_char by simp thus "\ (fst (VVV.cod \\\)) (fst (snd (VVV.cod \\\))) (snd (snd (VVV.cod \\\))) \ HoHV \\\ = \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))" using 1 \\\ comp_cod_arr \_def by (metis (no_types, lifting) HoVH_def HoHV.preserves_arr prod.collapse) qed show "\fgh. VVV.ide fgh \ iso (\ (fst fgh) (fst (snd fgh)) (snd (snd fgh)))" using \_def HoVH.preserves_ide HoVH_def by auto qed definition \ where "\ \ \a. a" sublocale bicategory vcomp hcomp \ \ src trg using hcomp_obj_self \_def hcomp_assoc VVV.arr_char VV.arr_char apply unfold_locales by (auto simp add: \_def ide_in_hom(2)) lemma is_bicategory: shows "bicategory vcomp hcomp \ \ src trg" .. sublocale strict_bicategory vcomp hcomp \ \ src trg proof show "\fgh. ide fgh \ lunit fgh = fgh" proof - fix fgh assume fgh: "ide fgh" have "fgh = lunit fgh" proof (intro lunit_eqI) show "ide fgh" using fgh by simp show "\fgh : trg fgh \ fgh \ fgh\" using fgh hcomp_def hcomp_trg_ide by auto show "trg fgh \ fgh = (\ (trg fgh) \ fgh) \ \' (trg fgh) (trg fgh) fgh" proof - have "(\ (trg fgh) \ fgh) \ \' (trg fgh) (trg fgh) fgh = (trg fgh \ fgh) \ \' (trg fgh) (trg fgh) fgh" using fgh \_def by metis also have "... = (trg fgh \ fgh) \ (trg fgh \ trg fgh \ fgh)" using fgh \_def by fastforce also have "... = trg fgh \ fgh" using fgh hcomp_obj_self hcomp_assoc by (simp add: hcomp_trg_ide) finally show ?thesis by simp qed qed thus "lunit fgh = fgh" by simp qed show "\fgh. ide fgh \ runit fgh = fgh" proof - fix fgh assume fgh: "ide fgh" have "fgh = runit fgh" proof (intro runit_eqI) show "ide fgh" using fgh by simp show "\fgh : fgh \ src fgh \ fgh\" using fgh hcomp_def hcomp_ide_src by auto show "fgh \ src fgh = (fgh \ \ (src fgh)) \ \ fgh (src fgh) (src fgh)" proof - have "(fgh \ \ (src fgh)) \ \ fgh (src fgh) (src fgh) = (fgh \ src fgh) \ \ fgh (src fgh) (src fgh)" using fgh \_def by metis also have "... = (fgh \ src fgh) \ (fgh \ src fgh \ src fgh)" using fgh \_def by fastforce also have "... = fgh \ src fgh" using fgh comp_arr_dom hcomp_obj_self by simp finally show ?thesis by simp qed qed thus "runit fgh = fgh" by simp qed show "\f g h. \ ide f; ide g; ide h; src f = trg g; src g = trg h \ \ ide (\ f g h)" using \_def VV.arr_char VVV.arr_char by auto qed theorem is_strict_bicategory: shows "strict_bicategory vcomp hcomp \ \ src trg" .. lemma iso_char: shows "iso \ \ arr \ \ B.iso (Map \)" and "iso \ \ inv \ = MkArr (Cod \) (Dom \) (B.inv (Map \))" proof - have 1: "iso \ \ arr \ \ B.iso (Map \)" proof - assume \: "iso \" obtain \ where \: "inverse_arrows \ \" using \ by auto have "B.inverse_arrows (Map \) (Map \)" proof show "B.ide (Map \ \\<^sub>B Map \)" proof - have "Map \ \\<^sub>B Map \ = Map (\ \ \)" using \ \ inverse_arrows_def Map_comp arr_char seq_char by (metis (no_types, lifting) ide_compE) moreover have "B.ide ..." using \ ide_char by blast ultimately show ?thesis by simp qed show "B.ide (Map \ \\<^sub>B Map \)" proof - have "Map \ \\<^sub>B Map \ = Map (\ \ \)" using \ \ inverse_arrows_def comp_char [of \ \] by simp moreover have "B.ide ..." using \ ide_char by blast ultimately show ?thesis by simp qed qed thus "arr \ \ B.iso (Map \)" using \ by auto qed let ?\ = "MkArr (Cod \) (Dom \) (B.inv (Map \))" have 2: "arr \ \ B.iso (Map \) \ iso \ \ inv \ = ?\" proof assume \: "arr \ \ B.iso (Map \)" have \: "\?\ : cod \ \ dom \\" using \ arr_char dom_char cod_char by auto have 4: "inverse_arrows \ ?\" proof show "ide (?\ \ \)" proof - have "?\ \ \ = dom \" using \ \ MkArr_Map comp_char seq_char B.comp_inv_arr' dom_char by auto thus ?thesis using \ by simp qed show "ide (\ \ ?\)" proof - have "\ \ ?\ = cod \" using \ \ MkArr_Map comp_char seq_char B.comp_arr_inv' cod_char by auto thus ?thesis using \ by simp qed qed thus "iso \" by auto show "inv \ = ?\" using 4 inverse_unique by simp qed have 3: "arr \ \ B.iso (Map \) \ iso \" using 2 by simp show "iso \ \ arr \ \ B.iso (Map \)" using 1 3 by blast show "iso \ \ inv \ = ?\" using 1 2 by blast qed subsection "The Strictness Theorem" text \ The Strictness Theorem asserts: ``Every bicategory is biequivalent to a strict bicategory.'' This amounts to an equivalent (and perhaps more desirable) formulation of the Coherence Theorem. In this section we prove the Strictness Theorem by constructing an equivalence pseudofunctor from a bicategory to its strictification. \ text \ We define a map \UP\ from the given bicategory \B\ to its strictification, and show that it is an equivalence pseudofunctor. The following auxiliary definition is not logically necessary, but it provides some terms that can be the targets of simplification rules and thereby enables some proofs to be done by simplification that otherwise could not be. Trying to eliminate it breaks some short proofs below, so I have kept it. \ definition UP\<^sub>0 where "UP\<^sub>0 a \ if B.obj a then MkIde \<^bold>\a\<^bold>\\<^sub>0 else null" lemma obj_UP\<^sub>0 [simp]: assumes "B.obj a" shows "obj (UP\<^sub>0 a)" unfolding obj_def using assms UP\<^sub>0_def ide_MkIde [of "\<^bold>\a\<^bold>\\<^sub>0"] src_def by simp lemma UP\<^sub>0_in_hom [intro]: assumes "B.obj a" shows "\UP\<^sub>0 a : UP\<^sub>0 a \ UP\<^sub>0 a\" and "\UP\<^sub>0 a : UP\<^sub>0 a \ UP\<^sub>0 a\" using assms obj_UP\<^sub>0 by blast+ lemma UP\<^sub>0_simps [simp]: assumes "B.obj a" shows "ide (UP\<^sub>0 a)" "arr (UP\<^sub>0 a)" and "src (UP\<^sub>0 a) = UP\<^sub>0 a" and "trg (UP\<^sub>0 a) = UP\<^sub>0 a" and "dom (UP\<^sub>0 a) = UP\<^sub>0 a" and "cod (UP\<^sub>0 a) = UP\<^sub>0 a" using assms obj_UP\<^sub>0 apply blast using assms obj_UP\<^sub>0 obj_simps by simp_all definition UP where "UP \ \ if B.arr \ then MkArr \<^bold>\B.dom \\<^bold>\ \<^bold>\B.cod \\<^bold>\ \ else null" lemma Dom_UP [simp]: assumes "B.arr \" shows "Dom (UP \) = \<^bold>\B.dom \\<^bold>\" using assms UP_def by fastforce lemma Cod_UP [simp]: assumes "B.arr \" shows "Cod (UP \) = \<^bold>\B.cod \\<^bold>\" using assms UP_def by fastforce lemma Map_UP [simp]: assumes "B.arr \" shows "Map (UP \) = \" using assms UP_def by fastforce lemma arr_UP: assumes "B.arr \" shows "arr (UP \)" using assms UP_def by (intro arrI, fastforce+) lemma UP_in_hom [intro]: assumes "B.arr \" shows "\UP \ : UP\<^sub>0 (src\<^sub>B \) \ UP\<^sub>0 (trg\<^sub>B \)\" and "\UP \ : UP (B.dom \) \ UP (B.cod \)\" using assms arr_UP UP_def UP\<^sub>0_def dom_char cod_char src_def trg_def by auto lemma UP_simps [simp]: assumes "B.arr \" shows "arr (UP \)" and "src (UP \) = UP\<^sub>0 (src\<^sub>B \)" and "trg (UP \) = UP\<^sub>0 (trg\<^sub>B \)" and "dom (UP \) = UP (B.dom \)" and "cod (UP \) = UP (B.cod \)" using assms arr_UP UP_in_hom by auto interpretation UP: "functor" V\<^sub>B vcomp UP using UP_def arr_UP UP_simps(4-5) arr_UP UP_def comp_char seq_char by unfold_locales auto interpretation UP: weak_arrow_of_homs V\<^sub>B src\<^sub>B trg\<^sub>B vcomp src trg UP proof fix \ assume \: "B.arr \" show "isomorphic (UP (src\<^sub>B \)) (src (UP \))" proof - let ?\ = "MkArr \<^bold>\src\<^sub>B \\<^bold>\ \<^bold>\src\<^sub>B \\<^bold>\\<^sub>0 (src\<^sub>B \)" have \: "\?\ : UP (src\<^sub>B \) \ src (UP \)\" proof show 1: "arr ?\" using \ by (intro arrI, auto) show "dom ?\ = UP (src\<^sub>B \)" using \ 1 dom_char UP_def by simp show "cod ?\ = src (UP \)" using \ 1 cod_char src_def by auto qed have "iso ?\" using \ \ iso_char src_def by auto thus ?thesis using \ isomorphic_def by auto qed show "isomorphic (UP (trg\<^sub>B \)) (trg (UP \))" proof - let ?\ = "MkArr \<^bold>\trg\<^sub>B \\<^bold>\ \<^bold>\trg\<^sub>B \\<^bold>\\<^sub>0 (trg\<^sub>B \)" have \: "\?\ : UP (trg\<^sub>B \) \ trg (UP \)\" proof show 1: "arr ?\" using \ by (intro arrI, auto) show "dom ?\ = UP (trg\<^sub>B \)" using \ 1 dom_char UP_def by simp show "cod ?\ = trg (UP \)" using \ 1 cod_char trg_def by auto qed have "iso ?\" using \ \ iso_char trg_def by auto thus ?thesis using \ isomorphic_def by auto qed qed interpretation HoUP_UP: composite_functor B.VV.comp VV.comp vcomp UP.FF \\\\. hcomp (fst \\) (snd \\)\ .. interpretation UPoH: composite_functor B.VV.comp V\<^sub>B vcomp \\\\. fst \\ \\<^sub>B snd \\\ UP .. abbreviation \\<^sub>o where "\\<^sub>o fg \ MkArr (\<^bold>\fst fg\<^bold>\ \<^bold>\ \<^bold>\snd fg\<^bold>\) \<^bold>\fst fg \\<^sub>B snd fg\<^bold>\ (fst fg \\<^sub>B snd fg)" interpretation \: transformation_by_components B.VV.comp vcomp HoUP_UP.map UPoH.map \\<^sub>o proof fix fg assume fg: "B.VV.ide fg" show "\\\<^sub>o fg : HoUP_UP.map fg \ UPoH.map fg\" proof (intro in_homI) show 1: "arr (\\<^sub>o fg)" using fg arr_char B.VV.ide_char B.VV.arr_char by auto show "dom (\\<^sub>o fg) = HoUP_UP.map fg" using 1 fg UP.FF_def B.VV.arr_char B.VV.ide_char dom_char hcomp_def B.can_Ide_self by simp show "cod (\\<^sub>o fg) = UPoH.map fg" using fg arr_char cod_char B.VV.ide_char B.VV.arr_char UP_def by auto qed next fix \\ assume \\: "B.VV.arr \\" show "\\<^sub>o (B.VV.cod \\) \ HoUP_UP.map \\ = UPoH.map \\ \ \\<^sub>o (B.VV.dom \\)" proof - have "\\<^sub>o (B.VV.cod \\) \ HoUP_UP.map \\ = MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) (fst \\ \\<^sub>B snd \\)" proof - have "\\<^sub>o (B.VV.cod \\) \ HoUP_UP.map \\ = MkArr (\<^bold>\B.cod (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.cod (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) (B.cod (fst \\) \\<^sub>B B.cod (snd \\)) \ MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.cod (snd \\)\<^bold>\) (fst \\ \\<^sub>B snd \\)" using \\ B.VV.arr_char arr_char UP.FF_def hcomp_def UP_def B.VV.cod_simp src_def trg_def B.can_in_hom B.can_Ide_self B.comp_arr_dom B.comp_cod_arr by auto also have "... = MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) ((B.cod (fst \\) \\<^sub>B B.cod (snd \\)) \\<^sub>B (fst \\ \\<^sub>B snd \\))" using \\ B.VV.arr_char by (intro comp_MkArr arr_MkArr, auto) also have "... = MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) (fst \\ \\<^sub>B snd \\)" using \\ B.VV.arr_char B.comp_cod_arr by auto finally show ?thesis by simp qed also have "... = UPoH.map \\ \ \\<^sub>o (B.VV.dom \\)" proof - have "UPoH.map \\ \ \\<^sub>o (B.VV.dom \\) = MkArr (\<^bold>\B.dom (fst \\) \\<^sub>B B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) (fst \\ \\<^sub>B snd \\) \ MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.dom (fst \\) \\<^sub>B B.dom (snd \\)\<^bold>\) (B.dom (fst \\) \\<^sub>B B.dom (snd \\))" using \\ B.VV.arr_char arr_char UP.FF_def hcomp_def UP_def B.VV.dom_simp src_def trg_def B.can_in_hom B.can_Ide_self B.comp_arr_dom B.comp_cod_arr by auto also have "... = MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) ((fst \\ \\<^sub>B snd \\) \\<^sub>B (B.dom (fst \\) \\<^sub>B B.dom (snd \\)))" using \\ B.VV.arr_char arr_MkArr apply (intro comp_MkArr arr_MkArr) by auto also have "... = MkArr (\<^bold>\B.dom (fst \\)\<^bold>\ \<^bold>\ \<^bold>\B.dom (snd \\)\<^bold>\) (\<^bold>\B.cod (fst \\) \\<^sub>B B.cod (snd \\)\<^bold>\) (fst \\ \\<^sub>B snd \\)" using \\ B.VV.arr_char B.comp_arr_dom by auto finally show ?thesis by simp qed finally show ?thesis by simp qed qed abbreviation cmp\<^sub>U\<^sub>P where "cmp\<^sub>U\<^sub>P \ \.map" lemma cmp\<^sub>U\<^sub>P_in_hom [intro]: assumes "B.arr (fst \\)" and "B.arr (snd \\)" and "src\<^sub>B (fst \\) = trg\<^sub>B (snd \\)" shows "\cmp\<^sub>U\<^sub>P \\ : UP\<^sub>0 (src\<^sub>B (snd \\)) \ UP\<^sub>0 (trg\<^sub>B (fst \\))\" and "\cmp\<^sub>U\<^sub>P \\ : UP (B.dom (fst \\)) \ UP (B.dom (snd \\)) \ UP (B.cod (fst \\) \\<^sub>B B.cod (snd \\))\" proof - let ?\ = "fst \\" and ?\ = "snd \\" show 1: "\cmp\<^sub>U\<^sub>P \\ : UP (B.dom ?\) \ UP (B.dom ?\) \ UP (B.cod ?\ \\<^sub>B B.cod ?\)\" proof show "arr (cmp\<^sub>U\<^sub>P \\)" using assms by auto show "dom (cmp\<^sub>U\<^sub>P \\) = UP (B.dom ?\) \ UP (B.dom ?\)" proof - have "B.VV.in_hom (?\, ?\) (B.dom ?\, B.dom ?\) (B.cod ?\, B.cod ?\)" using assms B.VV.in_hom_char B.VV.arr_char by auto hence "dom (cmp\<^sub>U\<^sub>P \\) = HoUP_UP.map (B.dom ?\, B.dom ?\)" by auto also have "... = UP (B.dom ?\) \ UP (B.dom ?\)" using assms UP.FF_def by fastforce finally show ?thesis by simp qed show "cod (cmp\<^sub>U\<^sub>P \\) = UP (B.cod ?\ \\<^sub>B B.cod ?\)" using assms B.VV.in_hom_char B.VV.arr_char B.VV.cod_simp by auto qed show "\cmp\<^sub>U\<^sub>P \\ : UP\<^sub>0 (src\<^sub>B ?\) \ UP\<^sub>0 (trg\<^sub>B ?\)\" using assms 1 src_dom [of "cmp\<^sub>U\<^sub>P \\"] trg_dom [of "cmp\<^sub>U\<^sub>P \\"] by fastforce qed lemma cmp\<^sub>U\<^sub>P_simps [simp]: assumes "B.arr (fst \\)" and "B.arr (snd \\)" and "src\<^sub>B (fst \\) = trg\<^sub>B (snd \\)" shows "arr (cmp\<^sub>U\<^sub>P \\)" and "src (cmp\<^sub>U\<^sub>P \\) = UP\<^sub>0 (src\<^sub>B (snd \\))" and "trg (cmp\<^sub>U\<^sub>P \\) = UP\<^sub>0 (trg\<^sub>B (fst \\))" and "dom (cmp\<^sub>U\<^sub>P \\) = UP (B.dom (fst \\)) \ UP (B.dom (snd \\))" and "cod (cmp\<^sub>U\<^sub>P \\) = UP (B.cod (fst \\) \\<^sub>B B.cod (snd \\))" using assms cmp\<^sub>U\<^sub>P_in_hom [of \\] by auto lemma cmp\<^sub>U\<^sub>P_ide_simps [simp]: assumes "B.ide (fst fg)" and "B.ide (snd fg)" and "src\<^sub>B (fst fg) = trg\<^sub>B (snd fg)" shows "Dom (cmp\<^sub>U\<^sub>P fg) = \<^bold>\fst fg\<^bold>\ \<^bold>\ \<^bold>\snd fg\<^bold>\" and "Cod (cmp\<^sub>U\<^sub>P fg) = \<^bold>\fst fg \\<^sub>B snd fg\<^bold>\" and "Map (cmp\<^sub>U\<^sub>P fg) = fst fg \\<^sub>B snd fg" using assms B.VV.ide_char B.VV.arr_char by auto interpretation \: natural_isomorphism B.VV.comp vcomp HoUP_UP.map UPoH.map cmp\<^sub>U\<^sub>P proof fix fg assume fg: "B.VV.ide fg" have "arr (cmp\<^sub>U\<^sub>P fg)" using fg \.preserves_reflects_arr [of fg] by simp thus "iso (cmp\<^sub>U\<^sub>P fg)" using fg iso_char by simp qed lemma cmp\<^sub>U\<^sub>P_ide_simp: assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g" shows "cmp\<^sub>U\<^sub>P (f, g) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g)" using assms B.VV.ide_char B.VV.arr_char by simp lemma cmp\<^sub>U\<^sub>P'_ide_simp: assumes "B.ide f" and "B.ide g" and "src\<^sub>B f = trg\<^sub>B g" shows "inv (cmp\<^sub>U\<^sub>P (f, g)) = MkArr \<^bold>\f \\<^sub>B g\<^bold>\ (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (f \\<^sub>B g)" using assms cmp\<^sub>U\<^sub>P_ide_simp iso_char \.components_are_iso [of "(f, g)"] B.VV.ide_char B.VV.arr_char by simp interpretation UP: pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \ \ src trg UP cmp\<^sub>U\<^sub>P proof fix f g h assume f: "B.ide f" and g: "B.ide g" and h: "B.ide h" and fg: "src\<^sub>B f = trg\<^sub>B g" and gh: "src\<^sub>B g = trg\<^sub>B h" show "UP \\<^sub>B[f, g, h] \ cmp\<^sub>U\<^sub>P (f \\<^sub>B g, h) \ (cmp\<^sub>U\<^sub>P (f, g) \ UP h) = cmp\<^sub>U\<^sub>P (f, g \\<^sub>B h) \ (UP f \ cmp\<^sub>U\<^sub>P (g, h)) \ \ (UP f) (UP g) (UP h)" proof - have "UP \\<^sub>B[f, g, h] \ cmp\<^sub>U\<^sub>P (f \\<^sub>B g, h) \ (cmp\<^sub>U\<^sub>P (f, g) \ UP h) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h)" proof - have 1: "UP.hseq\<^sub>D (MkIde \<^bold>\f\<^bold>\) (MkIde \<^bold>\g\<^bold>\)" using f g fg hseq_char src_def trg_def arr_char by auto have 2: "UP.hseq\<^sub>D (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g) \ MkIde (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\)) (MkIde \<^bold>\h\<^bold>\)" proof - have "MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g) \ MkIde (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g)" proof - have "MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g) \ MkIde (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (f \\<^sub>B g)" using f g fg by simp also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ ((f \\<^sub>B g) \\<^sub>B (f \\<^sub>B g))" using f g fg by (intro comp_MkArr arr_MkArr, auto) also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g)" using f g fg by simp finally show ?thesis by blast qed thus ?thesis using f g h fg gh arr_char src_def trg_def by auto qed have "UP \\<^sub>B[f, g, h] = MkArr \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ \\<^sub>B[f, g, h]" using f g h fg gh UP_def B.HoHV_def B.HoVH_def B.VVV.arr_char B.VV.arr_char B.VVV.dom_char B.VVV.cod_char by simp moreover have "cmp\<^sub>U\<^sub>P (f \\<^sub>B g, h) = MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ ((f \\<^sub>B g) \\<^sub>B h) \ MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) ((f \\<^sub>B g) \\<^sub>B h)" using f g h fg gh \.map_simp_ide \.map_def B.VV.arr_char UP.FF_def B.VV.cod_simp hcomp_def B.can_Ide_self by simp moreover have "cmp\<^sub>U\<^sub>P (f, g) \ UP h = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (B.inv \\<^sub>B[f, g, h])" proof - have "MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (B.can (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \\<^sub>B (f \\<^sub>B g) \\<^sub>B B.can (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\)) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (f \\<^sub>B g)" using f g fg B.can_Ide_self B.comp_arr_dom B.comp_cod_arr by simp moreover have "MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) (f \\<^sub>B g) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\f \\<^sub>B g\<^bold>\ (f \\<^sub>B g)" by (metis (no_types, lifting) 2 B.ideD(1) E.eval.simps(2-3) cod_MkArr comp_arr_ide f g ide_char' seq_char) moreover have "B.can ((\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\) \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) = B.inv \\<^sub>B[f, g, h]" using f g h fg gh B.canI_associator_0 B.inverse_arrows_can by simp ultimately show ?thesis using 1 2 f g h fg gh \.map_def UP_def hcomp_def UP.FF_def B.VV.arr_char B.can_Ide_self B.comp_cod_arr B.VV.cod_simp by simp qed ultimately have "UP \\<^sub>B[f, g, h] \ cmp\<^sub>U\<^sub>P (f \\<^sub>B g, h) \ (cmp\<^sub>U\<^sub>P (f, g) \ UP h) = MkArr \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ \\<^sub>B[f, g, h] \ MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ ((f \\<^sub>B g) \\<^sub>B h) \ MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) ((f \\<^sub>B g) \\<^sub>B h) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (B.inv \\<^sub>B[f, g, h])" using comp_assoc by presburger also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (\\<^sub>B[f, g, h] \\<^sub>B ((f \\<^sub>B g) \\<^sub>B h) \\<^sub>B ((f \\<^sub>B g) \\<^sub>B h) \\<^sub>B B.inv \\<^sub>B[f, g, h])" proof - have "Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (B.inv \\<^sub>B[f, g, h])) \ Arr (MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) ((f \\<^sub>B g) \\<^sub>B h)) \ Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (((f \\<^sub>B g) \\<^sub>B h) \\<^sub>B B.inv \\<^sub>B[f, g, h])) \ Arr (MkArr (\<^bold>\f \\<^sub>B g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ ((f \\<^sub>B g) \\<^sub>B h)) \ Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ (((f \\<^sub>B g) \\<^sub>B h) \\<^sub>B ((f \\<^sub>B g) \\<^sub>B h) \\<^sub>B B.inv \\<^sub>B[f, g, h])) \ Arr (MkArr \<^bold>\(f \\<^sub>B g) \\<^sub>B h\<^bold>\ \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ \\<^sub>B[f, g, h])" using f g h fg gh B.\.preserves_hom B.HoHV_def B.HoVH_def by auto thus ?thesis using f g h fg gh comp_def B.comp_arr_dom B.comp_cod_arr by simp qed also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h)" using B.comp_cod_arr B.comp_arr_inv' by (auto simp add: f fg g gh h) finally show ?thesis by simp qed also have "... = cmp\<^sub>U\<^sub>P (f, g \\<^sub>B h) \ (UP f \ cmp\<^sub>U\<^sub>P (g, h)) \ \ (UP f) (UP g) (UP h)" proof - have "cmp\<^sub>U\<^sub>P (f, g \\<^sub>B h) \ (UP f \ cmp\<^sub>U\<^sub>P (g, h)) \ \ (UP f) (UP g) (UP h) = cmp\<^sub>U\<^sub>P (f, g \\<^sub>B h) \ (MkIde \<^bold>\f\<^bold>\ \ cmp\<^sub>U\<^sub>P (g, h)) \ (MkIde \<^bold>\f\<^bold>\ \ MkIde \<^bold>\g\<^bold>\ \ MkIde \<^bold>\h\<^bold>\)" using f g h fg gh VVV.arr_char VV.arr_char arr_char src_def trg_def UP_def \_def by auto also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)" proof - have "cmp\<^sub>U\<^sub>P (f, g \\<^sub>B h) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h) \ MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)" using f g h fg gh \.map_simp_ide \.map_def UP.FF_def UP_def hcomp_def B.VV.arr_char B.can_Ide_self B.comp_arr_dom B.comp_cod_arr src_def trg_def arr_char B.VV.cod_simp by auto moreover have "cmp\<^sub>U\<^sub>P (g, h) = MkArr (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\g \\<^sub>B h\<^bold>\ (g \\<^sub>B h)" using g h gh cmp\<^sub>U\<^sub>P_ide_simp by blast moreover have "MkIde \<^bold>\f\<^bold>\ \ MkArr (\<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\g \\<^sub>B h\<^bold>\ (g \\<^sub>B h) = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)" using f g h fg gh hcomp_def arr_char src_def trg_def B.can_Ide_self B.comp_arr_dom B.comp_cod_arr by auto moreover have "MkIde \<^bold>\f\<^bold>\ \ MkIde \<^bold>\g\<^bold>\ \ MkIde \<^bold>\h\<^bold>\ = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)" proof - have "\f : f \\<^sub>B f\ \ \g : g \\<^sub>B g\ \ \h : h \\<^sub>B h\" using f g h by auto thus ?thesis using f g h fg gh hcomp_def arr_char src_def trg_def B.can_Ide_self B.comp_arr_dom B.comp_cod_arr by auto qed ultimately show ?thesis using comp_assoc by auto qed also have "... = MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h)" proof - have "Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)) \ Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g\<^bold>\ \<^bold>\ \<^bold>\h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)) \ Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) (f \\<^sub>B g \\<^sub>B h)) \ Arr (MkArr (\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\g \\<^sub>B h\<^bold>\) \<^bold>\f \\<^sub>B g \\<^sub>B h\<^bold>\ (f \\<^sub>B g \\<^sub>B h))" using f g h fg gh by auto thus ?thesis using f g h fg gh comp_def by auto qed finally show ?thesis by simp qed finally show ?thesis by blast qed qed lemma UP_is_pseudofunctor: shows "pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \ \ src trg UP cmp\<^sub>U\<^sub>P" .. lemma UP_map\<^sub>0_obj [simp]: assumes "B.obj a" shows "UP.map\<^sub>0 a = UP\<^sub>0 a" using assms UP.map\<^sub>0_def by auto interpretation UP: full_functor V\<^sub>B vcomp UP proof fix \ f g assume f: "B.ide f" and g: "B.ide g" assume \: "\\ : UP f \ UP g\" show "\\. \\ : f \\<^sub>B g\ \ UP \ = \" proof - have 1: "\Map \ : f \\<^sub>B g\" using f g \ UP_def arr_char in_hom_char by auto moreover have "UP (Map \) = \" proof - have "\ = MkArr (Dom \) (Cod \) (Map \)" using \ MkArr_Map by auto also have "... = UP (Map \)" using "1" B.arrI UP.as_nat_trans.preserves_hom UP_def \ in_hom_char by force finally show ?thesis by auto qed ultimately show ?thesis by blast qed qed interpretation UP: faithful_functor V\<^sub>B vcomp UP using arr_char UP_def by (unfold_locales, simp_all) interpretation UP: fully_faithful_functor V\<^sub>B vcomp UP .. lemma UP_is_fully_faithful_functor: shows "fully_faithful_functor V\<^sub>B vcomp UP" .. no_notation B.in_hom ("\_ : _ \\<^sub>B _\") (* Inherited from functor, I think. *) lemma Map_reflects_hhom: assumes "B.obj a" and "B.obj b" and "ide g" and "\g : UP.map\<^sub>0 a \ UP.map\<^sub>0 b\" shows "\Map g : a \\<^sub>B b\" proof have 1: "B.ide (Map g)" using assms ide_char by blast show "B.arr (Map g)" using 1 by simp show "src\<^sub>B (Map g) = a" proof - have "src\<^sub>B (Map g) = Map (src g)" using assms src_def apply simp by (metis (no_types, lifting) E.eval_simps(2) E.Ide_implies_Arr arr_char ideE) also have "... = Map (UP.map\<^sub>0 a)" using assms by (metis (no_types, lifting) in_hhomE) also have "... = a" using assms UP.map\<^sub>0_def UP_def [of a] src_def by auto finally show ?thesis by simp qed show "trg\<^sub>B (Map g) = b" proof - have "trg\<^sub>B (Map g) = Map (trg g)" using assms trg_def apply simp by (metis (no_types, lifting) E.eval_simps(3) E.Ide_implies_Arr arr_char ideE) also have "... = Map (UP.map\<^sub>0 b)" using assms by (metis (no_types, lifting) in_hhomE) also have "... = b" using assms UP.map\<^sub>0_def UP_def [of b] src_def by auto finally show ?thesis by simp qed qed lemma eval_Dom_ide [simp]: assumes "ide g" shows "\Dom g\ = Map g" using assms dom_char ideD by auto lemma Cod_ide: assumes "ide f" shows "Cod f = Dom f" using assms dom_char by auto lemma Map_preserves_objects: assumes "obj a" shows "B.obj (Map a)" proof - have "src\<^sub>B (Map a) = Map (src a)" using assms src_def apply simp using E.eval_simps(2) E.Ide_implies_Arr arr_char ideE by (metis (no_types, lifting) objE) also have 1: "... = \E.Src (Dom a)\" using assms src_def by auto also have "... = \\<^bold>\Map a\<^bold>\\<^sub>0\" using assms B.src.is_extensional 1 obj_simps(2) by force also have "... = Map a" using assms by auto finally have "src\<^sub>B (Map a) = Map a" by simp moreover have "B.arr (Map a)" using assms B.ideD arr_char by auto ultimately show ?thesis using B.obj_def by simp qed interpretation UP: equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \ \ src trg UP cmp\<^sub>U\<^sub>P proof (* UP is full, hence locally full. *) show "\f f' \. \ B.ide f; B.ide f'; src\<^sub>B f = src\<^sub>B f'; trg\<^sub>B f = trg\<^sub>B f'; \\ : UP f \ UP f'\ \ \ \\. \\ : f \\<^sub>B f'\ \ UP \ = \" using UP.is_full by simp (* UP is biessentially surjective on objects. *) show "\b. obj b \ \a. B.obj a \ equivalent_objects (UP.map\<^sub>0 a) b" proof - fix b assume b: "obj b" have 1: "B.obj (Map b)" using b Map_preserves_objects by simp have 3: "UP.map\<^sub>0 (Map b) = MkArr \<^bold>\Map b\<^bold>\\<^sub>0 \<^bold>\Map b\<^bold>\\<^sub>0 (Map b)" using b 1 UP.map\<^sub>0_def [of "Map b"] UP_def src_def arr_char by auto have 4: "b = MkArr (Dom b) (Dom b) (Map b)" using b objE eval_Dom_ide by (metis (no_types, lifting) dom_char ideD(2) obj_def) let ?\ = "MkArr \<^bold>\Map b\<^bold>\\<^sub>0 (Dom b) (Map b)" have \: "arr ?\" proof - have 2: "E.Obj (Dom b)" using b obj_char by blast have "E.Nml \<^bold>\Map b\<^bold>\\<^sub>0 \ E.Ide \<^bold>\Map b\<^bold>\\<^sub>0" using 1 by auto moreover have "E.Nml (Dom b) \ E.Ide (Dom b)" using b arr_char [of b] by auto moreover have "E.Src \<^bold>\Map b\<^bold>\\<^sub>0 = E.Src (Dom b)" using b 1 2 apply (cases "Dom b") apply simp_all by fastforce moreover have "E.Trg \<^bold>\Map b\<^bold>\\<^sub>0 = E.Trg (Dom b)" using b 1 2 apply (cases "Dom b") apply simp_all by fastforce moreover have "\Map b : \\<^bold>\Map b\<^bold>\\<^sub>0\ \\<^sub>B \Dom b\\" using b 1 by (elim objE, auto) ultimately show ?thesis using arr_char \E.Nml \<^bold>\Map b\<^bold>\\<^sub>0 \ E.Ide \<^bold>\Map b\<^bold>\\<^sub>0\ by auto qed hence "iso ?\" using 1 iso_char by auto moreover have "dom ?\ = UP.map\<^sub>0 (Map b)" using \ dom_char b 1 3 B.objE UP.map\<^sub>0_def UP_def src_def by auto moreover have "cod ?\ = b" using \ cod_char b 4 1 by auto ultimately have "isomorphic (UP.map\<^sub>0 (Map b)) b" using \ 3 4 isomorphic_def by blast moreover have 5: "obj (UP.map\<^sub>0 (Map b))" using 1 UP.map\<^sub>0_simps(2) by simp ultimately have 6: "UP.map\<^sub>0 (Map b) = b" using b isomorphic_objects_are_equal by simp have "equivalent_objects (UP.map\<^sub>0 (Map b)) b" using b 6 equivalent_objects_reflexive [of b] by simp thus "\a. B.obj a \ equivalent_objects (UP.map\<^sub>0 a) b" using 1 6 by auto qed (* UP is locally essentially surjective. *) show "\a b g. \ B.obj a; B.obj b; \g : UP.map\<^sub>0 a \ UP.map\<^sub>0 b\; ide g \ \ \f. \f : a \\<^sub>B b\ \ B.ide f \ isomorphic (UP f) g" proof - fix a b g assume a: "B.obj a" and b: "B.obj b" assume g_in_hhom: "\g : UP.map\<^sub>0 a \ UP.map\<^sub>0 b\" assume ide_g: "ide g" have 1: "B.ide (Map g)" using ide_g ide_char by blast have "arr (UP a)" using a by auto have "arr (UP b)" using b by auto have Map_g_eq: "Map g = \Dom g\" using ide_g by simp have Map_g_in_hhom: "\Map g : a \\<^sub>B b\" using a b ide_g g_in_hhom Map_reflects_hhom by simp let ?\ = "MkArr \<^bold>\Map g\<^bold>\ (Dom g) (Map g)" have \: "arr ?\" proof - have "\Map ?\ : \Dom ?\\ \\<^sub>B \Cod ?\\\" using 1 Map_g_eq by auto moreover have "E.Ide \<^bold>\Map g\<^bold>\ \ E.Nml \<^bold>\Map g\<^bold>\" using 1 by simp moreover have "E.Ide (Dom g) \ E.Nml (Dom g)" using ide_g arr_char ide_char by blast moreover have "E.Src \<^bold>\Map g\<^bold>\ = E.Src (Dom g)" using ide_g g_in_hhom src_def Map_g_in_hhom by (metis (no_types, lifting) B.ideD(2) B.in_hhom_def B.objE B.obj_def' Dom.simps(1) E.Src.simps(2) UP.map\<^sub>0_def \arr (UP a)\ a in_hhomE UP_def) moreover have "E.Trg \<^bold>\Map g\<^bold>\ = E.Trg (Dom g)" proof - have "E.Trg \<^bold>\Map g\<^bold>\ = \<^bold>\b\<^bold>\\<^sub>0" using Map_g_in_hhom by auto also have "... = E.Trg (Dom g)" proof - have "E.Trg (Dom g) = Dom (trg g)" using ide_g trg_def by simp also have "... = Dom (UP.map\<^sub>0 b)" using g_in_hhom by auto also have "... = \<^bold>\b\<^bold>\\<^sub>0" using b \arr (UP b)\ UP.map\<^sub>0_def src_def UP_def B.objE by auto finally show ?thesis by simp qed finally show ?thesis by simp qed ultimately show ?thesis using arr_char by simp qed have "\?\ : UP (Map g) \ g\" by (simp add: "1" \ ide_g in_hom_char) moreover have "iso ?\" using \ 1 iso_char by simp ultimately have "isomorphic (UP (Map g)) g" using isomorphic_def by auto thus "\f. \f : a \\<^sub>B b\ \ B.ide f \ isomorphic (UP f) g" using 1 Map_g_in_hhom by auto qed qed theorem UP_is_equivalence_pseudofunctor: shows "equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \ \ src trg UP cmp\<^sub>U\<^sub>P" .. text \ Next, we work out the details of the equivalence pseudofunctor \DN\ in the converse direction. \ definition DN where "DN \ \ if arr \ then Map \ else B.null" lemma DN_in_hom [intro]: assumes "arr \" shows "\DN \ : DN (src \) \\<^sub>B DN (trg \)\" and "\DN \ : DN (dom \) \\<^sub>B DN (cod \)\" using assms DN_def arr_char [of \] B.vconn_implies_hpar(1-2) E.eval_in_hom(1) B.in_hhom_def by auto lemma DN_simps [simp]: assumes "arr \" shows "B.arr (DN \)" and "src\<^sub>B (DN \) = DN (src \)" and "trg\<^sub>B (DN \) = DN (trg \)" and "B.dom (DN \) = DN (dom \)" and "B.cod (DN \) = DN (cod \)" using assms DN_in_hom by auto interpretation "functor" vcomp V\<^sub>B DN using DN_def seqE Map_comp seq_char by unfold_locales auto interpretation DN: weak_arrow_of_homs vcomp src trg V\<^sub>B src\<^sub>B trg\<^sub>B DN proof fix \ assume \: "arr \" show "B.isomorphic (DN (src \)) (src\<^sub>B (DN \))" proof - have "DN (src \) = src\<^sub>B (DN \)" using B.src.is_extensional DN_def DN_simps(2) by auto moreover have "B.ide (DN (src \))" using \ by simp ultimately show ?thesis using \ B.isomorphic_reflexive by auto qed show "B.isomorphic (DN (trg \)) (trg\<^sub>B (DN \))" proof - have "DN (trg \) = trg\<^sub>B (DN \)" using \B.isomorphic (DN (src \)) (src\<^sub>B (DN \))\ by fastforce moreover have "B.ide (DN (trg \))" using \ by simp ultimately show ?thesis using B.isomorphic_reflexive by auto qed qed interpretation "functor" VV.comp B.VV.comp DN.FF using DN.functor_FF by auto interpretation HoDN_DN: composite_functor VV.comp B.VV.comp V\<^sub>B DN.FF \\\\. H\<^sub>B (fst \\) (snd \\)\ .. interpretation DNoH: composite_functor VV.comp vcomp V\<^sub>B \\\\. fst \\ \ snd \\\ DN .. abbreviation \\<^sub>o where "\\<^sub>o fg \ B.can (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\ Dom (snd fg))" abbreviation \\<^sub>o' where "\\<^sub>o' fg \ B.can (Dom (fst fg) \<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg))" lemma \\<^sub>o_in_hom: assumes "VV.ide fg" shows "\\\<^sub>o fg : Map (fst fg) \\<^sub>B Map (snd fg) \\<^sub>B \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\\" and "\\\<^sub>o' fg : \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\ \\<^sub>B Map (fst fg) \\<^sub>B Map (snd fg)\" and "B.inverse_arrows (\\<^sub>o fg) (\\<^sub>o' fg)" proof - have 1: "E.Ide (Dom (fst fg) \<^bold>\ Dom (snd fg))" unfolding E.Ide.simps(3) apply (intro conjI) using assms VV.ide_char VV.arr_char arr_char apply simp using VV.arr_char VV.ideD(1) assms apply blast by (metis (no_types, lifting) VV.arrE VV.ideD(1) assms src_simps(1) trg_simps(1)) have 2: "E.Ide (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg))" using 1 by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char VV.ideD(1) arr_char assms) have 3: "\<^bold>\Dom (fst fg) \<^bold>\ Dom (snd fg)\<^bold>\ = \<^bold>\Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\<^bold>\" by (metis (no_types, lifting) E.Ide.simps(3) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml VV.arr_char VV.ideD(1) arr_char assms 1) have 4: "\Dom (fst fg) \<^bold>\ Dom (snd fg)\ = Map (fst fg) \\<^sub>B Map (snd fg)" using assms VV.ide_char VV.arr_char arr_char by simp show "\\\<^sub>o fg : Map (fst fg) \\<^sub>B Map (snd fg) \\<^sub>B \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\\" using 1 2 3 4 by auto show "\\\<^sub>o' fg : \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\ \\<^sub>B Map (fst fg) \\<^sub>B Map (snd fg)\" using 1 2 3 4 by auto show "B.inverse_arrows (\\<^sub>o fg) (\\<^sub>o' fg)" using 1 2 3 B.inverse_arrows_can by blast qed interpretation \: transformation_by_components VV.comp V\<^sub>B HoDN_DN.map DNoH.map \\<^sub>o proof fix fg assume fg: "VV.ide fg" have 1: "\Dom (fst fg) \<^bold>\ Dom (snd fg)\ = Map (fst fg) \\<^sub>B Map (snd fg)" using fg VV.ide_char VV.arr_char arr_char by simp show "\\\<^sub>o fg : HoDN_DN.map fg \\<^sub>B DNoH.map fg\" proof show "B.arr (\\<^sub>o fg)" using fg \\<^sub>o_in_hom by blast show "B.dom (\\<^sub>o fg) = HoDN_DN.map fg" proof - have "B.dom (\\<^sub>o fg) = Map (fst fg) \\<^sub>B Map (snd fg)" using fg \\<^sub>o_in_hom by blast also have "... = HoDN_DN.map fg" using fg DN.FF_def DN_def VV.arr_char src_def trg_def VV.ide_char by auto finally show ?thesis by simp qed show "B.cod (\\<^sub>o fg) = DNoH.map fg" proof - have "B.cod (\\<^sub>o fg) = \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\" using fg \\<^sub>o_in_hom by blast also have "... = DNoH.map fg" proof - have "DNoH.map fg = B.can (Cod (fst fg) \<^bold>\\<^bold>\\<^bold>\ Cod (snd fg)) (Cod (fst fg) \<^bold>\ Cod (snd fg)) \\<^sub>B (Map (fst fg) \\<^sub>B Map (snd fg)) \\<^sub>B B.can (Dom (fst fg) \<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg))" using fg DN_def Map_hcomp VV.arr_char apply simp using VV.ideD(1) by blast also have "... = B.can (Cod (fst fg) \<^bold>\\<^bold>\\<^bold>\ Cod (snd fg)) (Cod (fst fg) \<^bold>\ Cod (snd fg)) \\<^sub>B B.can (Dom (fst fg) \<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg))" proof - have "(Map (fst fg) \\<^sub>B Map (snd fg)) \\<^sub>B B.can (Dom (fst fg) \<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)) = B.can (Dom (fst fg) \<^bold>\ Dom (snd fg)) (Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg))" using fg 1 \\<^sub>o_in_hom B.comp_cod_arr by blast thus ?thesis by simp qed also have "... = \Dom (fst fg) \<^bold>\\<^bold>\\<^bold>\ Dom (snd fg)\" proof - have "B.can (Cod (fst fg) \<^bold>\\<^bold>\\<^bold>\ Cod (snd fg)) (Cod (fst fg) \<^bold>\ Cod (snd fg)) = \\<^sub>o fg" using fg VV.ide_char Cod_ide by simp thus ?thesis using fg 1 \\<^sub>o_in_hom [of fg] B.comp_arr_inv' by fastforce qed finally show ?thesis by simp qed finally show ?thesis by blast qed qed next show "\f. VV.arr f \ \\<^sub>o (VV.cod f) \\<^sub>B HoDN_DN.map f = DNoH.map f \\<^sub>B \\<^sub>o (VV.dom f)" proof - fix \\ assume \\: "VV.arr \\" show "\\<^sub>o (VV.cod \\) \\<^sub>B HoDN_DN.map \\ = DNoH.map \\ \\<^sub>B \\<^sub>o (VV.dom \\)" proof - have 1: "E.Ide (Dom (fst \\) \<^bold>\ Dom (snd \\))" unfolding E.Ide.simps(3) - apply (intro conjI) - using \\ VV.ide_char VV.arr_char arr_char - apply simp - using VV.arr_char VV.ideD(1) \\ - apply blast - by (metis (no_types, lifting) VV.arrE \\ src_simps(1) trg_simps(1)) + by (metis (no_types, lifting) VV.arrE \\ arrE src_simps(2) trg_simps(2)) have 2: "E.Ide (Dom (fst \\) \<^bold>\\<^bold>\\<^bold>\ Dom (snd \\))" using 1 by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char VV.ideD(1) arr_char \\) have 3: "\<^bold>\Dom (fst \\) \<^bold>\ Dom (snd \\)\<^bold>\ = \<^bold>\Dom (fst \\) \<^bold>\\<^bold>\\<^bold>\ Dom (snd \\)\<^bold>\" by (metis (no_types, lifting) E.Ide.simps(3) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml VV.arr_char arr_char \\ 1) have 4: "E.Ide (Cod (fst \\) \<^bold>\ Cod (snd \\))" unfolding E.Ide.simps(3) - apply (intro conjI) - using \\ VV.ide_char VV.arr_char arr_char - apply simp - using VV.arr_char VV.ideD(1) \\ - apply blast - by (metis (no_types, lifting) "1" E.Ide.simps(3) VV.arrE \\ arrE) + by (metis (no_types, lifting) "1" E.Ide.simps(3) VV.arrE \\ arr_char) have 5: "E.Ide (Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\))" using 4 by (meson E.Ide.simps(3) E.Ide_HcompNml VV.arr_char VV.ideD(1) arr_char \\) have 6: "\<^bold>\Cod (fst \\) \<^bold>\ Cod (snd \\)\<^bold>\ = \<^bold>\Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)\<^bold>\" by (metis (no_types, lifting) E.Ide.simps(3) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml VV.arr_char arr_char \\ 1) have A: "\\\<^sub>o' \\ : \Dom (fst \\) \<^bold>\\<^bold>\\<^bold>\ Dom (snd \\)\ \\<^sub>B \Dom (fst \\) \<^bold>\ Dom (snd \\)\\" using 1 2 3 by auto have B: "\Map (fst \\) \\<^sub>B Map (snd \\) : \Dom (fst \\) \<^bold>\ Dom (snd \\)\ \\<^sub>B \Cod (fst \\) \<^bold>\ Cod (snd \\)\\" using \\ VV.arr_char arr_char src_def trg_def E.Nml_implies_Arr E.eval_simps'(2-3) by auto have C: "\B.can (Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)) (Cod (fst \\) \<^bold>\ Cod (snd \\)) : \Cod (fst \\) \<^bold>\ Cod (snd \\)\ \\<^sub>B \Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)\\" using 4 5 6 by auto have "\\<^sub>o (VV.cod \\) \\<^sub>B HoDN_DN.map \\ = B.can (Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)) (Cod (fst \\) \<^bold>\ Cod (snd \\)) \\<^sub>B (Map (fst \\) \\<^sub>B Map (snd \\))" using \\ VV.arr_char VV.cod_char arr_char src_def trg_def cod_char DN.FF_def DN_def by auto also have "... = B.can (Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)) (Cod (fst \\) \<^bold>\ Cod (snd \\)) \\<^sub>B (Map (fst \\) \\<^sub>B Map (snd \\)) \\<^sub>B \\<^sub>o' \\ \\<^sub>B \\<^sub>o \\" using B \\ VV.arr_char arr_char src_def trg_def E.Ide_HcompNml E.Nml_HcompNml(1) B.can_Ide_self B.comp_arr_dom by auto also have "... = DNoH.map \\ \\<^sub>B \\<^sub>o (VV.dom \\)" proof - have "DNoH.map \\ \\<^sub>B \\<^sub>o (VV.dom \\) = B.can (Cod (fst \\) \<^bold>\\<^bold>\\<^bold>\ Cod (snd \\)) (Cod (fst \\) \<^bold>\ Cod (snd \\)) \\<^sub>B (Map (fst \\) \\<^sub>B Map (snd \\)) \\<^sub>B \\<^sub>o' \\ \\<^sub>B \\<^sub>o (VV.dom \\)" using \\ DN_def VV.arr_char B.comp_assoc by simp moreover have "\\<^sub>o (VV.dom \\) = \\<^sub>o \\" using \\ VV.dom_char VV.arr_char by auto ultimately show ?thesis using B.comp_assoc by simp qed finally show ?thesis by blast qed qed qed abbreviation cmp\<^sub>D\<^sub>N where "cmp\<^sub>D\<^sub>N \ \.map" interpretation \: natural_isomorphism VV.comp V\<^sub>B HoDN_DN.map DNoH.map cmp\<^sub>D\<^sub>N using \\<^sub>o_in_hom B.iso_def \.map_simp_ide apply unfold_locales apply auto by blast no_notation B.in_hom ("\_ : _ \\<^sub>B _\") lemma cmp\<^sub>D\<^sub>N_in_hom [intro]: assumes "arr (fst \\)" and "arr (snd \\)" and "src (fst \\) = trg (snd \\)" shows "\cmp\<^sub>D\<^sub>N \\ : DN (src (snd \\)) \\<^sub>B DN (trg (fst \\))\" and "\cmp\<^sub>D\<^sub>N \\ : DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\)) \\<^sub>B DN (cod (fst \\) \ cod (snd \\))\" proof - have 1: "VV.arr \\" using assms VV.arr_char by simp show 2: "\cmp\<^sub>D\<^sub>N \\ : DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\)) \\<^sub>B DN (cod (fst \\) \ cod (snd \\))\" proof - have "HoDN_DN.map (VV.dom \\) = DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\))" using assms 1 DN.FF_def VV.dom_simp by auto moreover have "DNoH.map (VV.cod \\) = DN (cod (fst \\) \ cod (snd \\))" using assms 1 VV.cod_simp by simp ultimately show ?thesis using assms 1 \.preserves_hom by auto qed show "\cmp\<^sub>D\<^sub>N \\ : DN (src (snd \\)) \\<^sub>B DN (trg (fst \\))\" using assms 2 B.src_dom [of "cmp\<^sub>D\<^sub>N \\"] B.trg_dom [of "cmp\<^sub>D\<^sub>N \\"] by (elim B.in_homE, intro B.in_hhomI) auto qed lemma cmp\<^sub>D\<^sub>N_simps [simp]: assumes "arr (fst \\)" and "arr (snd \\)" and "src (fst \\) = trg (snd \\)" shows "B.arr (cmp\<^sub>D\<^sub>N \\)" and "src\<^sub>B (cmp\<^sub>D\<^sub>N \\) = DN (src (snd \\))" and "trg\<^sub>B (cmp\<^sub>D\<^sub>N \\) = DN (trg (fst \\))" and "B.dom (cmp\<^sub>D\<^sub>N \\) = DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\))" and "B.cod (cmp\<^sub>D\<^sub>N \\) = DN (cod (fst \\) \ cod (snd \\))" proof show "VV.arr \\" using assms by blast have 1: "\cmp\<^sub>D\<^sub>N \\ : DN (src (snd \\)) \\<^sub>B DN (trg (fst \\))\" using assms by blast show "src\<^sub>B (cmp\<^sub>D\<^sub>N \\) = DN (src (snd \\))" using 1 by fast show "trg\<^sub>B (cmp\<^sub>D\<^sub>N \\) = DN (trg (fst \\))" using 1 by fast have 2: "\cmp\<^sub>D\<^sub>N \\ : DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\)) \\<^sub>B DN (cod (fst \\) \ cod (snd \\))\" using assms by blast show "B.dom (cmp\<^sub>D\<^sub>N \\) = DN (dom (fst \\)) \\<^sub>B DN (dom (snd \\))" using 2 by fast show "B.cod (cmp\<^sub>D\<^sub>N \\) = DN (cod (fst \\) \ cod (snd \\))" using 2 by fast qed interpretation DN: pseudofunctor vcomp hcomp \ \ src trg V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B DN cmp\<^sub>D\<^sub>N proof show "\f g h. \ ide f; ide g; ide h; src f = trg g; src g = trg h \ \ DN (\ f g h) \\<^sub>B cmp\<^sub>D\<^sub>N (f \ g, h) \\<^sub>B (cmp\<^sub>D\<^sub>N (f, g) \\<^sub>B DN h) = cmp\<^sub>D\<^sub>N (f, g \ h) \\<^sub>B (DN f \\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) \\<^sub>B \\<^sub>B[DN f, DN g, DN h]" proof - fix f g h assume f: "ide f" and g: "ide g" and h: "ide h" and fg: "src f = trg g" and gh: "src g = trg h" show "DN (\ f g h) \\<^sub>B cmp\<^sub>D\<^sub>N (f \ g, h) \\<^sub>B (cmp\<^sub>D\<^sub>N (f, g) \\<^sub>B DN h) = cmp\<^sub>D\<^sub>N (f, g \ h) \\<^sub>B (DN f \\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) \\<^sub>B \\<^sub>B[DN f, DN g, DN h]" proof - have 1: "E.Trg (Dom g) = E.Trg (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \ \E.Trg (Dom g)\ = \E.Trg (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)\" using f g h fg gh arr_char src_def trg_def E.Trg_HcompNml by (metis (no_types, lifting) ideD(1) src_simps(2) trg_simps(2)) have 2: "arr (MkArr (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) \\<^sub>B (Map f \\<^sub>B B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)))" proof - have "\B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) \\<^sub>B (Map f \\<^sub>B B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : EVAL (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B EVAL (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)\" proof (intro B.comp_in_homI) show 2: "\B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : EVAL (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B EVAL (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)\" using f g h fg gh 1 apply (intro B.can_in_hom) apply (metis (no_types, lifting) E.Ide_HcompNml E.Nml_HcompNml(1) arr_char ideD(1) src_simps(1) trg_simps(1)) apply (metis (no_types, lifting) E.Ide.simps(3) E.Ide_HcompNml ideD(1) arr_char src_simps(1) trg_simps(1)) by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml ideD(1) arr_char src_simps(1) trg_simps(1)) show "\B.can (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) : EVAL (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) \\<^sub>B EVAL (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)\" proof - have "E.Ide (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" using f g h fg gh 1 Cod_ide E.Ide_HcompNml arr_char apply simp by (metis (no_types, lifting) ideD(1) src_simps(1) trg_simps(1)) moreover have "E.Ide (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" using f g h fg gh by (metis (no_types, lifting) E.Ide.simps(3) E.Ide_HcompNml E.Nml_HcompNml(1) arr_char calculation ideD(1) src_simps(1) trg_simps(1)) moreover have "E.Nmlize (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) = E.Nmlize (Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" using f g h fg gh by (metis (no_types, lifting) E.Ide.simps(3) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml arr_char calculation(1) ideD(1) src_simps(1) trg_simps(1)) ultimately show ?thesis using B.can_in_hom [of "Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h" "Cod f \<^bold>\\<^bold>\\<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h"] by blast qed show "\Map f \\<^sub>B B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : EVAL (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B EVAL (Cod f \<^bold>\ Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)\" using f g h fg gh B.can_in_hom apply simp proof (intro B.hcomp_in_vhom B.comp_in_homI) show 1: "\B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : EVAL (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B EVAL (Dom g \<^bold>\ Dom h)\" using g h gh B.can_in_hom by (metis (no_types, lifting) E.Ide.simps(3) E.Ide_HcompNml E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml arr_char ideD(1) src_simps(1) trg_simps(1)) show "\B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) : EVAL (Cod g \<^bold>\ Cod h) \\<^sub>B EVAL (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)\" using g h gh B.can_in_hom by (metis (no_types, lifting) Cod_ide E.Ide.simps(3) E.Ide_HcompNml E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml arr_char ideD(1) src_simps(2) trg_simps(2)) show "\Map g \\<^sub>B Map h : EVAL (Dom g \<^bold>\ Dom h) \\<^sub>B EVAL (Cod g \<^bold>\ Cod h)\" proof show 2: "B.hseq (Map g) (Map h)" using g h gh by (metis (no_types, lifting) B.comp_null(1-2) B.hseq_char' B.ideD(1) Map_hcomp ideE ide_hcomp) show "B.dom (Map g \\<^sub>B Map h) = EVAL (Dom g \<^bold>\ Dom h)" using g h gh 2 by fastforce show "B.cod (Map g \\<^sub>B Map h) = EVAL (Cod g \<^bold>\ Cod h)" using g h gh 2 by fastforce qed show "\Map f : Map f \\<^sub>B EVAL (Cod f)\" using f arr_char Cod_ide by auto show "src\<^sub>B (Map f) = trg\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\" using f g h fg gh 1 2 src_def trg_def B.arrI B.hseqE B.not_arr_null B.trg.is_extensional B.trg.preserves_hom B.vconn_implies_hpar(2) B.vconn_implies_hpar(4) E.eval.simps(3) by (metis (no_types, lifting) Map_ide(1)) qed qed thus ?thesis using f g h fg gh arr_char src_def trg_def E.Nml_HcompNml E.Ide_HcompNml ideD(1) apply (intro arr_MkArr) by auto qed have 3: "E.Ide (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using g h gh ide_char arr_char src_def trg_def E.Ide_HcompNml Cod_ide by (metis (no_types, lifting) ideD(1) src_simps(2) trg_simps(2)) have 4: "E.Ide (Dom g \<^bold>\ Dom h)" by (metis (no_types, lifting) E.Ide.simps(3) arrE g gh h ideE src_simps(1) trg_simps(1)) have 5: "E.Nmlize (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Nmlize (Dom g \<^bold>\ Dom h)" using g h gh ide_char arr_char src_def trg_def E.Nml_HcompNml by (metis (no_types, lifting) 4 E.Ide.simps(3) E.Nmlize.simps(3) E.Nmlize_Nml ideD(1)) have 6: "E.Ide (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h fg gh arr_char src_def trg_def by (metis (no_types, lifting) 1 E.Nml_HcompNml(1) E.Ide_HcompNml ideD(1) src_simps(2) trg_simps(2)) have 7: "E.Ide (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h fg gh arr_char src_def trg_def by (metis (no_types, lifting) 1 3 E.Ide.simps(3) ideD(1) src_simps(2) trg_simps(2)) have 8: "E.Nmlize (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Nmlize (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h fg gh arr_char src_def trg_def 7 E.Nml_HcompNml(1) ideD(1) by auto have "DN (\ f g h) \\<^sub>B cmp\<^sub>D\<^sub>N (f \ g, h) \\<^sub>B (cmp\<^sub>D\<^sub>N (f, g) \\<^sub>B DN h) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have 9: "VVV.arr (f, g, h)" using f g h fg gh VVV.arr_char VV.arr_char arr_char ideD by simp have 10: "VV.ide (f, g)" using f g fg VV.ide_char by auto have 11: "VV.ide (hcomp f g, h)" using f g h fg gh VV.ide_char VV.arr_char by simp have 12: "arr (MkArr (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)))" proof (intro arr_MkArr) show "Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h \ IDE" using g h gh by (metis (no_types, lifting) 3 E.Nml_HcompNml(1) arr_char ideD(1) mem_Collect_eq src_simps(2) trg_simps(2)) show "Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h \ IDE" using g h gh Cod_ide \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h \ IDE\ by auto show "B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \ HOM (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" proof show "E.Src (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Src (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) \ E.Trg (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Trg (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) \ \B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\ \\<^sub>B \Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h\\" proof (intro conjI) show "E.Src (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Src (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" using g h gh Cod_ide by simp show "E.Trg (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Trg (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h)" using g h gh Cod_ide by simp show "\B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\ \\<^sub>B \Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h\\" proof (intro B.comp_in_homI) show "\B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\ \\<^sub>B \Dom g \<^bold>\ Dom h\\" using 3 4 5 by blast show "\Map g \\<^sub>B Map h : \Dom g \<^bold>\ Dom h\ \\<^sub>B \Cod g \<^bold>\ Cod h\\" using g h gh by (metis (no_types, lifting) 4 B.ide_in_hom(2) Cod_ide E.eval.simps(3) E.ide_eval_Ide Map_ide(2)) show "\B.can (Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h) (Cod g \<^bold>\ Cod h) : \Cod g \<^bold>\ Cod h\ \\<^sub>B \Cod g \<^bold>\\<^bold>\\<^bold>\ Cod h\\" using 3 4 5 Cod_ide g h by auto qed qed qed qed have "DN (\ f g h) = \Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\" proof - have "DN (\ f g h) = (B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B ((Map f \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h))) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h))" using f g h fg gh 1 2 9 12 DN_def \_def Cod_ide by simp also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B (Map f \\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" proof - have "\B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\ \\<^sub>B Map g \\<^sub>B Map h\" using g h 3 4 5 B.can_in_hom [of "Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h" "Dom g \<^bold>\ Dom h"] by simp hence "Map f \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = Map f \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using B.comp_cod_arr by auto also have "... = Map f \\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\" using f g h fg gh 3 4 5 B.can_Ide_self by auto finally have "Map f \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h) \\<^sub>B (Map g \\<^sub>B Map h) \\<^sub>B B.can (Dom g \<^bold>\ Dom h) (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = Map f \\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\" by simp thus ?thesis by simp qed also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" proof - have "\B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) : \Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\ \\<^sub>B Map f \\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\\" using f g h 6 7 8 B.can_in_hom [of "Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h" "Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h"] by simp hence "(Map f \\<^sub>B \Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) = B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using B.comp_cod_arr by auto thus ?thesis by simp qed also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h fg gh 6 7 8 by auto also have "... = \Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\" using f g h fg gh 6 B.can_Ide_self by blast finally show ?thesis by simp qed have "DN (\ f g h) \\<^sub>B cmp\<^sub>D\<^sub>N (f \ g, h) \\<^sub>B (cmp\<^sub>D\<^sub>N (f, g) \\<^sub>B DN h) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) \\<^sub>B (B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) (Dom f \<^bold>\ Dom g) \\<^sub>B Map h)" proof - have "DN (\ f g h) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h fg gh DN_def 1 4 6 7 B.can_Ide_self E.HcompNml_assoc E.Ide.simps(3) \DN (\ f g h) = \Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h\\ ide_char by (metis (no_types, lifting) arr_char ideD(1)) thus ?thesis using f g h fg gh 10 11 DN_def \.map_simp_ide by simp qed also have "... = (B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h)) \\<^sub>B (B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) (Dom f \<^bold>\ Dom g) \\<^sub>B Map h)" using B.comp_assoc by simp also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) \\<^sub>B B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have "B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) (Dom f \<^bold>\ Dom g) \\<^sub>B Map h = B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have "B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) (Dom f \<^bold>\ Dom g) \\<^sub>B Map h = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) (Dom f \<^bold>\ Dom g) \\<^sub>B B.can (Dom h) (Dom h)" using h B.can_Ide_self by fastforce also have "... = B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" using f g h 1 4 7 arr_char E.Nml_HcompNml(1) E.Src_HcompNml B.hcomp_can [of "Dom f \<^bold>\ Dom g" "Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g" "Dom h" "Dom h"] by (metis (no_types, lifting) E.Nmlize.simps(3) E.Nmlize_Nml E.Ide.simps(3) E.Ide_HcompNml E.Src.simps(3) arrE ideD(1)) finally show ?thesis by simp qed moreover have "B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B B.can ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have "E.Ide ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h)" using f g h 1 4 7 by (metis (no_types, lifting) E.Ide.simps(3) E.Ide_HcompNml E.Src_HcompNml arrE ideD(1)) moreover have "E.Ide ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 1 7 E.Ide_HcompNml E.Nml_HcompNml(1) arr_char calculation ideD(1) by auto moreover have "E.Ide (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 1 4 6 by blast moreover have "E.Nmlize ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) = E.Nmlize ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 1 4 7 by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml E.Ide.simps(3) arrE calculation(1) ideD(1)) moreover have "E.Nmlize ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\\<^bold>\\<^bold>\ Dom h) = E.Nmlize (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 1 4 7 E.HcompNml_assoc by fastforce ultimately show ?thesis using B.vcomp_can by simp qed ultimately show ?thesis by simp qed also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have "E.Ide ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" using 1 4 7 by simp moreover have "E.Ide ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h)" using f g 1 4 7 by (metis (no_types, lifting) E.Ide.simps(3) E.Ide_HcompNml E.Src_HcompNml arrE ideD(1)) moreover have "E.Ide (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 6 by blast moreover have "E.Nmlize ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h) = E.Nmlize ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h)" using f g h 1 7 E.Nml_HcompNml(1) by fastforce moreover have "E.Nmlize ((Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g) \<^bold>\ Dom h) = E.Nmlize (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h)" using f g h 1 4 7 by (metis (no_types, lifting) E.Nml_HcompNml(1) E.Nmlize.simps(3) E.Nmlize_Nml E.HcompNml_assoc E.Ide.simps(3) arrE ideD(1)) ultimately show ?thesis using B.vcomp_can by simp qed finally show ?thesis by simp qed also have "... = cmp\<^sub>D\<^sub>N (f, g \ h) \\<^sub>B (DN f \\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) \\<^sub>B \\<^sub>B[DN f, DN g, DN h]" proof - have "cmp\<^sub>D\<^sub>N (f, g \ h) \\<^sub>B (DN f \\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) \\<^sub>B \\<^sub>B[DN f, DN g, DN h] = (cmp\<^sub>D\<^sub>N (f, g \ h) \\<^sub>B (DN f \\<^sub>B cmp\<^sub>D\<^sub>N (g, h))) \\<^sub>B \\<^sub>B[DN f, DN g, DN h]" using B.comp_assoc by simp also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\ Dom h) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" proof - have "cmp\<^sub>D\<^sub>N (f, g \ h) \\<^sub>B (DN f \\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\ Dom h)" proof - have "cmp\<^sub>D\<^sub>N (f, g \ h) \\<^sub>B (DN f \\<^sub>B cmp\<^sub>D\<^sub>N (g, h)) = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B (Map f \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h))" using f g h fg gh VV.ide_char VV.arr_char DN_def by simp also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B (B.can (Dom f) (Dom f) \\<^sub>B B.can (Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom g \<^bold>\ Dom h))" proof - have "Map f = B.can (Dom f) (Dom f)" using f arr_char B.can_Ide_self [of "Dom f"] Map_ide by (metis (no_types, lifting) ide_char') thus ?thesis by simp qed also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) \\<^sub>B B.can (Dom f \<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\ Dom h)" using 1 4 5 7 B.hcomp_can by auto also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) (Dom f \<^bold>\ Dom g \<^bold>\ Dom h)" using 1 4 5 6 7 8 B.vcomp_can by auto finally show ?thesis by simp qed moreover have "\\<^sub>B[DN f, DN g, DN h] = B.can (Dom f \<^bold>\ Dom g \<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" using f g h 1 4 7 DN_def B.canE_associator(1) by auto ultimately show ?thesis by simp qed also have "... = B.can (Dom f \<^bold>\\<^bold>\\<^bold>\ Dom g \<^bold>\\<^bold>\\<^bold>\ Dom h) ((Dom f \<^bold>\ Dom g) \<^bold>\ Dom h)" using 1 4 5 6 7 8 E.Nmlize_Hcomp_Hcomp B.vcomp_can by simp finally show ?thesis by simp qed finally show ?thesis by blast qed qed qed lemma DN_is_pseudofunctor: shows "pseudofunctor vcomp hcomp \ \ src trg V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B DN cmp\<^sub>D\<^sub>N" .. interpretation faithful_functor vcomp V\<^sub>B DN using arr_char dom_char cod_char DN_def apply unfold_locales by (metis (no_types, lifting) Cod_dom Dom_cod MkArr_Map) no_notation B.in_hom ("\_ : _ \\<^sub>B _\") lemma DN_UP: assumes "B.arr \" shows "DN (UP \) = \" using assms UP_def DN_def arr_UP by auto interpretation DN: equivalence_pseudofunctor vcomp hcomp \ \ src trg V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B DN cmp\<^sub>D\<^sub>N proof (* DN is locally (but not globally) full. *) show "\f f' \. \ ide f; ide f'; src f = src f'; trg f = trg f'; \\ : DN f \\<^sub>B DN f'\ \ \ \\. \\ : f \ f'\ \ DN \ = \" proof - fix f f' \ assume f: "ide f" and f': "ide f'" and eq_src: "src f = src f'" and eq_trg: "trg f = trg f'" and \: "\\ : DN f \\<^sub>B DN f'\" show "\\. \\ : f \ f'\ \ DN \ = \" proof - let ?\ = "MkArr (Dom f) (Dom f') \" have \: "\?\ : f \ f'\" proof have "Map f = \Dom f\" using f by simp have "Map f' = \Dom f'\" using f' by simp have "Dom f' = Cod f'" using f' Cod_ide by simp show \: "arr ?\" proof - have "E.Nml (Dom ?\) \ E.Ide (Dom ?\)" proof - have "E.Nml (Dom f) \ E.Ide (Dom f)" using f ide_char arr_char by blast thus ?thesis using f by simp qed moreover have "E.Nml (Cod ?\) \ E.Ide (Cod ?\)" proof - have "E.Nml (Dom f') \ E.Ide (Dom f')" using f' ide_char arr_char by blast thus ?thesis using f' by simp qed moreover have "E.Src (Dom ?\) = E.Src (Cod ?\)" using f f' \ arr_char src_def eq_src ideD(1) by auto moreover have "E.Trg (Dom ?\) = E.Trg (Cod ?\)" using f f' \ arr_char trg_def eq_trg ideD(1) by auto moreover have "\Map ?\ : \Dom ?\\ \\<^sub>B \Cod ?\\\" proof - have "\\ : \Dom f\ \\<^sub>B \Dom f'\\" using f f' \ ide_char arr_char DN_def Cod_ide Map_ide by (metis (no_types, lifting) ideD(1)) thus ?thesis by simp qed ultimately show ?thesis using f f' \ ide_char arr_char by blast qed show "dom ?\ = f" using f \ dom_char MkArr_Map MkIde_Dom' by simp show "cod ?\ = f'" proof - have "cod ?\ = MkIde (Dom f')" using \ cod_char by simp also have "... = MkArr (Dom f') (Cod f') (Map f')" using f' by auto also have "... = f'" using f' MkArr_Map by simp finally show ?thesis by simp qed qed moreover have "DN ?\ = \" using \ DN_def by auto ultimately show ?thesis by blast qed qed (* DN is biessentially surjective on objects. *) show "\a'. B.obj a' \ \a. obj a \ B.equivalent_objects (DN.map\<^sub>0 a) a'" proof - fix a' assume a': "B.obj a'" have "obj (UP.map\<^sub>0 a')" using a' UP.map\<^sub>0_simps(1) by simp moreover have "B.equivalent_objects (DN.map\<^sub>0 (UP.map\<^sub>0 a')) a'" proof - have "arr (MkArr \<^bold>\a'\<^bold>\ \<^bold>\a'\<^bold>\ a')" using a' UP_def [of a'] UP.preserves_reflects_arr [of a'] by auto moreover have "arr (MkArr \<^bold>\a'\<^bold>\\<^sub>0 \<^bold>\a'\<^bold>\\<^sub>0 a')" using a' arr_char obj_simps by auto ultimately have "DN.map\<^sub>0 (UP.map\<^sub>0 a') = a'" using a' UP.map\<^sub>0_def DN.map\<^sub>0_def DN_def src_def by auto thus ?thesis using a' B.equivalent_objects_reflexive by simp qed ultimately show "\a. obj a \ B.equivalent_objects (DN.map\<^sub>0 a) a'" by blast qed (* DN is locally essentially surjective. *) show "\a b g. \ obj a; obj b; \g : DN.map\<^sub>0 a \\<^sub>B DN.map\<^sub>0 b\; B.ide g \ \ \f. \f : a \ b\ \ ide f \ B.isomorphic (DN f) g" proof - fix a b g assume a: "obj a" and b: "obj b" and g: "\g : DN.map\<^sub>0 a \\<^sub>B DN.map\<^sub>0 b\" and ide_g: "B.ide g" have "ide (UP g)" using ide_g UP.preserves_ide by simp moreover have "B.isomorphic (DN (UP g)) g" using ide_g DN_UP B.isomorphic_reflexive by simp moreover have "\UP g : a \ b\" proof show "arr (UP g)" using g UP.preserves_reflects_arr by auto show "src (UP g) = a" proof - have "src (UP g) = MkArr \<^bold>\src\<^sub>B g\<^bold>\\<^sub>0 \<^bold>\src\<^sub>B g\<^bold>\\<^sub>0 (src\<^sub>B g)" using ide_g src_def UP_def UP.preserves_reflects_arr [of g] B.ideD(1) by simp also have "... = a" proof - have "src\<^sub>B g = src\<^sub>B (DN.map\<^sub>0 a)" using a g B.in_hhom_def by simp also have "... = Map a" using a Map_preserves_objects DN.map\<^sub>0_def DN_def B.src_src B.obj_simps by auto finally have "src\<^sub>B g = Map a" by simp hence "MkArr \<^bold>\src\<^sub>B g\<^bold>\\<^sub>0 \<^bold>\src\<^sub>B g\<^bold>\\<^sub>0 (src\<^sub>B g) = MkArr \<^bold>\Map a\<^bold>\\<^sub>0 \<^bold>\Map a\<^bold>\\<^sub>0 (Map a)" by simp also have "... = a" using a obj_char [of a] MkIde_Dom [of a] apply (cases "Dom a") apply force by simp_all finally show ?thesis by simp qed finally show ?thesis by simp qed show "trg (UP g) = b" proof - have "trg (UP g) = MkArr \<^bold>\trg\<^sub>B g\<^bold>\\<^sub>0 \<^bold>\trg\<^sub>B g\<^bold>\\<^sub>0 (trg\<^sub>B g)" using ide_g trg_def UP_def UP.preserves_reflects_arr [of g] B.ideD(1) by simp also have "... = b" proof - have "trg\<^sub>B g = trg\<^sub>B (DN.map\<^sub>0 b)" using b g B.in_hhom_def by simp also have "... = Map b" using b Map_preserves_objects DN.map\<^sub>0_def DN_def B.src_src B.obj_simps by auto finally have "trg\<^sub>B g = Map b" by simp hence "MkArr \<^bold>\trg\<^sub>B g\<^bold>\\<^sub>0 \<^bold>\trg\<^sub>B g\<^bold>\\<^sub>0 (trg\<^sub>B g) = MkArr \<^bold>\Map b\<^bold>\\<^sub>0 \<^bold>\Map b\<^bold>\\<^sub>0 (Map b)" by simp also have "... = b" using b obj_char [of b] MkIde_Dom [of b] apply (cases "Dom b") apply force by simp_all finally show ?thesis by simp qed finally show ?thesis by simp qed qed ultimately show "\f. \f : a \ b\ \ ide f \ B.isomorphic (DN f) g" by blast qed qed theorem DN_is_equivalence_pseudofunctor: shows "equivalence_pseudofunctor vcomp hcomp \ \ src trg V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B DN cmp\<^sub>D\<^sub>N" .. text \ The following gives an explicit formula for a component of the unit isomorphism of the pseudofunctor \UP\ from a bicategory to its strictification. It is not currently being used -- I originally proved it in order to establish something that I later proved in a more abstract setting -- but it might be useful at some point. \ interpretation UP: equivalence_pseudofunctor V\<^sub>B H\<^sub>B \\<^sub>B \\<^sub>B src\<^sub>B trg\<^sub>B vcomp hcomp \ \ src trg UP cmp\<^sub>U\<^sub>P using UP_is_equivalence_pseudofunctor by auto lemma UP_unit_char: assumes "B.obj a" shows "UP.unit a = MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a" proof - have " MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a = UP.unit a" proof (intro UP.unit_eqI) show "B.obj a" using assms by simp have 0: "\a : a \\<^sub>B a\" using assms by auto have 1: "arr (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" apply (unfold arr_char, intro conjI) using assms by auto have 2: "arr (MkArr \<^bold>\a\<^bold>\ \<^bold>\a\<^bold>\ a)" apply (unfold arr_char, intro conjI) using assms by auto have 3: "arr (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\\<^sub>0 a)" apply (unfold arr_char, intro conjI) using assms by auto show "\MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a : UP.map\<^sub>0 a \ UP a\" proof show "arr (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" by fact show "dom (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a) = UP.map\<^sub>0 a" using assms 1 2 dom_char UP.map\<^sub>0_def UP_def src_def by auto show "cod (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a) = UP a" using assms 1 2 cod_char UP.map\<^sub>0_def UP_def src_def by auto qed show "iso (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" using assms 1 iso_char by auto show "MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ \ (UP.map\<^sub>0 a) = (UP \\<^sub>B[a] \ cmp\<^sub>U\<^sub>P (a, a)) \ (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" proof - have "MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ \ (UP.map\<^sub>0 a) = MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a" unfolding \_def UP.map\<^sub>0_def UP_def using assms 0 1 2 src_def by auto also have "... = (UP \\<^sub>B[a] \ cmp\<^sub>U\<^sub>P (a, a)) \ (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" proof - have "(UP \\<^sub>B[a] \ cmp\<^sub>U\<^sub>P (a, a)) \ (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a) = (MkArr \<^bold>\a \\<^sub>B a\<^bold>\ \<^bold>\a\<^bold>\ \\<^sub>B[a] \ MkArr (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) \<^bold>\a \\<^sub>B a\<^bold>\ (a \\<^sub>B a)) \ (MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a \ MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a)" using assms UP_def cmp\<^sub>U\<^sub>P_ide_simp by auto also have "... = (MkArr \<^bold>\a \\<^sub>B a\<^bold>\ \<^bold>\a\<^bold>\ \\<^sub>B[a] \ MkArr (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) \<^bold>\a \\<^sub>B a\<^bold>\ (a \\<^sub>B a)) \ MkArr \<^bold>\a\<^bold>\\<^sub>0 (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) (B.runit' a)" using assms 0 1 2 3 hcomp_def B.comp_cod_arr src_def trg_def B.can_Ide_self B.canE_unitor [of "\<^bold>\a\<^bold>\\<^sub>0"] B.comp_cod_arr by auto also have "... = MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ ((\\<^sub>B[a] \\<^sub>B (a \\<^sub>B a)) \\<^sub>B B.runit' a)" proof - have "MkArr \<^bold>\a \\<^sub>B a\<^bold>\ \<^bold>\a\<^bold>\ \\<^sub>B[a] \ MkArr (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) \<^bold>\a \\<^sub>B a\<^bold>\ (a \\<^sub>B a) = MkArr (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) \<^bold>\a\<^bold>\ (\\<^sub>B[a] \\<^sub>B (a \\<^sub>B a))" using assms by (intro comp_MkArr arr_MkArr) auto moreover have "MkArr (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) \<^bold>\a\<^bold>\ (\\<^sub>B[a] \\<^sub>B (a \\<^sub>B a)) \ MkArr \<^bold>\a\<^bold>\\<^sub>0 (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\a\<^bold>\) (B.runit' a) = MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ ((\\<^sub>B[a] \\<^sub>B (a \\<^sub>B a)) \\<^sub>B B.runit' a)" using assms 0 B.comp_arr_dom by (intro comp_MkArr arr_MkArr, auto) ultimately show ?thesis by argo qed also have "... = MkArr \<^bold>\a\<^bold>\\<^sub>0 \<^bold>\a\<^bold>\ a" using assms B.comp_arr_dom B.comp_arr_inv' B.iso_unit B.unitor_coincidence(2) by simp finally show ?thesis by argo qed finally show ?thesis by simp qed qed thus ?thesis by simp qed end subsection "Pseudofunctors into a Strict Bicategory" text \ In the special case of a pseudofunctor into a strict bicategory, we can obtain explicit formulas for the images of the units and associativities under the pseudofunctor, which only involve the structure maps of the pseudofunctor, since the units and associativities in the target bicategory are all identities. This is useful in applying strictification. \ locale pseudofunctor_into_strict_bicategory = pseudofunctor + D: strict_bicategory V\<^sub>D H\<^sub>D \\<^sub>D \\<^sub>D src\<^sub>D trg\<^sub>D begin lemma image_of_unitor: assumes "C.ide g" shows "F \\<^sub>C[g] = (D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" and "F \\<^sub>C[g] = (F g \\<^sub>D D.inv (unit (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g))" and "F (C.lunit' g) = \ (trg\<^sub>C g, g) \\<^sub>D (unit (trg\<^sub>C g) \\<^sub>D F g)" and "F (C.runit' g) = \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D unit (src\<^sub>C g))" proof - show A: "F \\<^sub>C[g] = (D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" proof - have 1: "\(D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g)) : F (trg\<^sub>C g \\<^sub>C g) \\<^sub>D F g\" proof show "\D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g : F (trg\<^sub>C g) \\<^sub>D F g \\<^sub>D F g\" using assms unit_char by (auto simp add: D.hcomp_obj_arr) show "\D.inv (\ (trg\<^sub>C g, g)) : F (trg\<^sub>C g \\<^sub>C g) \\<^sub>D F (trg\<^sub>C g) \\<^sub>D F g\" using assms cmp_in_hom(2) D.inv_is_inverse by simp qed have "(D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g)) = F g \\<^sub>D (D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" using 1 D.comp_cod_arr by auto also have "... = (F \\<^sub>C[g] \\<^sub>D \ (trg\<^sub>C g, g) \\<^sub>D (unit (trg\<^sub>C g) \\<^sub>D F g)) \\<^sub>D (D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" using assms lunit_coherence [of g] D.strict_lunit by simp also have "... = F \\<^sub>C[g] \\<^sub>D \ (trg\<^sub>C g, g) \\<^sub>D ((unit (trg\<^sub>C g) \\<^sub>D F g) \\<^sub>D (D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g)) \\<^sub>D D.inv (\ (trg\<^sub>C g, g))" using D.comp_assoc by simp also have "... = F \\<^sub>C[g]" proof - have "(unit (trg\<^sub>C g) \\<^sub>D F g) \\<^sub>D (D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g) = F (trg\<^sub>C g) \\<^sub>D F g" using assms unit_char D.whisker_right by (metis C.ideD(1) C.obj_trg C.trg.preserves_reflects_arr D.comp_arr_inv' unit_simps(5) preserves_arr preserves_ide) moreover have "\ (trg\<^sub>C g, g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g)) = F (trg\<^sub>C g \\<^sub>C g)" using assms D.comp_arr_inv D.inv_is_inverse by simp ultimately show ?thesis using assms D.comp_arr_dom D.comp_cod_arr unit_char by auto qed finally show ?thesis by simp qed show B: "F \\<^sub>C[g] = (F g \\<^sub>D D.inv (unit (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g))" proof - have 1: "\(F g \\<^sub>D D.inv (unit (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g)) : F (g \\<^sub>C src\<^sub>C g) \\<^sub>D F g\" proof show "\F g \\<^sub>D D.inv (unit (src\<^sub>C g)) : F g \\<^sub>D F (src\<^sub>C g) \\<^sub>D F g\" using assms unit_char by (auto simp add: D.hcomp_arr_obj) show "\D.inv (\ (g, src\<^sub>C g)) : F (g \\<^sub>C src\<^sub>C g) \\<^sub>D F g \\<^sub>D F (src\<^sub>C g)\" using assms cmp_in_hom(2) by simp qed have "(F g \\<^sub>D D.inv (unit (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g)) = F g \\<^sub>D (F g \\<^sub>D D.inv (unit (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g))" using 1 D.comp_cod_arr by auto also have "... = (F \\<^sub>C[g] \\<^sub>D \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D unit (src\<^sub>C g))) \\<^sub>D (F g \\<^sub>D D.inv (unit (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g))" using assms runit_coherence [of g] D.strict_runit by simp also have "... = F \\<^sub>C[g] \\<^sub>D (\ (g, src\<^sub>C g) \\<^sub>D ((F g \\<^sub>D unit (src\<^sub>C g)) \\<^sub>D (F g \\<^sub>D D.inv (unit (src\<^sub>C g))))) \\<^sub>D D.inv (\ (g, src\<^sub>C g))" using D.comp_assoc by simp also have "... = F \\<^sub>C[g]" proof - have "(F g \\<^sub>D unit (src\<^sub>C g)) \\<^sub>D (F g \\<^sub>D D.inv (unit (src\<^sub>C g))) = F g \\<^sub>D F (src\<^sub>C g)" using assms D.whisker_left unit_char by (metis C.ideD(1) C.obj_src C.src.preserves_ide D.comp_arr_inv' D.ideD(1) unit_simps(5) preserves_ide) moreover have "\ (g, src\<^sub>C g) \\<^sub>D D.inv (\ (g, src\<^sub>C g)) = F (g \\<^sub>C src\<^sub>C g)" using assms D.comp_arr_inv D.inv_is_inverse by simp ultimately show ?thesis using assms D.comp_arr_dom D.comp_cod_arr unit_char cmp_in_hom(2) [of g "src\<^sub>C g"] by auto qed finally show ?thesis by simp qed show "F (C.lunit' g) = \ (trg\<^sub>C g, g) \\<^sub>D (unit (trg\<^sub>C g) \\<^sub>D F g)" proof - have "F (C.lunit' g) = D.inv (F \\<^sub>C[g])" using assms C.iso_lunit preserves_inv by simp also have "... = D.inv ((D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g) \\<^sub>D D.inv (\ (trg\<^sub>C g, g)))" using A by simp also have "... = \ (trg\<^sub>C g, g) \\<^sub>D (unit (trg\<^sub>C g) \\<^sub>D F g)" proof - have "D.iso (D.inv (\ (trg\<^sub>C g, g))) \ D.inv (D.inv (\ (trg\<^sub>C g, g))) = \ (trg\<^sub>C g, g)" using assms by simp moreover have "D.iso (D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g) \ D.inv (D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g) = unit (trg\<^sub>C g) \\<^sub>D F g" using assms unit_char by simp moreover have "D.seq (D.inv (unit (trg\<^sub>C g)) \\<^sub>D F g) (D.inv (\ (trg\<^sub>C g, g)))" using assms unit_char by (metis A C.lunit_simps(1) preserves_arr) ultimately show ?thesis using D.inv_comp by simp qed finally show ?thesis by simp qed show "F (C.runit' g) = \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D unit (src\<^sub>C g))" proof - have "F (C.runit' g) = D.inv (F \\<^sub>C[g])" using assms C.iso_runit preserves_inv by simp also have "... = D.inv ((F g \\<^sub>D D.inv (unit (src\<^sub>C g))) \\<^sub>D D.inv (\ (g, src\<^sub>C g)))" using B by simp also have "... = \ (g, src\<^sub>C g) \\<^sub>D (F g \\<^sub>D unit (src\<^sub>C g))" proof - have "D.iso (D.inv (\ (g, src\<^sub>C g))) \ D.inv (D.inv (\ (g, src\<^sub>C g))) = \ (g, src\<^sub>C g)" using assms by simp moreover have "D.iso (F g \\<^sub>D D.inv (unit (src\<^sub>C g))) \ D.inv (F g \\<^sub>D D.inv (unit (src\<^sub>C g))) = F g \\<^sub>D unit (src\<^sub>C g)" using assms unit_char by simp moreover have "D.seq (F g \\<^sub>D D.inv (unit (src\<^sub>C g))) (D.inv (\ (g, src\<^sub>C g)))" using assms unit_char by (metis B C.runit_simps(1) preserves_arr) ultimately show ?thesis using D.inv_comp by simp qed finally show ?thesis by simp qed qed lemma image_of_associator: assumes "C.ide f" and "C.ide g" and "C.ide h" and "src\<^sub>C f = trg\<^sub>C g" and "src\<^sub>C g = trg\<^sub>C h" shows "F \\<^sub>C[f, g, h] = \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" and "F (C.\' f g h) = \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) \\<^sub>D (F f \\<^sub>D D.inv (\ (g, h))) \\<^sub>D D.inv (\ (f, g \\<^sub>C h))" proof - show 1: "F \\<^sub>C[f, g, h] = \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" proof - have 2: "D.seq (\ (f, g \\<^sub>C h)) ((F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h])" using assms D.assoc_in_hom(2) [of "F f" "F g" "F h"] cmp_simps(1,4) C.VV.cod_simp by (intro D.seqI) auto moreover have 3: "F \\<^sub>C[f, g, h] \\<^sub>D \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) = \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h]" using assms assoc_coherence [of f g h] by blast moreover have "D.iso (\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h))" using assms 2 3 C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def D.isos_compose by auto ultimately have "F \\<^sub>C[f, g, h] = (\ (f, g \\<^sub>C h) \\<^sub>D ((F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h])) \\<^sub>D D.inv (\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h))" using D.invert_side_of_triangle(2) [of "\ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h]" "F \\<^sub>C[f, g, h]" "\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h)"] by presburger also have "... = \ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" proof - have "D.inv (\ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h)) = (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h))" proof - have "D.seq (\ (f \\<^sub>C g, h)) (\ (f, g) \\<^sub>D F h)" using assms by fastforce thus ?thesis using assms D.inv_comp by simp qed moreover have "(F f \\<^sub>D \ (g, h)) \\<^sub>D \\<^sub>D[F f, F g, F h] = (F f \\<^sub>D \ (g, h))" using assms D.comp_arr_dom D.assoc_in_hom [of "F f" "F g" "F h"] cmp_in_hom by (metis "2" "3" D.comp_arr_ide D.hseq_char D.seqE D.strict_assoc cmp_simps(2) cmp_simps(3) preserves_ide) ultimately show ?thesis using D.comp_assoc by simp qed finally show ?thesis by simp qed show "F (C.\' f g h) = \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) \\<^sub>D (F f \\<^sub>D D.inv (\ (g, h))) \\<^sub>D D.inv (\ (f, g \\<^sub>C h))" proof - have "F (C.\' f g h) = D.inv (F \\<^sub>C[f, g, h])" using assms preserves_inv by simp also have "... = D.inv (\ (f, g \\<^sub>C h) \\<^sub>D (F f \\<^sub>D \ (g, h)) \\<^sub>D (D.inv (\ (f, g)) \\<^sub>D F h) \\<^sub>D D.inv (\ (f \\<^sub>C g, h)))" using 1 by simp also have "... = \ (f \\<^sub>C g, h) \\<^sub>D (\ (f, g) \\<^sub>D F h) \\<^sub>D (F f \\<^sub>D D.inv (\ (g, h))) \\<^sub>D D.inv (\ (f, g \\<^sub>C h))" using assms C.VV.arr_char FF_def D.hcomp_assoc D.comp_assoc C.VV.dom_simp C.VV.cod_simp (* OK, this is pretty cool, but not as cool as if I didn't have to direct it. *) by (simp add: D.inv_comp D.isos_compose) finally show ?thesis by simp qed qed end subsection "Internal Equivalences in a Strict Bicategory" text \ In this section we prove a useful fact about internal equivalences in a strict bicategory: namely, that if the ``right'' triangle identity holds for such an equivalence then the ``left'' does, as well. Later we will dualize this property, and use strictification to extend it to arbitrary bicategories. \ locale equivalence_in_strict_bicategory = strict_bicategory + equivalence_in_bicategory begin lemma triangle_right_implies_left: shows "(g \ \) \ (\ \ g) = g \ (\ \ f) \ (f \ \) = f" proof - text \ The formal proof here was constructed following the string diagram sketch below, which appears in \cite{nlab-zigzag-diagram} (see it also in context in \cite{nlab-adjoint-equivalence}). The diagram is reproduced here by permission of its author, Mike Shulman, who says (private communication): ``Just don't give the impression that the proof itself is due to me, because it's not. I don't know who first gave that proof; it's probably folklore.'' \begin{figure}[h] \includegraphics[width=6.5in]{triangle_right_implies_left.png} \end{figure} \ assume 1: "(g \ \) \ (\ \ g) = g" have 2: "(inv \ \ g) \ (g \ inv \) = g" proof - have "(inv \ \ g) \ (g \ inv \) = inv ((g \ \) \ (\ \ g))" using antipar inv_comp hcomp_assoc by simp also have "... = g" using 1 by simp finally show ?thesis by blast qed have "(\ \ f) \ (f \ \) = (\ \ f) \ (f \ (inv \ \ g) \ (g \ inv \) \ f) \ (f \ \)" proof - have "(f \ (inv \ \ g) \ (g \ inv \) \ f) \ (f \ \) = f \ \" using 2 ide_left ide_right antipar whisker_left by (metis comp_cod_arr unit_simps(1) unit_simps(3)) thus ?thesis by simp qed also have "... = (\ \ f) \ (f \ (inv \ \ g) \ (g \ inv \) \ f) \ (f \ \) \ (inv \ \ \)" proof - have "inv \ \ \ = src f" by (simp add: comp_inv_arr') thus ?thesis by (metis antipar(1) antipar(2) arrI calculation comp_ide_arr hcomp_arr_obj ideD(1) ide_left ide_right obj_src seqE strict_assoc' triangle_in_hom(1) vconn_implies_hpar(1)) qed also have "... = (\ \ (f \ (inv \ \ g) \ (g \ inv \)) \ ((f \ inv \) \ (f \ \))) \ (f \ \)" using ide_left ide_right antipar unit_is_iso by (metis "2" arr_inv calculation comp_arr_dom comp_inv_arr' counit_simps(1) counit_simps(2) dom_inv hcomp_arr_obj ideD(1) unit_simps(1) unit_simps(2) unit_simps(5) obj_trg seqI whisker_left) also have "... = (\ \ (f \ (inv \ \ g) \ (g \ inv \)) \ ((f \ inv \) \ ((inv \ \ f) \ (\ \ f)) \ (f \ \))) \ (f \ \)" proof - have "(inv \ \ f) \ (\ \ f) = (f \ g) \ f" using whisker_right [of f "inv \" \] counit_in_hom by (simp add: antipar(1) comp_inv_arr') thus ?thesis using hcomp_assoc comp_arr_dom by (metis comp_cod_arr ide_left local.unit_simps(1) local.unit_simps(3) whisker_left) qed also have "... = (((\ \ (f \ (inv \ \ g) \ (g \ inv \))) \ (f \ g)) \ (((f \ inv \) \ (inv \ \ f)) \ (\ \ f) \ (f \ \))) \ (f \ \)" using ide_left ide_right antipar comp_assoc whisker_right comp_cod_arr by (metis "2" comp_arr_dom counit_simps(1-2)) also have "... = (((\ \ (f \ (inv \ \ g) \ (g \ inv \))) \ ((f \ inv \) \ (inv \ \ f))) \ ((f \ g) \ (\ \ f) \ (f \ \))) \ (f \ \)" proof - have 3: "seq (\ \ (f \ (inv \ \ g) \ (g \ inv \))) (f \ g)" using 2 antipar by auto moreover have 4: "seq ((f \ inv \) \ (inv \ \ f)) ((\ \ f) \ (f \ \))" using antipar hcomp_assoc by auto ultimately show ?thesis using interchange by simp qed also have "... = ((\ \ (f \ (inv \ \ g) \ (g \ inv \))) \ ((f \ inv \) \ (inv \ \ f))) \ ((f \ g) \ (\ \ f) \ (f \ \)) \ (f \ \)" using comp_assoc by presburger also have "... = ((\ \ (f \ inv \) \ (inv \ \ f)) \ ((f \ (inv \ \ g) \ (g \ inv \)) \ f)) \ (f \ (g \ \) \ (\ \ g) \ f) \ (f \ \)" proof - have "(\ \ (f \ (inv \ \ g) \ (g \ inv \))) \ ((f \ inv \) \ (inv \ \ f)) = (\ \ (f \ inv \) \ (inv \ \ f)) \ ((f \ (inv \ \ g) \ (g \ inv \)) \ f)" proof - have "seq \ (f \ (inv \ \ g) \ (g \ inv \))" using 2 antipar by simp moreover have "seq ((f \ inv \) \ (inv \ \ f)) f" using antipar hcomp_assoc hcomp_obj_arr by auto ultimately show ?thesis using comp_assoc comp_arr_dom hcomp_obj_arr interchange [of \ "f \ (inv \ \ g) \ (g \ inv \)" "(f \ inv \) \ (inv \ \ f)" f] by simp qed moreover have "((f \ g) \ (\ \ f) \ (f \ \)) \ (f \ \) = (f \ (g \ \) \ (\ \ g) \ f) \ (f \ \)" proof - have "((f \ g) \ (\ \ f) \ (f \ \)) \ (f \ \) = (f \ g \ \ \ f) \ (f \ (g \ f) \ \) \ (f \ \ \ src f)" using antipar comp_assoc hcomp_assoc whisker_left hcomp_arr_obj by simp also have "... = (f \ g \ \ \ f) \ (f \ ((g \ f) \ \) \ (\ \ src f))" using antipar comp_arr_dom whisker_left hcomp_arr_obj by simp also have "... = (f \ g \ \ \ f) \ (f \ \ \ \)" using antipar comp_arr_dom comp_cod_arr hcomp_arr_obj interchange [of "g \ f" \ \ "src f"] by simp also have "... = (f \ g \ \ \ f) \ (f \ \ \ g \ f) \ (f \ src f \ \)" using antipar comp_arr_dom comp_cod_arr whisker_left interchange [of \ "src f" "g \ f" \] by simp also have "... = ((f \ g \ \ \ f) \ (f \ \ \ g \ f)) \ (f \ \)" using antipar comp_assoc by (simp add: hcomp_obj_arr) also have "... = (f \ (g \ \) \ (\ \ g) \ f) \ (f \ \)" using antipar comp_assoc whisker_left whisker_right hcomp_assoc by simp finally show ?thesis by blast qed ultimately show ?thesis by simp qed also have "... = (\ \ (f \ inv \) \ (inv \ \ f)) \ ((f \ (inv \ \ g) \ (g \ inv \) \ f) \ (f \ (g \ \) \ (\ \ g) \ f)) \ (f \ \)" using comp_assoc hcomp_assoc antipar(1) antipar(2) by auto also have "... = (\ \ (f \ inv \) \ (inv \ \ f)) \ ((f \ (inv \ \ g) \ (g \ inv \) \ (g \ \) \ (\ \ g) \ f)) \ (f \ \)" using ide_left ide_right antipar comp_cod_arr comp_assoc whisker_left by (metis "1" "2" comp_ide_self unit_simps(1) unit_simps(3)) also have "... = (\ \ (f \ inv \) \ (inv \ \ f)) \ (f \ \)" proof - have "(inv \ \ g) \ (g \ inv \) \ (g \ \) \ (\ \ g) = g" using ide_left ide_right antipar comp_arr_dom comp_assoc by (metis "1" "2" comp_ide_self) thus ?thesis using antipar comp_cod_arr by simp qed also have "... = (f \ inv \) \ ((inv \ \ f) \ (\ \ f)) \ (f \ \)" proof - have "(\ \ (f \ inv \) \ (inv \ \ f)) \ (f \ \) = (trg f \ \ \ (f \ inv \) \ (inv \ \ f)) \ (f \ \)" using hcomp_obj_arr comp_cod_arr by simp also have "... = ((trg f \ f \ inv \) \ (\ \ inv \ \ f)) \ (f \ \)" using antipar hcomp_arr_obj hcomp_assoc interchange by auto also have "... = (f \ inv \) \ ((inv \ \ f) \ (\ \ f)) \ (f \ \)" proof - have "(inv \ \ f) \ (\ \ f) = (trg f \ inv \ \ f) \ (\ \ trg f \ f)" using hseqI' by (simp add: hcomp_obj_arr) also have "... = \ \ inv \ \ f" using antipar comp_arr_dom comp_cod_arr interchange [of "trg f" \ "inv \ \ f" "trg f \ f"] by force finally have "(inv \ \ f) \ (\ \ f) = \ \ inv \ \ f" by simp moreover have "trg f \ f \ inv \ = f \ inv \" using hcomp_obj_arr [of "trg f" "f \ inv \"] by fastforce ultimately have "((trg f \ f \ inv \) \ (\ \ inv \ \ f)) \ (f \ \) = ((f \ inv \) \ ((inv \ \ f) \ (\ \ f))) \ (f \ \)" by simp thus ?thesis using comp_assoc by simp qed finally show ?thesis by simp qed also have "... = f \ inv \ \ \" proof - have "(inv \ \ f) \ (\ \ f) = f \ g \ f" using ide_left ide_right antipar counit_is_iso whisker_right hcomp_assoc by (metis comp_arr_dom comp_inv_arr' counit_simps(1) counit_simps(2) seqE) thus ?thesis using ide_left ide_right antipar unit_is_iso comp_cod_arr by (metis arr_inv dom_inv unit_simps(1) unit_simps(3) seqI whisker_left) qed also have "... = f \ src f" using antipar by (simp add: comp_inv_arr') also have "... = f" using hcomp_arr_obj by simp finally show ?thesis by simp qed end text \ Now we use strictification to generalize the preceding result to arbitrary bicategories. I originally attempted to generalize this proof directly from the strict case, by filling in the necessary canonical isomorphisms, but it turned out to be too daunting. The proof using strictification is still fairly tedious, but it is manageable. \ context equivalence_in_bicategory begin interpretation S: strictified_bicategory V H \ \ src trg .. notation S.vcomp (infixr "\\<^sub>S" 55) notation S.hcomp (infixr "\\<^sub>S" 53) notation S.in_hom ("\_ : _ \\<^sub>S _\") notation S.in_hhom ("\_ : _ \\<^sub>S _\") interpretation UP: equivalence_pseudofunctor V H \ \ src trg S.vcomp S.hcomp S.\ S.\ S.src S.trg S.UP S.cmp\<^sub>U\<^sub>P using S.UP_is_equivalence_pseudofunctor by auto interpretation UP: pseudofunctor_into_strict_bicategory V H \ \ src trg S.vcomp S.hcomp S.\ S.\ S.src S.trg S.UP S.cmp\<^sub>U\<^sub>P .. interpretation E: equivalence_in_bicategory S.vcomp S.hcomp S.\ S.\ S.src S.trg \S.UP f\ \S.UP g\ \S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.unit (src f)\ \S.inv (UP.unit (trg f)) \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)\ using equivalence_in_bicategory_axioms UP.preserves_equivalence [of f g \ \] by auto interpretation E: equivalence_in_strict_bicategory S.vcomp S.hcomp S.\ S.\ S.src S.trg \S.UP f\ \S.UP g\ \S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.unit (src f)\ \S.inv (UP.unit (trg f)) \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)\ .. lemma UP_triangle: shows "S.UP ((g \ \) \ \[g, f, g] \ (\ \ g)) = S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))" and "S.UP (\\<^sup>-\<^sup>1[g] \ \[g]) = (S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S (S.UP g \\<^sub>S UP.unit (src g))) \\<^sub>S (S.inv (UP.unit (trg g)) \\<^sub>S S.UP g) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))" and "S.UP ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) = S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" and "S.UP (\\<^sup>-\<^sup>1[f] \ \[f]) = (S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (UP.unit (trg f) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.unit (src f))) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" and "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g] \ S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f)) = (S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (UP.unit (trg f) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.unit (src f))) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" proof - show T1: "S.UP ((g \ \) \ \[g, f, g] \ (\ \ g)) = S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))" proof - have "S.UP ((g \ \) \ \[g, f, g] \ (\ \ g)) = S.UP (g \ \) \\<^sub>S S.UP \[g, f, g] \\<^sub>S S.UP (\ \ g)" using antipar by simp also have "... = (S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \) \\<^sub>S ((S.inv (S.cmp\<^sub>U\<^sub>P (g, f \ g)) \\<^sub>S S.cmp\<^sub>U\<^sub>P (g, f \ g)) \\<^sub>S (S.UP g \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))) \\<^sub>S (((S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP g) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g \ f, g)))) \\<^sub>S S.cmp\<^sub>U\<^sub>P (g \ f, g)) \\<^sub>S (S.UP \ \\<^sub>S S.UP g)) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))" proof - have "S.UP \[g, f, g] = S.cmp\<^sub>U\<^sub>P (g, f \ g) \\<^sub>S (S.UP g \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP g) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g \ f, g))" using ide_left ide_right antipar UP.image_of_associator [of g f g] by simp moreover have "S.UP (g \ \) = S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f \ g))" proof - have "S.seq (S.cmp\<^sub>U\<^sub>P (g, src g)) (S.UP g \\<^sub>S S.UP \)" using antipar UP.FF_def UP.cmp_in_hom [of g "src g"] apply (intro S.seqI) by auto moreover have "S.UP (g \ \) \\<^sub>S S.cmp\<^sub>U\<^sub>P (g, f \ g) = S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \)" using antipar UP.\.naturality [of "(g, \)"] UP.FF_def VV.arr_char VV.dom_simp VV.cod_simp by simp moreover have "S.iso (S.cmp\<^sub>U\<^sub>P (g, f \ g))" using antipar by simp ultimately show ?thesis - using antipar S.comp_assoc - S.invert_side_of_triangle(2) + using S.invert_side_of_triangle(2) [of "S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \)" "S.UP (g \ \)" - "S.cmp\<^sub>U\<^sub>P (g, f \ g)"] - by simp + "S.cmp\<^sub>U\<^sub>P (g, f \ g)"] S.comp_assoc + by presburger qed moreover have "S.UP (\ \ g) = (S.cmp\<^sub>U\<^sub>P (g \ f, g) \\<^sub>S (S.UP \ \\<^sub>S S.UP g)) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))" proof - have "S.UP (\ \ g) \\<^sub>S S.cmp\<^sub>U\<^sub>P (trg g, g) = S.cmp\<^sub>U\<^sub>P (g \ f, g) \\<^sub>S (S.UP \ \\<^sub>S S.UP g)" using antipar UP.\.naturality [of "(\, g)"] UP.FF_def VV.arr_char VV.dom_simp VV.cod_simp by simp moreover have "S.seq (S.cmp\<^sub>U\<^sub>P (g \ f, g)) (S.UP \ \\<^sub>S S.UP g)" using antipar UP.cmp_in_hom(2) by (intro S.seqI, auto) ultimately show ?thesis using antipar S.invert_side_of_triangle(2) by simp qed ultimately show ?thesis using S.comp_assoc by simp qed also have "... = S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S ((S.UP g \\<^sub>S S.UP \) \\<^sub>S (S.UP g \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))) \\<^sub>S ((S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP g) \\<^sub>S (S.UP \ \\<^sub>S S.UP g)) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))" proof - have "(S.inv (S.cmp\<^sub>U\<^sub>P (g \ f, g)) \\<^sub>S S.cmp\<^sub>U\<^sub>P (g \ f, g)) \\<^sub>S (S.UP \ \\<^sub>S S.UP g) = (S.UP \ \\<^sub>S S.UP g)" using antipar S.comp_inv_arr' S.comp_cod_arr by auto moreover have "(S.inv (S.cmp\<^sub>U\<^sub>P (g, f \ g)) \\<^sub>S S.cmp\<^sub>U\<^sub>P (g, f \ g)) \\<^sub>S (S.UP g \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) = (S.UP g \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))" proof - have "S.inv (S.cmp\<^sub>U\<^sub>P (g, f \ g)) \\<^sub>S S.cmp\<^sub>U\<^sub>P (g, f \ g) = S.UP g \\<^sub>S S.UP (f \ g)" using antipar S.comp_inv_arr' UP.cmp_in_hom by auto thus ?thesis using antipar VV.arr_char S.comp_cod_arr by simp qed ultimately show ?thesis using S.comp_assoc by simp qed also have "... = S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))" using antipar VV.arr_char S.whisker_left S.whisker_right by auto finally show ?thesis by simp qed show T2: "S.UP (\\<^sup>-\<^sup>1[g] \ \[g]) = (S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S (S.UP g \\<^sub>S UP.unit (src g))) \\<^sub>S (S.inv (UP.unit (trg g)) \\<^sub>S S.UP g) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))" using UP.image_of_unitor by simp show "S.UP ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) = S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" proof - have "S.UP ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) = S.UP (\ \ f) \\<^sub>S S.UP \\<^sup>-\<^sup>1[f, g, f] \\<^sub>S S.UP (f \ \)" using antipar by simp also have "... = S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.UP f) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (f \ g, f)) \\<^sub>S S.cmp\<^sub>U\<^sub>P (f \ g, f) \\<^sub>S (S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (f, g \ f)) \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g \ f) \\<^sub>S (S.UP f \\<^sub>S S.UP \)) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" proof - have "S.UP \\<^sup>-\<^sup>1[f, g, f] = S.cmp\<^sub>U\<^sub>P (f \ g, f) \\<^sub>S (S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, g \ f))" using ide_left ide_right antipar UP.image_of_associator by simp moreover have "S.UP (\ \ f) = S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.UP f) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f \ g, f))" proof - have "S.seq (S.cmp\<^sub>U\<^sub>P (trg f, f)) (S.UP \ \\<^sub>S S.UP f)" using antipar UP.FF_def VV.ide_char VV.arr_char UP.cmp_in_hom [of "trg f" f] apply (intro S.seqI) by auto moreover have "S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.UP f) = S.UP (\ \ f) \\<^sub>S S.cmp\<^sub>U\<^sub>P (f \ g, f)" using antipar UP.\.naturality [of "(\, f)"] UP.FF_def VV.arr_char VV.dom_simp VV.cod_simp by simp moreover have "S.iso (S.cmp\<^sub>U\<^sub>P (f \ g, f))" using antipar by simp ultimately show ?thesis using antipar S.comp_assoc S.invert_side_of_triangle(2) [of "S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.UP f)" "S.UP (\ \ f)" "S.cmp\<^sub>U\<^sub>P (f \ g, f)"] by metis qed moreover have "S.UP (f \ \) = (S.cmp\<^sub>U\<^sub>P (f, g \ f) \\<^sub>S (S.UP f \\<^sub>S S.UP \)) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" proof - have "S.cmp\<^sub>U\<^sub>P (f, g \ f) \\<^sub>S (S.UP f \\<^sub>S S.UP \) = S.UP (f \ \) \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, src f)" using antipar UP.\.naturality [of "(f, \)"] UP.FF_def VV.arr_char VV.dom_simp VV.cod_simp by simp moreover have "S.seq (S.cmp\<^sub>U\<^sub>P (f, g \ f)) (S.UP f \\<^sub>S S.UP \)" using antipar by (intro S.seqI, auto) ultimately show ?thesis using antipar S.invert_side_of_triangle(2) by auto qed ultimately show ?thesis using S.comp_assoc by simp qed also have "... = S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S ((S.UP \ \\<^sub>S S.UP f) \\<^sub>S (S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f)) \\<^sub>S ((S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \\<^sub>S (S.UP f \\<^sub>S S.UP \)) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" proof - have "(S.inv (S.cmp\<^sub>U\<^sub>P (f \ g, f)) \\<^sub>S S.cmp\<^sub>U\<^sub>P (f \ g, f)) \\<^sub>S (S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) = (S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f)" using antipar S.comp_cod_arr VV.arr_char S.comp_inv_arr' by auto moreover have "(S.inv (S.cmp\<^sub>U\<^sub>P (f, g \ f)) \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g \ f)) \\<^sub>S (S.UP f \\<^sub>S S.UP \) = (S.UP f \\<^sub>S S.UP \)" using antipar S.comp_inv_arr' S.comp_cod_arr by auto ultimately show ?thesis using S.comp_assoc by simp qed also have "... = S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" using antipar VV.arr_char S.whisker_left S.whisker_right by auto finally show ?thesis by simp qed show "S.UP (\\<^sup>-\<^sup>1[f] \ \[f]) = (S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (UP.unit (trg f) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.unit (src f))) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" using UP.image_of_unitor by simp show "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g] \ S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f)) = (S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (UP.unit (trg f) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.unit (src f))) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" proof - assume A: "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" have B: "(S.UP g \\<^sub>S S.inv (UP.unit (src g)) \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.unit (trg g) \\<^sub>S S.UP g) = S.UP g" proof - show "(S.UP g \\<^sub>S S.inv (UP.unit (src g)) \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.unit (trg g) \\<^sub>S S.UP g) = S.UP g" proof - have 2: "S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g)) = (S.cmp\<^sub>U\<^sub>P (g, src g) \\<^sub>S (S.UP g \\<^sub>S UP.unit (src g))) \\<^sub>S (S.inv (UP.unit (trg g)) \\<^sub>S S.UP g) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))" using A T1 T2 by simp show ?thesis proof - have 8: "S.seq (S.cmp\<^sub>U\<^sub>P (g, src g)) ((S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g)))" using antipar VV.arr_char S.hcomp_assoc by (metis (no_types, lifting) S.arr_UP T1 arrI triangle_in_hom(2)) have 7: "S.seq (S.cmp\<^sub>U\<^sub>P (g, src g)) ((S.UP g \\<^sub>S UP.unit (src g)) \\<^sub>S (S.inv (UP.unit (trg g)) \\<^sub>S S.UP g) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g)))" using antipar 2 8 S.comp_assoc by auto have 5: "(S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) = (S.UP g \\<^sub>S UP.unit (src g)) \\<^sub>S (S.inv (UP.unit (trg g)) \\<^sub>S S.UP g)" proof - have "((S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g)) = ((S.UP g \\<^sub>S UP.unit (src g)) \\<^sub>S (S.inv (UP.unit (trg g)) \\<^sub>S S.UP g)) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g))" proof - have "S.mono (S.cmp\<^sub>U\<^sub>P (g, src g))" using antipar S.iso_is_section S.section_is_mono by simp thus ?thesis using 2 8 7 S.monoE S.comp_assoc by presburger qed moreover have "S.epi (S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g)))" using antipar S.iso_is_retraction S.retraction_is_epi by simp moreover have "S.seq ((S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)) (S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g)))" using S.comp_assoc S.seq_char 8 by presburger moreover have "S.seq ((S.UP g \\<^sub>S UP.unit (src g)) \\<^sub>S (S.inv (UP.unit (trg g)) \\<^sub>S S.UP g)) (S.inv (S.cmp\<^sub>U\<^sub>P (trg g, g)))" using antipar calculation(1,3) by presburger ultimately show ?thesis using 2 S.epiE by blast qed have 6: "S.seq (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)" using antipar VV.arr_char S.hcomp_assoc by auto have 3: "(S.UP g \\<^sub>S S.inv (UP.unit (src g))) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g) = (S.inv (UP.unit (trg g)) \\<^sub>S S.UP g)" proof - have "S.seq (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)" using 6 by simp moreover have "(S.UP g \\<^sub>S UP.unit (src g)) \\<^sub>S (S.inv (UP.unit (trg g)) \\<^sub>S S.UP g) = (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)" using 5 by argo moreover have "S.iso (S.UP g \\<^sub>S UP.unit (src g))" using antipar UP.unit_char S.UP_map\<^sub>0_obj by simp moreover have "S.inv (S.UP g \\<^sub>S UP.unit (src g)) = S.UP g \\<^sub>S S.inv (UP.unit (src g))" using antipar UP.unit_char S.UP_map\<^sub>0_obj by simp ultimately show ?thesis using S.invert_side_of_triangle(1) [of "(S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g)) \\<^sub>S (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)" "S.UP g \\<^sub>S UP.unit (src g)" "S.inv (UP.unit (trg g)) \\<^sub>S S.UP g"] by presburger qed have 4: "((S.UP g \\<^sub>S S.inv (UP.unit (src g))) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))) \\<^sub>S ((S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)) \\<^sub>S (UP.unit (trg g) \\<^sub>S S.UP g) = S.UP g" proof - have "(((S.UP g \\<^sub>S S.inv (UP.unit (src g))) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))) \\<^sub>S ((S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g)) \\<^sub>S (UP.unit (trg g) \\<^sub>S S.UP g)) = (((S.UP g \\<^sub>S S.inv (UP.unit (src g))) \\<^sub>S (S.UP g \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g))) \\<^sub>S ((S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S S.UP g))) \\<^sub>S (UP.unit (trg g) \\<^sub>S S.UP g)" using S.comp_assoc by simp also have "... = (S.inv (UP.unit (trg g)) \\<^sub>S S.UP g) \\<^sub>S (UP.unit (trg g) \\<^sub>S S.UP g)" using 3 S.comp_assoc by auto also have "... = S.inv (UP.unit (trg g)) \\<^sub>S UP.unit (trg g) \\<^sub>S S.UP g" using UP.unit_char(2) S.whisker_right by auto also have "... = UP.map\<^sub>0 (trg g) \\<^sub>S S.UP g" using UP.unit_char [of "trg g"] S.comp_inv_arr S.inv_is_inverse by simp also have "... = S.UP g" using S.UP_map\<^sub>0_obj by (simp add: S.hcomp_obj_arr) finally show ?thesis by blast qed thus ?thesis using antipar S.whisker_left S.whisker_right UP.unit_char S.comp_assoc by simp qed qed qed show "S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f)) = (S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S (UP.unit (trg f) \\<^sub>S S.UP f)) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.unit (src f))) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" proof - have "(S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \) = (UP.unit (trg f) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.unit (src f)))" proof - have 2: "(S.inv (UP.unit (trg f)) \\<^sub>S S.UP f) \\<^sub>S ((S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \)) \\<^sub>S (S.UP f \\<^sub>S UP.unit (src f)) = S.UP f" proof - have "S.UP f = (S.inv (UP.unit (trg f)) \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.unit (src f))" using B antipar E.triangle_right_implies_left by argo also have "... = (S.inv (UP.unit (trg f)) \\<^sub>S S.UP f) \\<^sub>S ((S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \)) \\<^sub>S (S.UP f \\<^sub>S UP.unit (src f))" proof - have "S.inv (UP.unit (trg f)) \\<^sub>S S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f = (S.inv (UP.unit (trg f)) \\<^sub>S S.UP f) \\<^sub>S (S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f)" using UP.unit_char S.whisker_right by simp moreover have "S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \ \\<^sub>S UP.unit (src f) = (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \) \\<^sub>S (S.UP f \\<^sub>S UP.unit (src f))" using antipar UP.unit_char S.whisker_left S.comp_assoc by simp ultimately show ?thesis using S.comp_assoc by presburger qed finally show ?thesis by argo qed show ?thesis proof - have "((S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \)) \\<^sub>S (S.UP f \\<^sub>S UP.unit (src f)) = (UP.unit (trg f) \\<^sub>S S.UP f)" proof - have "S.inv (S.inv (UP.unit (trg f)) \\<^sub>S S.UP f) \\<^sub>S S.UP f = UP.unit (trg f) \\<^sub>S S.UP f" using UP.unit_char S.comp_arr_dom S.UP_map\<^sub>0_obj by (simp add: S.hcomp_obj_arr) moreover have "S.arr (S.UP f)" by simp moreover have "S.iso (S.inv (UP.unit (trg f)) \\<^sub>S S.UP f)" using S.UP_map\<^sub>0_obj by (simp add: UP.unit_char(2)) ultimately show ?thesis using 2 S.invert_side_of_triangle(1) [of "S.UP f" "S.inv (UP.unit (trg f)) \\<^sub>S S.UP f" "((S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \)) \\<^sub>S (S.UP f \\<^sub>S UP.unit (src f))"] by presburger qed moreover have "S.hseq (UP.unit (trg f)) (S.UP f) \ S.iso (S.UP f \\<^sub>S UP.unit (src f)) \ S.inv (S.UP f \\<^sub>S UP.unit (src f)) = S.UP f \\<^sub>S S.inv (UP.unit (src f))" using UP.unit_char S.UP_map\<^sub>0_obj by simp ultimately show ?thesis using S.invert_side_of_triangle(2) [of "UP.unit (trg f) \\<^sub>S S.UP f" "(S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \)" "S.UP f \\<^sub>S UP.unit (src f)"] by presburger qed qed hence "S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S ((S.UP \ \\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \\<^sub>S S.UP \)) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f)) = S.cmp\<^sub>U\<^sub>P (trg f, f) \\<^sub>S ((UP.unit (trg f) \\<^sub>S S.UP f) \\<^sub>S (S.UP f \\<^sub>S S.inv (UP.unit (src f)))) \\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))" by simp thus ?thesis using S.comp_assoc by simp qed qed qed lemma triangle_right_implies_left: assumes "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" shows "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" proof - have "par ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) (\\<^sup>-\<^sup>1[f] \ \[f])" using antipar by simp moreover have "S.UP ((\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \)) = S.UP (\\<^sup>-\<^sup>1[f] \ \[f])" using assms UP_triangle(3-5) by simp ultimately show "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" using UP.is_faithful by blast qed text \ We \emph{really} don't want to go through the ordeal of proving a dual version of \UP_triangle(5)\, do we? So let's be smart and dualize via the opposite bicategory. \ lemma triangle_left_implies_right: assumes "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" shows "(g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" proof - interpret Cop: op_bicategory V H \ \ src trg .. interpret Eop: equivalence_in_bicategory V Cop.H Cop.\ \ Cop.src Cop.trg g f \ \ using antipar unit_in_hom counit_in_hom by (unfold_locales) simp_all have "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f] \ (g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" using antipar Cop.lunit_ide_simp Cop.runit_ide_simp Cop.assoc_ide_simp VVV.ide_char VVV.arr_char VV.arr_char Eop.triangle_right_implies_left by simp thus ?thesis using assms by simp qed lemma triangle_left_iff_right: shows "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f] \ (g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" using triangle_left_implies_right triangle_right_implies_left by auto end text \ We might as well specialize the dual result back to the strict case while we are at it. \ context equivalence_in_strict_bicategory begin lemma triangle_left_iff_right: shows "(\ \ f) \ (f \ \) = f \ (g \ \) \ (\ \ g) = g" proof - have "(\ \ f) \ (f \ \) = f \ (\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = \\<^sup>-\<^sup>1[f] \ \[f]" proof - have "\\<^sup>-\<^sup>1[f] \ \[f] = f" using strict_lunit strict_runit by simp moreover have "(\ \ f) \ \\<^sup>-\<^sup>1[f, g, f] \ (f \ \) = (\ \ f) \ (f \ \)" using antipar strict_assoc assoc'_in_hom(2) [of f g f] comp_cod_arr by auto ultimately show ?thesis by simp qed also have "... \ (g \ \) \ \[g, f, g] \ (\ \ g) = \\<^sup>-\<^sup>1[g] \ \[g]" using triangle_left_iff_right by blast also have "... \ (g \ \) \ (\ \ g) = g" proof - have "\\<^sup>-\<^sup>1[g] \ \[g] = g" using strict_lunit strict_runit by simp moreover have "(g \ \) \ \[g, f, g] \ (\ \ g) = (g \ \) \ (\ \ g)" using antipar strict_assoc assoc_in_hom(2) [of g f g] comp_cod_arr by auto ultimately show ?thesis by simp qed finally show ?thesis by simp qed end end diff --git a/thys/Bicategory/Subbicategory.thy b/thys/Bicategory/Subbicategory.thy --- a/thys/Bicategory/Subbicategory.thy +++ b/thys/Bicategory/Subbicategory.thy @@ -1,1494 +1,1492 @@ (* Title: Subbicategory Author: Eugene W. Stark , 2019 Maintainer: Eugene W. Stark *) section "Sub-Bicategories" text \ In this section we give a construction of a sub-bicategory in terms of a predicate on the arrows of an ambient bicategory that has certain closure properties with respect to that bicategory. While the construction given here is likely to be of general use, it is not the most general sub-bicategory construction that one could imagine, because it requires that the sub-bicategory actually contain the unit and associativity isomorphisms of the ambient bicategory. Our main motivation for including this construction here is to apply it to exploit the fact that the sub-bicategory of endo-arrows of a fixed object is a monoidal category, which will enable us to transfer to bicategories a result about unit isomorphisms in monoidal categories. \ theory Subbicategory imports Bicategory begin subsection "Construction" locale subbicategory = B: bicategory V H \\<^sub>B \ src\<^sub>B trg\<^sub>B + subcategory V Arr for V :: "'a comp" (infixr "\\<^sub>B" 55) and H :: "'a comp" (infixr "\\<^sub>B" 55) and \\<^sub>B :: "'a \ 'a \ 'a \ 'a" ("\\<^sub>B[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src\<^sub>B :: "'a \ 'a" and trg\<^sub>B :: "'a \ 'a" and Arr :: "'a \ bool" + assumes src_closed: "Arr f \ Arr (src\<^sub>B f)" and trg_closed: "Arr f \ Arr (trg\<^sub>B f)" and hcomp_closed: "\ Arr f; Arr g; trg\<^sub>B f = src\<^sub>B g \ \ Arr (g \\<^sub>B f)" and assoc_closed: "\ Arr f \ B.ide f; Arr g \ B.ide g; Arr h \ B.ide h; src\<^sub>B f = trg\<^sub>B g; src\<^sub>B g = trg\<^sub>B h \ \ Arr (\\<^sub>B f g h)" and assoc'_closed: "\ Arr f \ B.ide f; Arr g \ B.ide g; Arr h \ B.ide h; src\<^sub>B f = trg\<^sub>B g; src\<^sub>B g = trg\<^sub>B h \ \ Arr (B.inv (\\<^sub>B f g h))" and lunit_closed: "\ Arr f; B.ide f \ \ Arr (B.\ f)" and lunit'_closed: "\ Arr f; B.ide f \ \ Arr (B.inv (B.\ f))" and runit_closed: "\ Arr f; B.ide f \ \ Arr (B.\ f)" and runit'_closed: "\ Arr f; B.ide f \ \ Arr (B.inv (B.\ f))" begin notation B.in_hom ("\_ : _ \\<^sub>B _\") notation comp (infixr "\" 55) definition hcomp (infixr "\" 53) where "g \ f = (if arr f \ arr g \ trg\<^sub>B f = src\<^sub>B g then g \\<^sub>B f else null)" definition src where "src \ = (if arr \ then src\<^sub>B \ else null)" definition trg where "trg \ = (if arr \ then trg\<^sub>B \ else null)" interpretation src: endofunctor \(\)\ src using src_def null_char inclusion arr_char src_closed trg_closed dom_closed cod_closed dom_simp cod_simp apply unfold_locales apply auto[4] by (metis B.src.as_nat_trans.preserves_comp_2 comp_char seq_char) interpretation trg: endofunctor \(\)\ trg using trg_def null_char inclusion arr_char src_closed trg_closed dom_closed cod_closed dom_simp cod_simp apply unfold_locales apply auto[4] by (metis B.trg.as_nat_trans.preserves_comp_2 comp_char seq_char) interpretation horizontal_homs \(\)\ src trg using src_def trg_def src.preserves_arr trg.preserves_arr null_char ide_char arr_char inclusion by (unfold_locales, simp_all) interpretation "functor" VV.comp \(\)\ \\\\. fst \\ \ snd \\\ using hcomp_def VV.arr_char src_def trg_def arr_char hcomp_closed dom_char cod_char VV.dom_char VV.cod_char apply unfold_locales apply auto[2] proof - fix f assume f: "VV.arr f" show "dom (fst f \ snd f) = fst (VV.dom f) \ snd (VV.dom f)" proof - have "dom (fst f \ snd f) = B.dom (fst f) \\<^sub>B B.dom (snd f)" proof - have "dom (fst f \ snd f) = B.dom (fst f \ snd f)" using f dom_char by (simp add: arr_char hcomp_closed hcomp_def) also have "... = B.dom (fst f) \\<^sub>B B.dom (snd f)" using f by (metis (no_types, lifting) B.hcomp_simps(3) B.hseqI' VV.arrE arrE hcomp_def inclusion src_def trg_def) finally show ?thesis by blast qed also have "... = fst (VV.dom f) \ snd (VV.dom f)" using f VV.arr_char VV.dom_char arr_char hcomp_def B.seq_if_composable dom_closed apply simp by (metis (no_types, lifting) dom_char) finally show ?thesis by simp qed show "cod (fst f \ snd f) = fst (VV.cod f) \ snd (VV.cod f)" proof - have "cod (fst f \ snd f) = B.cod (fst f) \\<^sub>B B.cod (snd f)" using f VV.arr_char arr_char cod_char hcomp_def src_def trg_def src_closed trg_closed hcomp_closed inclusion B.hseq_char arrE by auto also have "... = fst (VV.cod f) \ snd (VV.cod f)" using f VV.arr_char VV.cod_char arr_char hcomp_def B.seq_if_composable cod_closed apply simp by (metis (no_types, lifting) cod_char) finally show ?thesis by simp qed next fix f g assume fg: "VV.seq g f" show "fst (VV.comp g f) \ snd (VV.comp g f) = (fst g \ snd g) \ (fst f \ snd f)" proof - have "fst (VV.comp g f) \ snd (VV.comp g f) = fst g \ fst f \ snd g \ snd f" using fg VV.seq_char VV.comp_char VxV.comp_char VxV.not_Arr_Null by (metis (no_types, lifting) VxV.seqE prod.sel(1) prod.sel(2)) also have "... = (fst g \\<^sub>B fst f) \\<^sub>B (snd g \\<^sub>B snd f)" using fg comp_char hcomp_def VV.seq_char inclusion arr_char seq_char B.hseq_char by (metis (no_types, lifting) B.hseq_char' VxV.seq_char null_char) also have 1: "... = (fst g \\<^sub>B snd g) \\<^sub>B (fst f \\<^sub>B snd f)" proof - have "src\<^sub>B (fst g) = trg\<^sub>B (snd g)" by (metis (no_types, lifting) VV.arrE VV.seq_char fg src_def trg_def) thus ?thesis using fg VV.seq_char VV.arr_char arr_char seq_char inclusion B.interchange by (meson VxV.seqE) qed also have "... = (fst g \ snd g) \ (fst f \ snd f)" using fg comp_char hcomp_def VV.seq_char VV.arr_char arr_char seq_char inclusion B.hseq_char' hcomp_closed src_def trg_def by (metis (no_types, lifting) 1) finally show ?thesis by auto qed qed interpretation horizontal_composition \(\)\ \(\)\ src trg using arr_char src_def trg_def src_closed trg_closed apply (unfold_locales) using hcomp_def inclusion not_arr_null by auto abbreviation \ where "\ \ \ \ \ if VVV.arr (\, \, \) then \\<^sub>B \ \ \ else null" abbreviation (input) \\<^sub>S\<^sub>B where "\\<^sub>S\<^sub>B \\\ \ \ (fst \\\) (fst (snd \\\)) (snd (snd \\\))" lemma assoc_closed': assumes "VVV.arr \\\" shows "Arr (\\<^sub>S\<^sub>B \\\)" proof - have 1: "B.VVV.arr \\\" using assms VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char arr_char src_def trg_def inclusion by auto show "Arr (\\<^sub>S\<^sub>B \\\)" proof - have "Arr (\\<^sub>S\<^sub>B \\\) = Arr ((fst \\\ \\<^sub>B fst (snd \\\) \\<^sub>B snd (snd \\\)) \\<^sub>B \\<^sub>S\<^sub>B (B.VVV.dom \\\))" using assms 1 B.\_def B.assoc_is_natural_1 [of "fst \\\" "fst (snd \\\)" "snd (snd \\\)"] VV.arr_char VVV.arr_char B.VVV.arr_char B.VV.arr_char B.VVV.dom_char B.VV.dom_char apply simp by (metis (no_types, lifting) arr_char dom_char dom_closed src.preserves_dom trg.preserves_dom) also have "..." proof (intro comp_closed) show "Arr (fst \\\ \\<^sub>B fst (snd \\\) \\<^sub>B snd (snd \\\))" using assms 1 B.VVV.arr_char B.VV.arr_char hcomp_closed by (metis (no_types, lifting) B.H.preserves_reflects_arr B.trg_hcomp VV.arr_char VVV.arrE arr_char) show "B.cod (\ (fst (B.VVV.dom \\\)) (fst (snd (B.VVV.dom \\\))) (snd (snd (B.VVV.dom \\\)))) = B.dom (fst \\\ \\<^sub>B fst (snd \\\) \\<^sub>B snd (snd \\\))" using assms 1 VVV.arr_char VV.arr_char B.VxVxV.dom_char B.VVV.dom_simp B.VVV.cod_simp apply simp by (metis (no_types, lifting) B.VV.arr_char B.VVV.arrE B.\.preserves_reflects_arr B.assoc_is_natural_1 B.seqE arr_dom dom_char src_dom trg_dom) show "Arr (\ (fst (B.VVV.dom \\\)) (fst (snd (B.VVV.dom \\\))) (snd (snd (B.VVV.dom \\\))))" proof - have "VVV.arr (B.VVV.dom \\\)" using 1 B.VVV.dom_char B.VVV.arr_char B.VV.arr_char VVV.arr_char VV.arr_char apply simp by (metis (no_types, lifting) VVV.arrE arr_dom assms dom_simp src_dom trg_dom) moreover have "Arr (\\<^sub>B (B.dom (fst \\\)) (B.dom (fst (snd \\\))) (B.dom (snd (snd \\\))))" proof - have "B.VVV.ide (B.VVV.dom \\\)" using 1 B.VVV.ide_dom by blast thus ?thesis using assms B.\_def B.VVV.arr_char B.VV.arr_char B.VVV.ide_char B.VV.ide_char dom_closed assoc_closed by (metis (no_types, lifting) "1" B.ide_dom B.src_dom B.trg_dom VV.arr_char VVV.arrE arr_char) qed ultimately show ?thesis using 1 B.VVV.ide_dom assoc_closed B.VVV.dom_char apply simp by (metis (no_types, lifting) B.VV.arr_char B.VVV.arrE B.VVV.inclusion B.VxV.dom_char B.VxVxV.arrE B.VxVxV.dom_char prod.sel(1) prod.sel(2)) qed qed finally show ?thesis by blast qed qed lemma lunit_closed': assumes "Arr f" shows "Arr (B.\ f)" proof - have 1: "arr f \ arr (B.\ (B.dom f))" using assms arr_char lunit_closed dom_closed B.ide_dom inclusion by simp moreover have "B.dom f = B.cod (B.\ (B.dom f))" using 1 arr_char B.\.preserves_cod inclusion by simp moreover have "B.\ f = f \ B.\ (B.dom f)" using assms 1 B.\.is_natural_1 inclusion comp_char arr_char by simp ultimately show ?thesis using arr_char comp_closed cod_char seqI dom_simp by auto qed lemma runit_closed': assumes "Arr f" shows "Arr (B.\ f)" proof - have 1: "arr f \ arr (B.\ (B.dom f))" using assms arr_char runit_closed dom_closed B.ide_dom inclusion by simp moreover have "B.dom f = B.cod (B.\ (B.dom f))" using 1 arr_char B.\.preserves_cod inclusion by simp moreover have "B.\ f = f \ B.\ (B.dom f)" using assms 1 B.\.is_natural_1 inclusion comp_char arr_char by simp ultimately show ?thesis using arr_char comp_closed cod_char seqI dom_simp by auto qed interpretation natural_isomorphism VVV.comp \(\)\ HoHV HoVH \\<^sub>S\<^sub>B proof fix \\\ show "\ VVV.arr \\\ \ \\<^sub>S\<^sub>B \\\ = null" by simp assume \\\: "VVV.arr \\\" have 1: "B.VVV.arr \\\" using \\\ VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char arr_char src_def trg_def inclusion by auto show "dom (\\<^sub>S\<^sub>B \\\) = HoHV (VVV.dom \\\)" proof - have "dom (\\<^sub>S\<^sub>B \\\) = B.HoHV (B.VVV.dom \\\)" using \\\ 1 arr_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char B.\_def assoc_closed' dom_simp by simp also have "... = HoHV (VVV.dom \\\)" proof - have "HoHV (VVV.dom \\\) = HoHV (VxVxV.dom \\\)" using \\\ VVV.dom_char VV.arr_char src_def trg_def VVV.arr_char by auto also have "... = B.HoHV (B.VVV.dom \\\)" using \\\ VVV.dom_char VVV.arr_char VV.arr_char src_def trg_def HoHV_def B.HoHV_def arr_char B.VVV.arr_char B.VVV.dom_char B.VV.arr_char dom_closed hcomp_closed hcomp_def inclusion dom_simp by auto finally show ?thesis by simp qed finally show ?thesis by simp qed show "cod (\\<^sub>S\<^sub>B \\\) = HoVH (VVV.cod \\\)" proof - have "cod (\\<^sub>S\<^sub>B \\\) = B.HoVH (B.VVV.cod \\\)" using \\\ 1 arr_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char B.\_def assoc_closed' cod_simp by simp also have "... = HoVH (VVV.cod \\\)" proof - have "HoVH (VVV.cod \\\) = HoVH (VxVxV.cod \\\)" using \\\ VVV.cod_char VV.arr_char src_def trg_def VVV.arr_char by auto also have "... = B.HoVH (B.VVV.cod \\\)" using \\\ VVV.cod_char VV.arr_char src_def trg_def VVV.arr_char HoVH_def B.HoVH_def arr_char B.VVV.arr_char B.VVV.cod_char B.VV.arr_char cod_closed hcomp_closed hcomp_def inclusion cod_simp by simp finally show ?thesis by simp qed finally show ?thesis by simp qed have 3: "Arr (fst \\\) \ Arr (fst (snd \\\)) \ Arr (snd (snd \\\)) \ src\<^sub>B (fst \\\) = trg\<^sub>B (fst (snd \\\)) \ src\<^sub>B (fst (snd \\\)) = trg\<^sub>B (snd (snd \\\))" using \\\ VVV.arr_char VV.arr_char src_def trg_def arr_char by auto show "HoVH \\\ \ \\<^sub>S\<^sub>B (VVV.dom \\\) = \\<^sub>S\<^sub>B \\\" proof - have "\\<^sub>S\<^sub>B \\\ = (fst \\\ \\<^sub>B fst (snd \\\) \\<^sub>B snd (snd \\\)) \\<^sub>B \\<^sub>B (B.dom (fst \\\)) (B.dom (fst (snd \\\))) (B.dom (snd (snd \\\)))" using 3 inclusion B.assoc_is_natural_1 [of "fst \\\" "fst (snd \\\)" "snd (snd \\\)"] by (simp add: \\\) also have "... = (fst \\\ \ fst (snd \\\) \ snd (snd \\\)) \ \\<^sub>B (dom (fst \\\)) (dom (fst (snd \\\))) (dom (snd (snd \\\)))" using 1 3 \\\ hcomp_closed assoc_closed dom_closed hcomp_def comp_def inclusion comp_char dom_char VVV.arr_char VV.arr_char apply simp using B.hcomp_simps(2-3) arr_char by presburger also have "... = HoVH \\\ \ \\<^sub>S\<^sub>B (VVV.dom \\\)" using \\\ B.\_def HoVH_def VVV.dom_char VV.dom_char VxVxV.dom_char apply simp by (metis (no_types, lifting) VV.arr_char VVV.arrE VVV.arr_dom VxV.dom_char dom_simp) finally show ?thesis by argo qed show "\\<^sub>S\<^sub>B (VVV.cod \\\) \ HoHV \\\ = \\<^sub>S\<^sub>B \\\" proof - have "\\<^sub>S\<^sub>B \\\ = \\<^sub>B (B.cod (fst \\\)) (B.cod (fst (snd \\\))) (B.cod (snd (snd \\\))) \\<^sub>B (fst \\\ \\<^sub>B fst (snd \\\)) \\<^sub>B snd (snd \\\)" using 3 inclusion B.assoc_is_natural_2 [of "fst \\\" "fst (snd \\\)" "snd (snd \\\)"] by (simp add: \\\) also have "... = \\<^sub>B (cod (fst \\\)) (cod (fst (snd \\\))) (cod (snd (snd \\\))) \ - ((fst \\\ \ fst (snd \\\)) \ snd (snd \\\)) " + ((fst \\\ \ fst (snd \\\)) \ snd (snd \\\))" using 1 3 \\\ hcomp_closed assoc_closed cod_closed hcomp_def comp_def inclusion comp_char cod_char VVV.arr_char VV.arr_char by auto also have "... = \\<^sub>S\<^sub>B (VVV.cod \\\) \ HoHV \\\" using \\\ B.\_def HoHV_def VVV.cod_char VV.cod_char VxVxV.cod_char VVV.arr_char VV.arr_char arr_cod src_cod trg_cod by simp finally show ?thesis by argo qed next fix fgh assume fgh: "VVV.ide fgh" show "iso (\\<^sub>S\<^sub>B fgh)" proof - have 1: "B.arr (fst (snd fgh)) \ B.arr (snd (snd fgh)) \ src\<^sub>B (fst (snd fgh)) = trg\<^sub>B (snd (snd fgh)) \ src\<^sub>B (fst fgh) = trg\<^sub>B (fst (snd fgh))" using fgh VVV.ide_char VVV.arr_char VV.arr_char src_def trg_def arr_char inclusion by auto have 2: "B.ide (fst fgh) \ B.ide (fst (snd fgh)) \ B.ide (snd (snd fgh))" using fgh VVV.ide_char ide_char by blast have "\\<^sub>S\<^sub>B fgh = \\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh))" using fgh B.\_def by simp moreover have "B.VVV.ide fgh" using fgh 1 2 VVV.ide_char B.VVV.ide_char VVV.arr_char B.VVV.arr_char src_def trg_def inclusion arr_char B.VV.arr_char by simp moreover have "Arr (\\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh)))" using fgh 1 VVV.ide_char VVV.arr_char VV.arr_char src_def trg_def arr_char assoc_closed' B.\_def by simp moreover have "Arr (B.inv (\\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh))))" using fgh 1 VVV.ide_char VVV.arr_char VV.arr_char src_def trg_def arr_char assoc'_closed by (simp add: VVV.arr_char "2" B.VVV.ide_char calculation(2)) ultimately show ?thesis using fgh iso_char B.\.components_are_iso by auto qed qed interpretation L: endofunctor \(\)\ L using endofunctor_L by auto interpretation R: endofunctor \(\)\ R using endofunctor_R by auto interpretation L: faithful_functor \(\)\ \(\)\ L proof fix f f' assume par: "par f f'" assume eq: "L f = L f'" have "B.par f f'" using par inclusion arr_char dom_simp cod_simp by fastforce moreover have "B.L f = B.L f'" proof - have "\a. Arr a \ B.arr a" by (simp add: inclusion) moreover have 1: "\a. arr a \ (if arr a then hseq (trg a) a else arr null)" using L.preserves_arr by presburger moreover have "Arr f \ Arr (trg f) \ trg\<^sub>B f = src\<^sub>B (trg f)" by (simp add: \B.par f f'\ arrE par trg_closed trg_def) ultimately show ?thesis by (metis \B.par f f'\ eq hcomp_def hseq_char' par trg_def) qed ultimately show "f = f'" using B.L.is_faithful by blast qed interpretation L: full_functor \(\)\ \(\)\ L proof fix f f' \ assume f: "ide f" and f': "ide f'" and \: "\\ : L f \ L f'\" have 1: "L f = trg\<^sub>B f \\<^sub>B f \ L f' = trg\<^sub>B f' \\<^sub>B f'" using f f' hcomp_def trg_def arr_char ide_char trg_closed by simp have 2: "\\ : trg\<^sub>B f \\<^sub>B f \\<^sub>B trg\<^sub>B f' \\<^sub>B f'\" using 1 f f' \ hcomp_def trg_def src_def inclusion dom_char cod_char hseq_char' arr_char ide_char trg_closed null_char by (simp add: arr_char in_hom_char) show "\\. \\ : f \ f'\ \ L \ = \" proof - let ?\ = "B.\ f' \\<^sub>B \ \\<^sub>B B.inv (B.\ f)" have \: "\?\ : f \ f'\ \ \?\ : f \\<^sub>B f'\" proof - have "\?\ : f \\<^sub>B f'\" using f f' \ 2 B.\_ide_simp lunit'_closed lunit_closed' ide_char by auto thus ?thesis using f f' \ in_hom_char arr_char comp_closed ide_char lunit'_closed lunit_closed by (metis (no_types, lifting) B.arrI B.seqE in_homE) qed have \_eq: "?\ = B.\ f' \ \ \ B.inv (B.\ f)" proof - have "?\ = B.\ f' \ \ \\<^sub>B B.inv (B.\ f)" using f f' \ \ arr_char inclusion comp_char comp_closed ide_char lunit'_closed lunit_closed by (metis (no_types, lifting) B.seqE in_homE) also have "... = B.\ f' \ \ \ B.inv (B.\ f)" using f f' \ \ arr_char inclusion comp_char comp_closed ide_char lunit'_closed lunit_closed by (metis (no_types, lifting) B.seqE in_homE) finally show ?thesis by simp qed have "L ?\ = \" proof - have "L ?\ = trg\<^sub>B ?\ \\<^sub>B ?\" using \ \_eq hcomp_def trg_def inclusion arr_char trg_closed by auto also have "... = (trg\<^sub>B ?\ \\<^sub>B ?\) \\<^sub>B (B.inv (B.\ f) \\<^sub>B B.\ f)" proof - have "B.inv (B.\ f) \\<^sub>B B.\ f = trg\<^sub>B f \\<^sub>B f" using f ide_char B.comp_inv_arr B.inv_is_inverse by auto moreover have "B.dom (trg\<^sub>B ?\ \\<^sub>B ?\) = trg\<^sub>B f \\<^sub>B f" proof - have "B.dom (trg\<^sub>B ?\) = trg\<^sub>B f" using f \ B.vconn_implies_hpar(2) by force moreover have "B.dom ?\ = f" using \ by blast ultimately show ?thesis using B.hcomp_simps [of "trg\<^sub>B ?\" ?\] by (metis (no_types, lifting) B.hseqI' B.ideD(1) B.src_trg B.trg.preserves_reflects_arr B.trg_dom f ide_char) qed ultimately show ?thesis using \ \_eq B.comp_arr_dom in_hom_char by auto qed also have "... = ((trg\<^sub>B ?\ \\<^sub>B ?\) \\<^sub>B B.inv (B.\ f)) \\<^sub>B B.\ f" using B.comp_assoc by simp also have "... = (B.inv (B.\ f') \\<^sub>B ?\) \\<^sub>B B.\ f" using \ \_eq B.\'.naturality [of ?\] by auto also have "... = (B.inv (B.\ f') \\<^sub>B B.\ f') \\<^sub>B \ \\<^sub>B (B.inv (B.\ f) \\<^sub>B B.\ f)" using \ \_eq arr_char arrI comp_simp B.comp_assoc by metis also have "... = \" using f f' \ 2 B.comp_arr_dom B.comp_cod_arr ide_char B.\.components_are_iso B.\_ide_simp B.comp_inv_arr' by auto finally show ?thesis by blast qed thus ?thesis using \ by auto qed qed interpretation R: faithful_functor \(\)\ \(\)\ R proof fix f f' assume par: "par f f'" assume eq: "R f = R f'" have "B.par f f'" using par inclusion arr_char dom_simp cod_simp by fastforce moreover have "B.R f = B.R f'" proof - have "\a. Arr a \ B.arr a" by (simp add: inclusion) moreover have 1: "\a. arr a \ (if arr a then hseq a (src a) else arr null)" using R.preserves_arr by presburger moreover have "arr f \ arr (src f) \ trg\<^sub>B (src f) = src\<^sub>B f" by (meson 1 hcomp_def hseq_char' par) ultimately show ?thesis by (metis \B.par f f'\ eq hcomp_def hseq_char' src_def) qed ultimately show "f = f'" using B.R.is_faithful by blast qed interpretation R: full_functor \(\)\ \(\)\ R proof fix f f' \ assume f: "ide f" and f': "ide f'" and \: "\\ : R f \ R f'\" have 1: "R f = f \\<^sub>B src\<^sub>B f \ R f' = f' \\<^sub>B src\<^sub>B f'" using f f' hcomp_def src_def arr_char ide_char src_closed by simp have 2: "\\ : f \\<^sub>B src\<^sub>B f \\<^sub>B f' \\<^sub>B src\<^sub>B f'\" using 1 f f' \ hcomp_def trg_def src_def inclusion dom_char cod_char hseq_char' arr_char ide_char trg_closed null_char by (simp add: arr_char in_hom_char) show "\\. \\ : f \ f'\ \ R \ = \" proof - let ?\ = "B.\ f' \\<^sub>B \ \\<^sub>B B.inv (B.\ f)" have \: "\?\ : f \ f'\ \ \?\ : f \\<^sub>B f'\" proof - have "\?\ : f \\<^sub>B f'\" using f f' \ 2 B.\_ide_simp runit'_closed runit_closed' ide_char by auto thus ?thesis - using f f' \ in_hom_char [of ?\ f f'] arr_char comp_closed ide_char - runit'_closed runit_closed - apply auto - by fastforce + by (metis (no_types, lifting) B.arrI B.seqE \ arrE arrI comp_closed f f' + ide_char in_hom_char runit'_closed runit_closed') qed have \_eq: "?\ = B.\ f' \ \ \ B.inv (B.\ f)" proof - have "?\ = B.\ f' \ \ \\<^sub>B B.inv (B.\ f)" using f f' \ \ arr_char inclusion comp_char comp_closed ide_char runit'_closed runit_closed by (metis (no_types, lifting) B.seqE in_homE) also have "... = B.\ f' \ \ \ B.inv (B.\ f)" using f f' \ \ arr_char inclusion comp_char comp_closed ide_char runit'_closed runit_closed by (metis (no_types, lifting) B.arrI B.comp_in_homE in_hom_char) finally show ?thesis by simp qed have "R ?\ = \" proof - have "R ?\ = ?\ \\<^sub>B src\<^sub>B ?\" using \ \_eq hcomp_def src_def inclusion arr_char src_closed by auto also have "... = (?\ \\<^sub>B src\<^sub>B ?\) \\<^sub>B (B.inv (B.\ f) \\<^sub>B B.\ f)" proof - have "B.inv (B.\ f) \\<^sub>B B.\ f = f \\<^sub>B src\<^sub>B f" using f ide_char B.comp_inv_arr B.inv_is_inverse by auto moreover have "B.dom (?\ \\<^sub>B src\<^sub>B ?\) = f \\<^sub>B src\<^sub>B f" using f \ \_eq ide_char arr_char B.src_dom [of ?\] by (metis (no_types, lifting) B.R.as_nat_trans.preserves_comp_2 B.R.preserves_seq B.dom_src B.hcomp_simps(3) B.in_homE) ultimately show ?thesis using \ \_eq B.comp_arr_dom in_hom_char by auto qed also have "... = ((?\ \\<^sub>B src\<^sub>B ?\) \\<^sub>B B.inv (B.\ f)) \\<^sub>B B.\ f" using B.comp_assoc by simp also have "... = (B.inv (B.\ f') \\<^sub>B ?\) \\<^sub>B B.\ f" using \ \_eq B.\'.naturality [of ?\] by auto also have "... = (B.inv (B.\ f') \\<^sub>B B.\ f') \\<^sub>B \ \\<^sub>B (B.inv (B.\ f) \\<^sub>B B.\ f)" using \ \_eq arr_char arrI comp_simp B.comp_assoc by metis also have "... = \" using f f' \ 2 B.comp_arr_dom B.comp_cod_arr ide_char B.\.components_are_iso B.\_ide_simp B.comp_inv_arr' by auto finally show ?thesis by blast qed thus ?thesis using \ by blast qed qed interpretation bicategory \(\)\ \(\)\ \ \ src trg proof show "\a. obj a \ \\[a] : a \ a \ a\" proof (intro in_homI) fix a assume a: "obj a" have 1: "Arr (\ a)" using a obj_def src_def trg_def in_hom_char B.unit_in_hom arr_char hcomp_def B.obj_def ide_char objE hcomp_closed by (metis (no_types, lifting) B.\_ide_simp B.unitor_coincidence(1) inclusion lunit_closed) show 2: "arr \[a]" using 1 arr_char by simp show "dom \[a] = a \ a" using a 2 dom_char by (metis (full_types) B.objI_trg B.unit_simps(4) R.preserves_reflects_arr hcomp_def hseq_char' inclusion objE obj_simps(1) subcategory.arrE subcategory_axioms trg_def) show "cod \[a] = a" using a 2 cod_char by (metis B.obj_def' B.unit_simps(5) inclusion objE obj_simps(1) subcategory.arrE subcategory_axioms trg_def) qed show "\a. obj a \ iso (\ a)" proof - fix a assume a: "obj a" have 1: "trg\<^sub>B a = src\<^sub>B a" using a obj_def src_def trg_def B.obj_def arr_char by (metis horizontal_homs.objE horizontal_homs_axioms) have 2: "Arr (\ a)" using a 1 obj_def src_def trg_def in_hom_char B.unit_in_hom arr_char hcomp_def B.obj_def ide_char objE hcomp_closed by (metis (no_types, lifting) B.\_ide_simp B.unitor_coincidence(1) inclusion lunit_closed) have "iso (B.\ a)" using a 2 obj_def B.iso_unit iso_char arr_char lunit_closed lunit'_closed B.iso_lunit apply simp by (metis (no_types, lifting) B.\.components_are_iso B.ide_src inclusion src_def) thus "iso (\ a)" using a 2 obj_def B.iso_unit iso_char arr_char B.unitor_coincidence apply simp by (metis (no_types, lifting) B.\_ide_simp B.ide_src B.obj_src inclusion src_def) qed show "\f g h k. \ ide f; ide g; ide h; ide k; src f = trg g; src g = trg h; src h = trg k \ \ (f \ \ g h k) \ \ f (g \ h) k \ (\ f g h \ k) = \ f g (h \ k) \ \ (f \ g) h k" using B.pentagon VVV.arr_char VV.arr_char hcomp_def assoc_closed arr_char comp_char hcomp_closed comp_closed ide_char inclusion src_def trg_def by simp qed proposition is_bicategory: shows "bicategory (\) (\) \ \ src trg" .. lemma obj_char: shows "obj a \ arr a \ B.obj a" proof assume a: "obj a" show "arr a \ B.obj a" using a obj_def B.obj_def src_def arr_char inclusion by metis next assume a: "arr a \ B.obj a" have "src a = a" using a src_def by auto thus "obj a" using a obj_def by simp qed lemma hcomp_char: shows "hcomp = (\f g. if arr f \ arr g \ src f = trg g then f \\<^sub>B g else null)" using hcomp_def src_def trg_def by metis lemma assoc_simp: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "\ f g h = \\<^sub>B f g h" using assms VVV.arr_char VV.arr_char by auto lemma assoc'_simp: assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h" shows "\' f g h = B.\' f g h" proof - have "\' f g h = B.inv (\\<^sub>B f g h)" using assms inv_char by fastforce also have "... = B.\' f g h" using assms ide_char src_def trg_def B.VVV.ide_char B.VVV.arr_char B.VV.arr_char by force finally show ?thesis by blast qed lemma lunit_simp: assumes "ide f" shows "lunit f = B.lunit f" proof - have "B.lunit f = lunit f" proof (intro lunit_eqI) show "ide f" by fact show 1: "\B.lunit f : trg f \ f \ f\" proof show 2: "arr (B.lunit f)" using assms arr_char lunit_closed by (simp add: arr_char B.\_ide_simp ide_char) show "dom (B.lunit f) = trg f \ f" using assms 2 dom_char hcomp_char ide_char src_trg trg.preserves_arr trg_def by auto show "cod (B.lunit f) = f" using assms 2 in_hom_char by (simp add: cod_simp ide_char) qed show "trg f \ B.lunit f = (\[trg f] \ f) \ \' (trg f) (trg f) f" proof - have "trg f \ B.lunit f = trg\<^sub>B f \\<^sub>B B.lunit f" using assms 1 arr_char hcomp_char by (metis (no_types, lifting) ideD(1) src_trg trg.preserves_reflects_arr trg_def vconn_implies_hpar(2,4)) also have "... = (\[trg f] \\<^sub>B f) \\<^sub>B B.\' (trg f) (trg f) f" using assms ide_char B.lunit_char(2) trg_def by simp also have "... = (\[trg f] \\<^sub>B f) \\<^sub>B \' (trg f) (trg f) f" using assms assoc'_simp [of "trg f" "trg f" f] by simp also have "... = (\[trg f] \ f) \\<^sub>B \' (trg f) (trg f) f" using assms hcomp_char by simp also have "... = (\[trg f] \ f) \ \' (trg f) (trg f) f" using assms seq_char [of "\[trg f] \ f" "\' (trg f) (trg f) f"] comp_char [of "\[trg f] \ f" "\' (trg f) (trg f) f"] by simp finally show ?thesis by blast qed qed thus ?thesis by simp qed lemma lunit'_simp: assumes "ide f" shows "lunit' f = B.lunit' f" using assms inv_char [of "lunit f"] lunit_simp by fastforce lemma runit_simp: assumes "ide f" shows "runit f = B.runit f" proof - have "B.runit f = runit f" proof (intro runit_eqI) show "ide f" by fact show 1: "\B.runit f : f \ src f \ f\" proof show 2: "arr (B.runit f)" using assms arr_char runit_closed by (simp add: arr_char B.\_ide_simp ide_char) show "dom (B.runit f) = f \ src f" using assms 2 dom_char hcomp_char by (metis (no_types, lifting) B.runit_simps(4) ide_char src.preserves_reflects_arr src_def trg_src) show "cod (B.runit f) = f" using assms 2 in_hom_char by (simp add: cod_simp ide_char) qed show "B.runit f \ src f = (f \ \[src f]) \ \ f (src f) (src f)" proof - have "B.runit f \ src f = B.runit f \\<^sub>B src\<^sub>B f" using assms 1 arr_char hcomp_char by (metis (no_types, lifting) ideD(1) src.preserves_reflects_arr src_def trg_src vconn_implies_hpar(1,3)) also have "... = (f \\<^sub>B \[src f]) \\<^sub>B \\<^sub>B f (src f) (src f)" using assms ide_char B.runit_char(2) src_def by simp also have "... = (f \\<^sub>B \[src f]) \\<^sub>B \ f (src f) (src f)" using assms assoc_simp by simp also have "... = (f \ \[src f]) \\<^sub>B \ f (src f) (src f)" using assms 1 hcomp_char by simp also have "... = (f \ \[src f]) \ \ f (src f) (src f)" proof - have "B.seq (f \ \[src f]) (\ f (src f) (src f))" using assms seq_char [of "f \ \[src f]" "\ f (src f) (src f)"] by simp thus ?thesis using assms comp_char [of "f \ \[src f]" "\ f (src f) (src f)"] by simp qed finally show ?thesis by blast qed qed thus ?thesis by simp qed lemma runit'_simp: assumes "ide f" shows "runit' f = B.runit' f" using assms inv_char [of "runit f"] runit_simp by fastforce lemma comp_eqI [intro]: assumes "seq f g" and "f = f'" and "g = g'" shows "f \ g = f' \\<^sub>B g'" using assms comp_char ext ext not_arr_null by auto lemma comp_eqI' [intro]: assumes "seq f g" and "f = f'" and "g = g'" shows "f \\<^sub>B g = f' \ g'" using assms comp_char ext ext not_arr_null by auto lemma hcomp_eqI [intro]: assumes "hseq f g" and "f = f'" and "g = g'" shows "f \ g = f' \\<^sub>B g'" using assms hcomp_char not_arr_null by auto lemma hcomp_eqI' [intro]: assumes "hseq f g" and "f = f'" and "g = g'" shows "f \\<^sub>B g = f' \ g'" using assms hcomp_char not_arr_null by auto lemma arr_compI: assumes "seq f g" shows "arr (f \\<^sub>B g)" using assms seq_char dom_char cod_char by (metis (no_types, lifting) comp_simp) lemma arr_hcompI: assumes "hseq f g" shows "arr (f \\<^sub>B g)" using assms hseq_char src_def trg_def hcomp_eqI' by auto end sublocale subbicategory \ bicategory \(\)\ \(\)\ \ \ src trg using is_bicategory by auto subsection "The Sub-bicategory of Endo-arrows of an Object" text \ We now consider the sub-bicategory consisting of all arrows having the same object \a\ both as their source and their target and we show that the resulting structure is a monoidal category. We actually prove a slightly more general result, in which the unit of the monoidal category is taken to be an arbitrary isomorphism \\\ : w \\<^sub>B w \ w\\ with \w\ isomorphic to \a\, rather than the particular choice \\\[a] : a \\<^sub>B a \ a\\ made by the ambient bicategory. \ locale subbicategory_at_object = B: bicategory V H \\<^sub>B \ src\<^sub>B trg\<^sub>B + subbicategory V H \\<^sub>B \ src\<^sub>B trg\<^sub>B \\\. B.arr \ \ src\<^sub>B \ = a \ trg\<^sub>B \ = a\ for V :: "'a comp" (infixr "\\<^sub>B" 55) and H :: "'a comp" (infixr "\\<^sub>B" 55) and \\<^sub>B :: "'a \ 'a \ 'a \ 'a" ("\\<^sub>B[_, _, _]") and \ :: "'a \ 'a" ("\[_]") and src\<^sub>B :: "'a \ 'a" and trg\<^sub>B :: "'a \ 'a" and a :: "'a" and w :: "'a" and \ :: "'a" + assumes obj_a: "B.obj a" and isomorphic_a_w: "B.isomorphic a w" and \_in_vhom: "\\ : w \\<^sub>B w \ w\" and \_is_iso: "B.iso \" begin notation hcomp (infixr "\" 53) lemma arr_simps: assumes "arr \" shows "src \ = a" and "trg \ = a" apply (metis (no_types, lifting) arrE assms src_def) by (metis (no_types, lifting) arrE assms trg_def) lemma \_simps [simp]: shows "arr \" and "src \ = a" and "trg \ = a" and "dom \ = w \\<^sub>B w" and "cod \ = w" using isomorphic_a_w \_in_vhom in_hom_char arr_simps by auto lemma ide_w: shows "B.ide w" using isomorphic_a_w B.isomorphic_def by auto lemma w_simps [simp]: shows "ide w" and "B.ide w" and "src w = a" and "trg w = a" and "src\<^sub>B w = a" and "trg\<^sub>B w = a" and "dom w = w" and "cod w = w" proof - show w: "ide w" using \_in_vhom ide_cod by blast show "B.ide w" using w ide_char by simp obtain \ where \: "\\ : a \\<^sub>B w\ \ B.iso \" using isomorphic_a_w B.isomorphic_def by auto show "src\<^sub>B w = a" using obj_a w \ B.src_cod by force show "trg\<^sub>B w = a" using obj_a w \ B.src_cod by force show "src w = a" using \src\<^sub>B w = a\ w ide_w src_def by simp show "trg w = a" using \src\<^sub>B w = a\ w ide_w trg_def by (simp add: \trg\<^sub>B w = a\) show "dom w = w" using w by simp show "cod w = w" using w by simp qed lemma VxV_arr_eq_VV_arr: shows "VxV.arr f \ VV.arr f" using inclusion VxV.arr_char VV.arr_char arr_char src_def trg_def by auto lemma VxV_comp_eq_VV_comp: shows "VxV.comp = VV.comp" proof - have "\f g. VxV.comp f g = VV.comp f g" proof - fix f g show "VxV.comp f g = VV.comp f g" unfolding VV.comp_def using VxV.comp_char arr_simps(1) arr_simps(2) apply (cases "seq (fst f) (fst g)", cases "seq (snd f) (snd g)") by (elim seqE) auto qed thus ?thesis by blast qed lemma VxVxV_arr_eq_VVV_arr: shows "VxVxV.arr f \ VVV.arr f" using VVV.arr_char VV.arr_char src_def trg_def inclusion arr_char by auto lemma VxVxV_comp_eq_VVV_comp: shows "VxVxV.comp = VVV.comp" proof - have "\f g. VxVxV.comp f g = VVV.comp f g" proof - fix f g show "VxVxV.comp f g = VVV.comp f g" proof (cases "VxVxV.seq f g") assume 1: "\ VxVxV.seq f g" have "VxVxV.comp f g = VxVxV.null" using 1 VxVxV.ext by blast also have "... = (null, null, null)" using VxVxV.null_char VxV.null_char by simp also have "... = VVV.null" using VVV.null_char VV.null_char by simp also have "... = VVV.comp f g" proof - have "\ VVV.seq f g" using 1 VVV.seq_char by blast thus ?thesis by (metis (no_types, lifting) VVV.ext) qed finally show ?thesis by simp next assume 1: "VxVxV.seq f g" have 2: "B.arr (fst f) \ B.arr (fst (snd f)) \ B.arr (snd (snd f)) \ src\<^sub>B (fst f) = a \ src\<^sub>B (fst (snd f)) = a \ src\<^sub>B (snd (snd f)) = a \ trg\<^sub>B (fst f) = a \ trg\<^sub>B (fst (snd f)) = a \ trg\<^sub>B (snd (snd f)) = a" using 1 VxVxV.seq_char VxV.seq_char arr_char by blast have 3: "B.arr (fst g) \ B.arr (fst (snd g)) \ B.arr (snd (snd g)) \ src\<^sub>B (fst g) = a \ src\<^sub>B (fst (snd g)) = a \ src\<^sub>B (snd (snd g)) = a \ trg\<^sub>B (fst g) = a \ trg\<^sub>B (fst (snd g)) = a \ trg\<^sub>B (snd (snd g)) = a" using 1 VxVxV.seq_char VxV.seq_char arr_char by blast have 4: "B.seq (fst f) (fst g) \ B.seq (fst (snd f)) (fst (snd g)) \ B.seq (snd (snd f)) (snd (snd g))" using 1 VxVxV.seq_char VxV.seq_char seq_char by blast have 5: "VxVxV.comp f g = (fst f \ fst g, fst (snd f) \ fst (snd g), snd (snd f) \ snd (snd g))" using 1 2 3 4 VxVxV.seqE VxVxV.comp_char VxV.comp_char seq_char arr_char by (metis (no_types, lifting)) also have "... = VVV.comp f g" using 1 VVV.comp_char VVV.arr_char VV.arr_char apply simp using 2 3 5 arrI arr_simps(1) arr_simps(2) by presburger finally show ?thesis by blast qed qed thus ?thesis by blast qed interpretation H: "functor" VxV.comp \(\)\ \\\\. fst \\ \ snd \\\ using H.functor_axioms hcomp_def VxV_comp_eq_VV_comp by simp interpretation H: binary_endofunctor \(\)\ \\\\. fst \\ \ snd \\\ .. lemma HoHV_eq_ToTC: shows "HoHV = H.ToTC" using HoHV_def H.ToTC_def VVV.arr_char VV.arr_char src_def trg_def inclusion arr_char by auto lemma HoVH_eq_ToCT: shows "HoVH = H.ToCT" using HoVH_def H.ToCT_def VVV.arr_char VV.arr_char src_def trg_def inclusion arr_char by auto interpretation ToTC: "functor" VxVxV.comp \(\)\ H.ToTC using HoHV_eq_ToTC VxVxV_comp_eq_VVV_comp HoHV.functor_axioms by simp interpretation ToCT: "functor" VxVxV.comp \(\)\ H.ToCT using HoVH_eq_ToCT VxVxV_comp_eq_VVV_comp HoVH.functor_axioms by simp interpretation \: natural_isomorphism VxVxV.comp \(\)\ H.ToTC H.ToCT \ unfolding \_def using \.natural_isomorphism_axioms HoHV_eq_ToTC HoVH_eq_ToCT \_def VxVxV_comp_eq_VVV_comp by simp interpretation L: endofunctor \(\)\ \\f. fst (w, f) \ snd (w, f)\ proof fix f show "\ arr f \ fst (w, f) \ snd (w, f) = null" using arr_char hcomp_def by auto assume f: "arr f" show "hseq (fst (w, f)) (snd (w, f))" using f hseq_char arr_char src_def trg_def \_in_vhom cod_char by simp show "dom (fst (w, f) \ snd (w, f)) = fst (w, dom f) \ snd (w, dom f)" using f arr_char hcomp_def dom_simp by simp show "cod (fst (w, f) \ snd (w, f)) = fst (w, cod f) \ snd (w, cod f)" using f arr_char hcomp_def cod_simp by simp next fix f g assume fg: "seq g f" show "fst (w, g \ f) \ snd (w, g \ f) = (fst (w, g) \ snd (w, g)) \ (fst (w, f) \ snd (w, f))" by (simp add: fg whisker_left) qed interpretation L': equivalence_functor \(\)\ \(\)\ \\f. fst (w, f) \ snd (w, f)\ proof - obtain \ where \: "\\ : w \\<^sub>B a\ \ B.iso \" using isomorphic_a_w B.isomorphic_symmetric by force have "\\ : w \ a\" using \ in_hom_char by (metis (no_types, lifting) B.in_homE B.src_cod B.src_src B.trg_cod B.trg_trg \_in_vhom arr_char arr_cod cod_simp) hence \: "\\ : w \\<^sub>B a\ \ B.iso \ \ \\ : w \ a\ \ iso \" using \ iso_char arr_char by auto interpret \: natural_isomorphism \(\)\ \(\)\ \\f. fst (w, f) \ snd (w, f)\ map \\f. \ f \ (\ \ dom f)\ proof fix \ show "\ arr \ \ \ \ \ (\ \ dom \) = null" using \ arr_char dom_char ext apply simp using comp_null(2) hcomp_def by fastforce assume \: "arr \" have 0: "in_hhom (dom \) a a" using \ arr_char src_dom trg_dom src_def trg_def dom_simp by simp have 1: "in_hhom \ a a" using \ arr_char src_dom trg_dom src_def trg_def by auto have 2: "hseq \ (B.dom \)" using \ 0 1 dom_simp by (intro hseqI) auto have 3: "seq (\ \) (\ \ dom \)" proof (intro seqI') show "\\ \ dom \ : w \ dom \ \ a \ dom \\" by (metis (no_types, lifting) 0 \ \ hcomp_in_vhom ide_dom ide_in_hom(2) in_hhom_def w_simps(3)) show "\\ \ : a \ dom \ \ cod \\" using \ 2 \.preserves_hom [of \ "dom \" "cod \"] arr_simps(2) arr_cod by fastforce qed show "dom (\ \ \ (\ \ dom \)) = fst (w, dom \) \ snd (w, dom \)" proof - have "dom (\ \ \ (\ \ dom \)) = dom \ \ dom \" using \ 3 hcomp_simps(3) dom_comp dom_dom apply (elim seqE) by auto also have "... = fst (w, dom \) \ snd (w, dom \)" using \_in_vhom \ by (metis (no_types, lifting) in_homE prod.sel(1) prod.sel(2)) finally show ?thesis by simp qed show "cod (\ \ \ (\ \ dom \)) = map (cod \)" proof - have "seq (\ \) (\ \ dom \)" using 3 by simp hence "cod (\ \ \ (\ \ dom \)) = cod (\ \)" using cod_comp by blast also have "... = map (cod \)" using \ by blast finally show ?thesis by blast qed show "map \ \ \ (dom \) \ (\ \ dom (dom \)) = \ \ \ (\ \ dom \)" proof - have "map \ \ \ (dom \) \ (\ \ dom (dom \)) = (map \ \ \ (dom \)) \ (\ \ dom \)" using \ comp_assoc by simp also have "... = \ \ \ (\ \ dom \)" using \ \ \.is_natural_1 by auto finally show ?thesis by blast qed show "(\ (cod \) \ (\ \ dom (cod \))) \ (fst (w, \) \ snd (w, \)) = \ \ \ (\ \ dom \)" proof - have "(\ (cod \) \ (\ \ dom (cod \))) \ (fst (w, \) \ snd (w, \)) = (\ (cod \) \ (\ \ B.cod \)) \ (w \ \)" using \ \ dom_char arr_char \_in_vhom cod_simp by simp also have "... = \ (cod \) \ (\ \ w \ B.cod \ \ \)" proof - have "seq \ w" using \ \_in_vhom w_simps(1) by blast moreover have 2: "seq (B.cod \) \" using \ seq_char cod_simp by (simp add: comp_cod_arr) moreover have "src \ = trg (B.cod \)" using \ \ 2 by (metis (no_types, lifting) arr_simps(2) seqE vconn_implies_hpar(1) w_simps(3)) ultimately show ?thesis using interchange comp_assoc by simp qed also have "... = \ (cod \) \ (\ \ \)" using \ \ \_in_vhom comp_arr_dom comp_cod_arr cod_simp apply (elim conjE in_homE) by auto also have "... = (\ (cod \) \ (cod \ \ \)) \ (\ \ dom \)" proof - have 1: "seq (cod \) \" using \ arr_cod_iff_arr dom_cod iso_is_arr seqI by presburger moreover have 2: "seq \ (dom \)" using \ by (simp add: comp_arr_dom) moreover have "src (cod \) = trg \" using \ \ arr_cod arr_simps(1-2) iso_is_arr by auto ultimately show ?thesis using 1 2 interchange [of "cod \" \ \ "dom \"] comp_arr_dom comp_cod_arr comp_assoc by fastforce qed also have "... = \ \ \ (\ \ dom \)" proof - have "L \ = cod \ \ \" using \ \ arr_simps(2) in_homE by auto hence "\ (cod \) \ (cod \ \ \) = \ \" using \ \.is_natural_2 [of \] by simp thus ?thesis by simp qed finally show ?thesis by simp qed next show "\f. ide f \ iso (\ f \ (\ \ dom f))" proof - fix f assume f: "ide f" have "iso (\ f)" using f iso_lunit by simp moreover have "iso (\ \ dom f)" using \ f src_def trg_def ide_char arr_char apply (intro iso_hcomp, simp_all) by (metis (no_types, lifting) in_homE) moreover have "seq (\ f) (\ \ dom f)" proof (intro seqI') show " \\ f : a \ f \ f\" using f lunit_in_hom(2) \_ide_simp ide_char arr_char trg_def by simp show "\\ \ dom f : w \ f \ a \ f\" using \ f ide_char arr_char hcomp_def src_def trg_def obj_a ide_in_hom in_hom_char by (intro hcomp_in_vhom, auto) qed ultimately show "iso (\ f \ (\ \ dom f))" using isos_compose by simp qed qed show "equivalence_functor (\) (\) (\f. fst (w, f) \ snd (w, f))" using \.natural_isomorphism_axioms L.isomorphic_to_identity_is_equivalence by simp qed interpretation L: equivalence_functor \(\)\ \(\)\ \\f. fst (cod \, f) \ snd (cod \, f)\ proof - have "(\f. fst (cod \, f) \ snd (cod \, f)) = (\f. fst (w, f) \ snd (w, f))" using \_in_vhom by simp thus "equivalence_functor (\) (\) (\f. fst (cod \, f) \ snd (cod \, f))" using L'.equivalence_functor_axioms by simp qed interpretation R: endofunctor \(\)\ \\f. fst (f, w) \ snd (f, w)\ proof fix f show "\ arr f \ fst (f, w) \ snd (f, w) = null" using arr_char hcomp_def by auto assume f: "arr f" show "hseq (fst (f, w)) (snd (f, w))" using f hseq_char arr_char src_def trg_def \_in_vhom cod_char isomorphic_a_w B.isomorphic_def in_hom_char by simp show "dom (fst (f, w) \ snd (f, w)) = fst (dom f, w) \ snd (dom f, w)" using f arr_char dom_char cod_char hcomp_def \_in_vhom by simp show "cod (fst (f, w) \ snd (f, w)) = fst (cod f, w) \ snd (cod f, w)" using f arr_char dom_char cod_char hcomp_def \_in_vhom by simp next fix f g assume fg: "seq g f" have 1: "a \\<^sub>B a = a" using obj_a by auto show "fst (g \ f, w) \ snd (g \ f, w) = (fst (g, w) \ snd (g, w)) \ (fst (f, w) \ snd (f, w))" by (simp add: fg whisker_right) qed interpretation R': equivalence_functor \(\)\ \(\)\ \\f. fst (f, w) \ snd (f, w)\ proof - obtain \ where \: "\\ : w \\<^sub>B a\ \ B.iso \" using isomorphic_a_w B.isomorphic_symmetric by force have "\\ : w \ a\" using \ in_hom_char by (metis (no_types, lifting) B.in_homE B.src_cod B.src_src B.trg_cod B.trg_trg \_in_vhom arr_char arr_cod cod_simp) hence \: "\\ : w \\<^sub>B a\ \ B.iso \ \ \\ : w \ a\ \ iso \" using \ iso_char arr_char by auto interpret \: natural_isomorphism comp comp \\f. fst (f, w) \ snd (f, w)\ map \\f. \ f \ (dom f \ \)\ proof fix \ show "\ arr \ \ \ \ \ (dom \ \ \) = null" using \ arr_char dom_char ext apply simp using comp_null(2) hcomp_def by fastforce assume \: "arr \" have 0: "in_hhom (dom \) a a" using \ arr_char src_dom trg_dom src_def trg_def dom_simp by simp have 1: "in_hhom \ a a" using \ arr_char src_dom trg_dom src_def trg_def by auto have 2: "hseq (B.dom \) \" using \ 0 1 dom_simp hseqI by auto have 3: "seq (\ \) (dom \ \ \)" proof (intro seqI') show "\dom \ \ \ : dom \ \ w \ dom \ \ a\" by (metis (no_types, lifting) "0" "1" \ \ hcomp_in_vhom hseqI hseq_char ide_dom ide_in_hom(2) vconn_implies_hpar(2)) show "\\ \ : dom \ \ a \ cod \\" using \ 2 \.preserves_hom [of \ "dom \" "cod \"] arr_simps(2) arr_cod dom_simp cod_simp by fastforce qed show "dom (\ \ \ (dom \ \ \)) = fst (dom \, w) \ snd (dom \, w)" proof - have "dom (\ \ \ (dom \ \ \)) = dom \ \ dom \" using \ 3 hcomp_simps(3) dom_comp dom_dom apply (elim seqE) by auto also have "... = fst (dom \, w) \ snd (dom \, w)" using \_in_vhom \ by (metis (no_types, lifting) in_homE prod.sel(1) prod.sel(2)) finally show ?thesis by simp qed show "cod (\ \ \ (dom \ \ \)) = map (cod \)" proof - have "seq (\ \) (dom \ \ \)" using 3 by simp hence "cod (\ \ \ (dom \ \ \)) = cod (\ \)" using cod_comp by blast also have "... = map (cod \)" using \ by blast finally show ?thesis by blast qed show "map \ \ \ (dom \) \ (dom (dom \) \ \) = \ \ \ (dom \ \ \)" proof - have "map \ \ \ (dom \) \ (dom (dom \) \ \) = (map \ \ \ (dom \)) \ (dom (dom \) \ \)" using comp_assoc by simp also have "... = (map \ \ \ (dom \)) \ (dom \ \ \)" using \ dom_dom by simp also have "... = \ \ \ (dom \ \ \)" using \ \ \.is_natural_1 by auto finally show ?thesis by blast qed show "(\ (cod \) \ (dom (cod \) \ \)) \ (fst (\, w) \ snd (\, w)) = \ \ \ (dom \ \ \)" proof - have "(\ (cod \) \ (dom (cod \) \ \)) \ (fst (\, w) \ snd (\, w)) = (\ (cod \) \ (B.cod \ \ \)) \ (\ \ w)" using \ \ dom_char arr_char \_in_vhom cod_simp by simp also have "... = \ (cod \) \ (B.cod \ \ \ \ \ \ w)" proof - have 2: "seq \ w" using \ \_in_vhom w_simps(1) by blast moreover have "seq (B.cod \) \" using \ seq_char cod_simp by (simp add: comp_cod_arr) moreover have "src (B.cod \) = trg \" using \ \ 2 using arr_simps(1) calculation(2) seq_char vconn_implies_hpar(2) by force ultimately show ?thesis using interchange comp_assoc by simp qed also have "... = \ (cod \) \ (\ \ \)" using \ \ \_in_vhom comp_arr_dom comp_cod_arr cod_simp apply (elim conjE in_homE) by auto also have "... = (\ (cod \) \ (\ \ cod \)) \ (dom \ \ \)" proof - have "(\ \ cod \) \ (dom \ \ \) = \ \ \" proof - have "seq \ (dom \)" using \ by (simp add: comp_arr_dom) moreover have "seq (cod \) \" using \ iso_is_arr arr_cod dom_cod by auto moreover have "src \ = trg (cod \)" using \ \ 2 by (metis (no_types, lifting) arr_simps(1) arr_simps(2) calculation(2) seqE) ultimately show ?thesis using \ \ iso_is_arr comp_arr_dom comp_cod_arr interchange [of \ "dom \" "cod \" \] by simp qed thus ?thesis using comp_assoc by simp qed also have "... = \ \ \ (dom \ \ \)" proof - have "\ \ cod \ = R \" using \ \ arr_simps(1) in_homE by auto hence "\ (cod \) \ (\ \ cod \) = \ \" using \ \ \.is_natural_2 by simp thus ?thesis by simp qed finally show ?thesis by simp qed next show "\f. ide f \ iso (\ f \ (dom f \ \))" proof - fix f assume f: "ide f" have 1: "iso (\ f)" using f iso_lunit by simp moreover have 2: "iso (dom f \ \)" using \ f src_def trg_def ide_char arr_char apply (intro iso_hcomp, simp_all) by (metis (no_types, lifting) in_homE) moreover have "seq (\ f) (dom f \ \)" proof (intro seqI') show "\\ f : f \ a \ f\" using f runit_in_hom(2) \_ide_simp ide_char arr_char src_def by simp show "\dom f \ \ : f \ w \ f \ a\" using \ f ide_char arr_char hcomp_def src_def trg_def obj_a ide_in_hom in_hom_char by (intro hcomp_in_vhom, auto) qed ultimately show "iso (\ f \ (dom f \ \))" using isos_compose by simp qed qed show "equivalence_functor (\) (\) (\f. fst (f, w) \ snd (f, w))" using \.natural_isomorphism_axioms R.isomorphic_to_identity_is_equivalence by simp qed interpretation R: equivalence_functor \(\)\ \(\)\ \\f. fst (f, cod \) \ snd (f, cod \)\ proof - have "(\f. fst (f, cod \) \ snd (f, cod \)) = (\f. fst (f, w) \ snd (f, w))" using \_in_vhom by simp thus "equivalence_functor (\) (\) (\f. fst (f, cod \) \ snd (f, cod \))" using R'.equivalence_functor_axioms by simp qed interpretation M: monoidal_category \(\)\ \\\\. fst \\ \ snd \\\ \ \ proof show "\\ : fst (cod \, cod \) \ snd (cod \, cod \) \ cod \\" using \_in_vhom hcomp_def arr_char by auto show "iso \" using \_is_iso iso_char arr_char inv_char \_in_vhom by auto show "\f g h k. \ ide f; ide g; ide h; ide k \ \ (fst (f, \ (g, h, k)) \ snd (f, \ (g, h, k))) \ \ (f, hcomp (fst (g, h)) (snd (g, h)), k) \ (fst (\ (f, g, h), k) \ snd (\ (f, g, h), k)) = \ (f, g, fst (h, k) \ snd (h, k)) \ \ (fst (f, g) \ snd (f, g), h, k)" proof - fix f g h k assume f: "ide f" and g: "ide g" and h: "ide h" and k: "ide k" have 1: "VVV.arr (f, g, h) \ VVV.arr (g, h, k)" using f g h k VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char by simp have 2: "VVV.arr (f, g \ h, k)" using f g h k 1 HoHV_def VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char VxV.arrI VxVxV.arrI VxVxV_comp_eq_VVV_comp hseqI' by auto have 3: "VVV.arr (f, g, h \ k)" using f g h k 1 VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char VxV.arrI VxVxV.arrI VxVxV_comp_eq_VVV_comp H.preserves_reflects_arr hseqI' by auto have 4: "VVV.arr (f \ g, h, k)" using f g h k VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char hseq_char VxV.arrI VxVxV.arrI VxVxV_comp_eq_VVV_comp by force have "(fst (f, \ (g, h, k)) \ snd (f, \ (g, h, k))) \ \ (f, fst (g, h) \ snd (g, h), k) \ (fst (\ (f, g, h), k) \ snd (\ (f, g, h), k)) = (f \ \\<^sub>B[g, h, k]) \ \\<^sub>B[f, g \ h, k] \ (\\<^sub>B[f, g, h] \ k)" unfolding \_def by (simp add: 1 2) also have "... = (f \\<^sub>B \\<^sub>B g h k) \ \\<^sub>B f (g \\<^sub>B h) k \ (\\<^sub>B f g h \\<^sub>B k)" unfolding hcomp_def using f g h k src_def trg_def arr_char using assoc_closed ide_char by auto also have "... = (f \\<^sub>B \\<^sub>B g h k) \\<^sub>B \\<^sub>B f (g \\<^sub>B h) k \\<^sub>B (\\<^sub>B f g h \\<^sub>B k)" proof - have "arr (f \\<^sub>B \\<^sub>B g h k)" using ide_char arr_char assoc_closed f g h hcomp_closed k by simp moreover have "arr (\\<^sub>B f (g \\<^sub>B h) k)" using ide_char arr_char assoc_closed f g h hcomp_closed k by simp moreover have "arr (\\<^sub>B f g h \\<^sub>B k)" using ide_char arr_char assoc_closed f g h hcomp_closed k by simp moreover have "arr (\\<^sub>B f (g \\<^sub>B h) k \\<^sub>B (\\<^sub>B f g h \\<^sub>B k))" unfolding arr_char apply (intro conjI) using ide_char arr_char assoc_closed f g h hcomp_closed k B.HoHV_def B.HoVH_def apply (intro B.seqI) apply simp_all proof - have 1: "B.arr (\\<^sub>B f (g \\<^sub>B h) k \\<^sub>B \\<^sub>B f g h \\<^sub>B k)" using f g h k ide_char arr_char B.HoHV_def B.HoVH_def apply (intro B.seqI) by auto show "src\<^sub>B (\\<^sub>B f (g \\<^sub>B h) k \\<^sub>B \\<^sub>B f g h \\<^sub>B k) = a" using 1 f g h k arr_char B.src_vcomp B.vseq_implies_hpar(1) by fastforce show "trg\<^sub>B (\\<^sub>B f (g \\<^sub>B h) k \\<^sub>B \\<^sub>B f g h \\<^sub>B k) = a" using "1" arr_char calculation(2-3) by auto qed ultimately show ?thesis using B.ext comp_char by (metis (no_types, lifting)) qed also have "... = \\<^sub>B f g (h \\<^sub>B k) \\<^sub>B \\<^sub>B (f \\<^sub>B g) h k" using f g h k src_def trg_def arr_char ide_char B.pentagon using "4" \_def hcomp_def by auto also have "... = \\<^sub>B f g (h \\<^sub>B k) \ \\<^sub>B (f \\<^sub>B g) h k" proof - have "arr (\\<^sub>B (f \\<^sub>B g) h k)" using ide_char arr_char assoc_closed f g h hcomp_closed k by simp moreover have "arr (\\<^sub>B f g (h \\<^sub>B k))" using ide_char arr_char assoc_closed f g h hcomp_closed k by fastforce ultimately show ?thesis using B.ext comp_char by auto qed also have "... = \\<^sub>B[f, g, fst (h, k) \ snd (h, k)] \ \\<^sub>B[fst (f, g) \ snd (f, g), h, k]" unfolding hcomp_def using f g h k src_def trg_def arr_char ide_char by simp also have "... = \ (f, g, fst (h, k) \ snd (h, k)) \ \ (fst (f, g) \ snd (f, g), h, k)" unfolding \_def using 1 2 3 4 by simp finally show "(fst (f, \ (g, h, k)) \ snd (f, \ (g, h, k))) \ \ (f, fst (g, h) \ snd (g, h), k) \ (fst (\ (f, g, h), k) \ snd (\ (f, g, h), k)) = \ (f, g, fst (h, k) \ snd (h, k)) \ \ (fst (f, g) \ snd (f, g), h, k)" by simp qed qed proposition is_monoidal_category: shows "monoidal_category (\) (\\\. fst \\ \ snd \\) \ \" .. end text \ In a bicategory, the ``objects'' are essentially arbitrarily chosen representatives of their isomorphism classes. Choosing any other representatives results in an equivalent structure. Each object \a\ is additionally equipped with an arbitrarily chosen unit isomorphism \\\ : a \ a \ a\\. For any \(a, \)\ and \(a', \')\, where \a\ and \a'\ are isomorphic to the same object, there exists a unique isomorphism \\\: a \ a'\\ that is compatible with the chosen unit isomorphisms \\\ and \\'\. We have already proved this property for monoidal categories, which are bicategories with just one ``object''. Here we use that already-proven property to establish its generalization to arbitary bicategories, by exploiting the fact that if \a\ is an object in a bicategory, then the sub-bicategory consisting of all \\\ such that \src \ = a = trg \\, is a monoidal category. At some point it would potentially be nicer to transfer the proof for monoidal categories to obtain a direct, ``native'' proof of this fact for bicategories. \ lemma (in bicategory) unit_unique_upto_unique_iso: assumes "obj a" and "isomorphic a w" and "\\ : w \ w \ w\" and "iso \" shows "\!\. \\ : a \ w\ \ iso \ \ \ \ \[a] = \ \ (\ \ \)" proof - have \_in_hhom: "\\ : a \ a\" using assms apply (intro in_hhomI) apply auto apply (metis src_cod in_homE isomorphic_implies_hpar(3) objE) by (metis trg_cod in_homE isomorphic_implies_hpar(4) objE) interpret S: subbicategory V H \ \ src trg \\\. arr \ \ src \ = a \ trg \ = a\ using assms iso_unit in_homE isoE isomorphicE VVV.arr_char VV.arr_char apply unfold_locales apply auto[7] proof fix f g h assume f: "(arr f \ src f = a \ trg f = a) \ ide f" and g: "(arr g \ src g = a \ trg g = a) \ ide g" and h: "(arr h \ src h = a \ trg h = a) \ ide h" and fg: "src f = trg g" and gh: "src g = trg h" show "arr (\[f, g, h])" using assms f g h fg gh by auto show "src (\[f, g, h]) = a \ trg (\[f, g, h]) = a" using assms f g h fg gh by auto show "arr (inv (\[f, g, h])) \ src (inv (\[f, g, h])) = a \ trg (inv (\[f, g, h])) = a" using assms f g h fg gh \.preserves_hom src_dom trg_dom by simp next fix f assume f: "arr f \ src f = a \ trg f = a" assume ide_left: "ide f" show "arr (\ f) \ src (\ f) = a \ trg (\ f) = a" using f assms(1) \.preserves_hom src_cod [of "\ f"] trg_cod [of "\ f"] by simp show "arr (inv (\ f)) \ src (inv (\ f)) = a \ trg (inv (\ f)) = a" using f ide_left assms(1) \'.preserves_hom src_dom [of "\'.map f"] trg_dom [of "\'.map f"] by simp show "arr (\ f) \ src (\ f) = a \ trg (\ f) = a" using f assms(1) \.preserves_hom src_cod [of "\ f"] trg_cod [of "\ f"] by simp show "arr (inv (\ f)) \ src (inv (\ f)) = a \ trg (inv (\ f)) = a" using f ide_left assms(1) \'.preserves_hom src_dom [of "\'.map f"] trg_dom [of "\'.map f"] by simp qed interpret S: subbicategory_at_object V H \ \ src trg a a \\[a]\ proof show "obj a" by fact show "isomorphic a a" using assms(1) isomorphic_reflexive by blast show "S.in_hom \[a] (a \ a) a" using S.arr_char S.in_hom_char assms(1) by fastforce show "iso \[a]" using assms iso_unit by simp qed interpret S\<^sub>\: subbicategory_at_object V H \ \ src trg a w \ proof show "obj a" by fact show "iso \" by fact show "isomorphic a w" using assms by simp show "S.in_hom \ (w \ w) w" using assms S.arr_char S.dom_char S.cod_char \_in_hhom by (intro S.in_homI, auto) qed interpret M: monoidal_category S.comp \\\\. S.hcomp (fst \\) (snd \\)\ S.\ \\[a]\ using S.is_monoidal_category by simp interpret M\<^sub>\: monoidal_category S.comp \\\\. S.hcomp (fst \\) (snd \\)\ S.\ \ using S\<^sub>\.is_monoidal_category by simp interpret M: monoidal_category_with_alternate_unit S.comp \\\\. S.hcomp (fst \\) (snd \\)\ S.\ \\[a]\ \ .. have 1: "M\<^sub>\.unity = w" using assms M\<^sub>\.unity_def S.cod_char S.arr_char by (metis (no_types, lifting) S.in_homE S\<^sub>\.\_in_vhom) have 2: "M.unity = a" using assms M.unity_def S.cod_char S.arr_char by simp have "\!\. S.in_hom \ a w \ S.iso \ \ S.comp \ \[a] = S.comp \ (M.tensor \ \)" using assms 1 2 M.unit_unique_upto_unique_iso M.unity_def M\<^sub>\.unity_def S.cod_char by simp show "\!\. \\ : a \ w\ \ iso \ \ \ \ \[a] = \ \ (\ \ \)" proof - have 1: "\\. S.in_hom \ a w \ \\ : a \ w\" using assms S.in_hom_char S.arr_char by (metis (no_types, lifting) S.ideD(1) S.w_simps(1) S\<^sub>\.w_simps(1) in_homE src_dom trg_dom) moreover have "\\. S.in_hom \ a w \ S.iso \ \ iso \" using assms S.in_hom_char S.arr_char S.iso_char by auto moreover have "\\. S.in_hom \ a w \ M.tensor \ \ = \ \ \" using assms S.in_hom_char S.arr_char S.hcomp_def by simp moreover have "\\. S.in_hom \ a w \ S.comp \ \[a] = \ \ \[a]" using assms S.in_hom_char S.comp_char by auto moreover have "\\. S.in_hom \ a w \ S.comp \ (M.tensor \ \) = \ \ (\ \ \)" using assms S.in_hom_char S.arr_char S.hcomp_def S.comp_char S.dom_char S.cod_char by (metis (no_types, lifting) M\<^sub>\.arr_tensor S\<^sub>\.\_simps(1) calculation(3) ext) ultimately show ?thesis by (metis (no_types, lifting) M.unit_unique_upto_unique_iso M.unity_def M\<^sub>\.unity_def S.\_in_vhom S.in_homE S\<^sub>\.\_in_vhom) qed qed end diff --git a/thys/Category3/EquivalenceOfCategories.thy b/thys/Category3/EquivalenceOfCategories.thy --- a/thys/Category3/EquivalenceOfCategories.thy +++ b/thys/Category3/EquivalenceOfCategories.thy @@ -1,915 +1,913 @@ (* Title: EquivalenceOfCategories Author: Eugene W. Stark , 2017 Maintainer: Eugene W. Stark *) chapter "Equivalence of Categories" text \ In this chapter we define the notions of equivalence and adjoint equivalence of categories and establish some properties of functors that are part of an equivalence. \ theory EquivalenceOfCategories imports Adjunction begin locale equivalence_of_categories = C: category C + D: category D + F: "functor" D C F + G: "functor" C D G + \: natural_isomorphism D D D.map "G o F" \ + \: natural_isomorphism C C "F o G" C.map \ for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and F :: "'d \ 'c" and G :: "'c \ 'd" and \ :: "'d \ 'd" and \ :: "'c \ 'c" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") lemma C_arr_expansion: assumes "C.arr f" shows "\ (C.cod f) \\<^sub>C F (G f) \\<^sub>C C.inv (\ (C.dom f)) = f" and "C.inv (\ (C.cod f)) \\<^sub>C f \\<^sub>C \ (C.dom f) = F (G f)" proof - have \_dom: "C.inverse_arrows (\ (C.dom f)) (C.inv (\ (C.dom f)))" using assms C.inv_is_inverse by auto have \_cod: "C.inverse_arrows (\ (C.cod f)) (C.inv (\ (C.cod f)))" using assms C.inv_is_inverse by auto have "\ (C.cod f) \\<^sub>C F (G f) \\<^sub>C C.inv (\ (C.dom f)) = (\ (C.cod f) \\<^sub>C F (G f)) \\<^sub>C C.inv (\ (C.dom f))" using C.comp_assoc by force also have 1: "... = (f \\<^sub>C \ (C.dom f)) \\<^sub>C C.inv (\ (C.dom f))" using assms \.naturality by simp also have 2: "... = f" using assms \_dom C.comp_arr_inv C.comp_arr_dom C.comp_assoc by force finally show "\ (C.cod f) \\<^sub>C F (G f) \\<^sub>C C.inv (\ (C.dom f)) = f" by blast show "C.inv (\ (C.cod f)) \\<^sub>C f \\<^sub>C \ (C.dom f) = F (G f)" using assms 1 2 \_dom \_cod C.invert_side_of_triangle C.isoI C.iso_inv_iso by metis qed lemma G_is_faithful: shows "faithful_functor C D G" proof fix f f' assume par: "C.par f f'" and eq: "G f = G f'" show "f = f'" proof - have "C.inv (\ (C.cod f)) \ C.hom (C.cod f) (F (G (C.cod f))) \ C.iso (C.inv (\ (C.cod f)))" using par by auto moreover have 1: "\ (C.dom f) \ C.hom (F (G (C.dom f))) (C.dom f) \ C.iso (\ (C.dom f))" using par by auto ultimately have 2: "f \\<^sub>C \ (C.dom f) = f' \\<^sub>C \ (C.dom f)" using par C_arr_expansion eq C.iso_is_section C.section_is_mono by (metis C_arr_expansion(1) eq) show ?thesis proof - have "C.epi (\ (C.dom f))" using 1 par C.iso_is_retraction C.retraction_is_epi by blast thus ?thesis using 2 par by auto qed qed qed lemma G_is_essentially_surjective: shows "essentially_surjective_functor C D G" proof fix b assume b: "D.ide b" have "C.ide (F b) \ D.isomorphic (G (F b)) b" proof show "C.ide (F b)" using b by simp show "D.isomorphic (G (F b)) b" proof (unfold D.isomorphic_def) have "\D.inv (\ b) : G (F b) \\<^sub>D b\ \ D.iso (D.inv (\ b))" using b by auto thus "\f. \f : G (F b) \\<^sub>D b\ \ D.iso f" by blast qed qed thus "\a. C.ide a \ D.isomorphic (G a) b" by blast qed interpretation \_inv: inverse_transformation C C \F o G\ C.map \ .. interpretation \_inv: inverse_transformation D D D.map \G o F\ \ .. interpretation GF: equivalence_of_categories D C G F \_inv.map \_inv.map .. lemma F_is_faithful: shows "faithful_functor D C F" using GF.G_is_faithful by simp lemma F_is_essentially_surjective: shows "essentially_surjective_functor D C F" using GF.G_is_essentially_surjective by simp lemma G_is_full: shows "full_functor C D G" proof fix a a' g assume a: "C.ide a" and a': "C.ide a'" assume g: "\g : G a \\<^sub>D G a'\" show "\f. \f : a \\<^sub>C a'\ \ G f = g" proof have \a: "C.inverse_arrows (\ a) (C.inv (\ a))" using a C.inv_is_inverse by auto have \a': "C.inverse_arrows (\ a') (C.inv (\ a'))" using a' C.inv_is_inverse by auto let ?f = "\ a' \\<^sub>C F g \\<^sub>C C.inv (\ a)" have f: "\?f : a \\<^sub>C a'\" using a a' g \a \a' \.preserves_hom [of a' a' a'] \_inv.preserves_hom [of a a a] by fastforce moreover have "G ?f = g" proof - interpret F: faithful_functor D C F using F_is_faithful by auto have "F (G ?f) = F g" proof - have "F (G ?f) = C.inv (\ a') \\<^sub>C ?f \\<^sub>C \ a" using f C_arr_expansion(2) [of "?f"] by auto also have "... = (C.inv (\ a') \\<^sub>C \ a') \\<^sub>C F g \\<^sub>C C.inv (\ a) \\<^sub>C \ a" using a a' f g C.comp_assoc by fastforce also have "... = F g" using a a' g \a \a' C.comp_inv_arr C.comp_arr_dom C.comp_cod_arr by auto finally show ?thesis by blast qed moreover have "D.par (G (\ a' \\<^sub>C F g \\<^sub>C C.inv (\ a))) g" using f g by fastforce ultimately show ?thesis using f g F.is_faithful by blast qed ultimately show "\?f : a \\<^sub>C a'\ \ G ?f = g" by blast qed qed end (* I'm not sure why I had to close and re-open the context here in order to * get the G_is_full fact in the interpretation GF. *) context equivalence_of_categories begin interpretation \_inv: inverse_transformation C C \F o G\ C.map \ .. interpretation \_inv: inverse_transformation D D D.map \G o F\ \ .. interpretation GF: equivalence_of_categories D C G F \_inv.map \_inv.map .. lemma F_is_full: shows "full_functor D C F" using GF.G_is_full by simp end text \ Traditionally the term "equivalence of categories" is also used for a functor that is part of an equivalence of categories. However, it seems best to use that term for a situation in which all of the structure of an equivalence is explicitly given, and to have a different term for one of the functors involved. \ locale equivalence_functor = C: category C + D: category D + "functor" C D G for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and G :: "'c \ 'd" + assumes induces_equivalence: "\F \ \. equivalence_of_categories C D F G \ \" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") end sublocale equivalence_of_categories \ equivalence_functor C D G using equivalence_of_categories_axioms by (unfold_locales, blast) text \ An equivalence functor is fully faithful and essentially surjective. \ sublocale equivalence_functor \ fully_faithful_functor C D G proof - obtain F \ \ where 1: "equivalence_of_categories C D F G \ \" using induces_equivalence by blast interpret equivalence_of_categories C D F G \ \ using 1 by auto show "fully_faithful_functor C D G" using G_is_full G_is_faithful fully_faithful_functor.intro by auto qed sublocale equivalence_functor \ essentially_surjective_functor C D G proof - obtain F \ \ where 1: "equivalence_of_categories C D F G \ \" using induces_equivalence by blast interpret equivalence_of_categories C D F G \ \ using 1 by auto show "essentially_surjective_functor C D G" using G_is_essentially_surjective by auto qed lemma (in inverse_functors) induce_equivalence: shows "equivalence_of_categories A B F G B.map A.map" using inv inv' A.is_extensional B.is_extensional B.comp_arr_dom B.comp_cod_arr A.comp_arr_dom A.comp_cod_arr by unfold_locales auto lemma (in invertible_functor) is_equivalence: shows "equivalence_functor A B G" using equivalence_functor_axioms.intro equivalence_functor_def equivalence_of_categories_def induce_equivalence by blast lemma (in identity_functor) is_equivalence: shows "equivalence_functor C C map" proof - interpret inverse_functors C C map map using map_def by unfold_locales auto interpret invertible_functor C C map using inverse_functors_axioms by unfold_locales blast show ?thesis using is_equivalence by blast qed text \ A special case of an equivalence functor is an endofunctor \F\ equipped with a natural isomorphism from \F\ to the identity functor. \ context endofunctor begin lemma isomorphic_to_identity_is_equivalence: assumes "natural_isomorphism A A F A.map \" shows "equivalence_functor A A F" proof - interpret \: natural_isomorphism A A F A.map \ using assms by auto interpret \': inverse_transformation A A F A.map \ .. interpret F\': natural_isomorphism A A F \F o F\ \F o \'.map\ proof - interpret F\': natural_transformation A A F \F o F\ \F o \'.map\ using \'.natural_transformation_axioms functor_axioms horizontal_composite [of A A A.map F \'.map A F F F] by simp show "natural_isomorphism A A F (F o F) (F o \'.map)" apply unfold_locales using \'.components_are_iso by fastforce qed interpret F\'o\': vertical_composite A A A.map F \F o F\ \'.map \F o \'.map\ .. interpret F\'o\': natural_isomorphism A A A.map \F o F\ F\'o\'.map using \'.natural_isomorphism_axioms F\'.natural_isomorphism_axioms natural_isomorphisms_compose by fast interpret inv_F\'o\': inverse_transformation A A A.map \F o F\ F\'o\'.map .. interpret F: equivalence_of_categories A A F F F\'o\'.map inv_F\'o\'.map .. show ?thesis .. qed end text \ An adjoint equivalence is an equivalence of categories that is also an adjunction. \ locale adjoint_equivalence = unit_counit_adjunction C D F G \ \ + \: natural_isomorphism D D D.map "G o F" \ + \: natural_isomorphism C C "F o G" C.map \ for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and F :: "'d \ 'c" and G :: "'c \ 'd" and \ :: "'d \ 'd" and \ :: "'c \ 'c" text \ An adjoint equivalence is clearly an equivalence of categories. \ sublocale adjoint_equivalence \ equivalence_of_categories .. context adjoint_equivalence begin text \ The triangle identities for an adjunction reduce to inverse relations when \\\ and \\\ are natural isomorphisms. \ lemma triangle_G': assumes "C.ide a" shows "D.inverse_arrows (\ (G a)) (G (\ a))" proof show "D.ide (G (\ a) \\<^sub>D \ (G a))" using assms triangle_G G\o\G.map_simp_ide by fastforce thus "D.ide (\ (G a) \\<^sub>D G (\ a))" using assms D.section_retraction_of_iso [of "G (\ a)" "\ (G a)"] by auto qed lemma triangle_F': assumes "D.ide b" shows "C.inverse_arrows (F (\ b)) (\ (F b))" proof show "C.ide (\ (F b) \\<^sub>C F (\ b))" using assms triangle_F \FoF\.map_simp_ide by auto thus "C.ide (F (\ b) \\<^sub>C \ (F b))" using assms C.section_retraction_of_iso [of "\ (F b)" "F (\ b)"] by auto qed text \ An adjoint equivalence can be dualized by interchanging the two functors and inverting the natural isomorphisms. This is somewhat awkward to prove, but probably useful to have done it once and for all. \ lemma dual_equivalence: assumes "adjoint_equivalence C D F G \ \" shows "adjoint_equivalence D C G F (inverse_transformation.map C C (C.map) \) (inverse_transformation.map D D (G o F) \)" proof - interpret adjoint_equivalence C D F G \ \ using assms by auto interpret \': inverse_transformation C C \F o G\ C.map \ .. interpret \': inverse_transformation D D D.map \G o F\ \ .. interpret G\': natural_transformation C D G \G o F o G\ \G o \'.map\ proof - have "natural_transformation C D G (G o (F o G)) (G o \'.map)" using G.as_nat_trans.natural_transformation_axioms \'.natural_transformation_axioms horizontal_composite by fastforce thus "natural_transformation C D G (G o F o G) (G o \'.map)" using o_assoc by metis qed interpret \'G: natural_transformation C D \G o F o G\ G \\'.map o G\ using \'.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms horizontal_composite by fastforce interpret \'F: natural_transformation D C F \F o G o F\ \\'.map o F\ using \'.natural_transformation_axioms F.as_nat_trans.natural_transformation_axioms horizontal_composite by fastforce interpret F\': natural_transformation D C \F o G o F\ F \F o \'.map\ proof - have "natural_transformation D C (F o (G o F)) F (F o \'.map)" using \'.natural_transformation_axioms F.as_nat_trans.natural_transformation_axioms horizontal_composite by fastforce thus "natural_transformation D C (F o G o F) F (F o \'.map)" using o_assoc by metis qed interpret F\'o\'F: vertical_composite D C F \(F o G) o F\ F \\'.map o F\ \F o \'.map\ .. interpret \'GoG\': vertical_composite C D G \G o F o G\ G \G o \'.map\ \\'.map o G\ .. show ?thesis proof show "\'GoG\'.map = G" proof (intro NaturalTransformation.eqI) show "natural_transformation C D G G G" using G.as_nat_trans.natural_transformation_axioms by auto show "natural_transformation C D G G \'GoG\'.map" using \'GoG\'.natural_transformation_axioms by auto show "\a. C.ide a \ \'GoG\'.map a = G a" proof - fix a assume a: "C.ide a" show "\'GoG\'.map a = G a" - using a \'GoG\'.map_simp_ide triangle_G' - \.components_are_iso \.components_are_iso G.preserves_ide + using a \'GoG\'.map_simp_ide triangle_G' G.preserves_ide \'.inverts_components \'.inverts_components D.inverse_unique G.preserves_inverse_arrows G\o\G.map_simp_ide D.inverse_arrows_sym triangle_G by (metis o_apply) qed qed show "F\'o\'F.map = F" proof (intro NaturalTransformation.eqI) show "natural_transformation D C F F F" using F.as_nat_trans.natural_transformation_axioms by auto show "natural_transformation D C F F F\'o\'F.map" using F\'o\'F.natural_transformation_axioms by auto show "\b. D.ide b \ F\'o\'F.map b = F b" proof - fix b assume b: "D.ide b" show "F\'o\'F.map b = F b" using b F\'o\'F.map_simp_ide \FoF\.map_simp_ide triangle_F triangle_F' - \.components_are_iso \.components_are_iso G.preserves_ide \'.inverts_components \'.inverts_components F.preserves_ide C.inverse_unique F.preserves_inverse_arrows C.inverse_arrows_sym by (metis o_apply) qed qed qed qed end text \ Every fully faithful and essentially surjective functor underlies an adjoint equivalence. To prove this without repeating things that were already proved in @{theory Category3.Adjunction}, we first show that a fully faithful and essentially surjective functor is a left adjoint functor, and then we show that if the left adjoint in a unit-counit adjunction is fully faithful and essentially surjective, then the unit and counit are natural isomorphisms; hence the adjunction is in fact an adjoint equivalence. \ locale fully_faithful_and_essentially_surjective_functor = C: category C + D: category D + fully_faithful_functor C D F + essentially_surjective_functor C D F for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\\<^sub>D" 55) and F :: "'c \ 'd" begin notation C.in_hom ("\_ : _ \\<^sub>C _\") notation D.in_hom ("\_ : _ \\<^sub>D _\") lemma is_left_adjoint_functor: shows "left_adjoint_functor C D F" proof fix y assume y: "D.ide y" let ?x = "SOME x. C.ide x \ (\e. D.iso e \ \e : F x \\<^sub>D y\)" let ?e = "SOME e. D.iso e \ \e : F ?x \\<^sub>D y\" have "\x e. D.iso e \ terminal_arrow_from_functor C D F x y e" proof - have "\x. D.iso ?e \ terminal_arrow_from_functor C D F x y ?e" proof - have x: "C.ide ?x \ (\e. D.iso e \ \e : F ?x \\<^sub>D y\)" using y essentially_surjective someI_ex [of "\x. C.ide x \ (\e. D.iso e \ \e : F x \\<^sub>D y\)"] by blast hence e: "D.iso ?e \ \?e : F ?x \\<^sub>D y\" using someI_ex [of "\e. D.iso e \ \e : F ?x \\<^sub>D y\"] by blast interpret arrow_from_functor C D F ?x y ?e using x e by (unfold_locales, simp) interpret terminal_arrow_from_functor C D F ?x y ?e proof fix x' f assume 1: "arrow_from_functor C D F x' y f" interpret f: arrow_from_functor C D F x' y f using 1 by simp have f: "\f: F x' \\<^sub>D y\" by (meson f.arrow) show "\!g. is_coext x' f g" proof let ?g = "SOME g. \g : x' \\<^sub>C ?x\ \ F g = D.inv ?e \\<^sub>D f" have g: "\?g : x' \\<^sub>C ?x\ \ F ?g = D.inv ?e \\<^sub>D f" using f e x f.arrow is_full D.comp_in_homI D.inv_in_hom someI_ex [of "\g. \g : x' \\<^sub>C ?x\ \ F g = D.inv ?e \\<^sub>D f"] by auto show 1: "is_coext x' f ?g" proof - have "\?g : x' \\<^sub>C ?x\" using g by simp moreover have "?e \\<^sub>D F ?g = f" proof - have "?e \\<^sub>D F ?g = ?e \\<^sub>D D.inv ?e \\<^sub>D f" using g by simp also have "... = (?e \\<^sub>D D.inv ?e) \\<^sub>D f" using e f D.inv_in_hom by (metis D.comp_assoc) also have "... = f" proof - have "?e \\<^sub>D D.inv ?e = y" using e D.comp_arr_inv D.inv_is_inverse by auto thus ?thesis using f D.comp_cod_arr by auto qed finally show ?thesis by blast qed ultimately show ?thesis unfolding is_coext_def by simp qed show "\g'. is_coext x' f g' \ g' = ?g" proof - fix g' assume g': "is_coext x' f g'" have 2: "\g' : x' \\<^sub>C ?x\ \ ?e \\<^sub>D F g' = f" using g' is_coext_def by simp have 3: "\?g : x' \\<^sub>C ?x\ \ ?e \\<^sub>D F ?g = f" using 1 is_coext_def by simp have "F g' = F ?g" using e f 2 3 D.iso_is_section D.section_is_mono D.monoE D.arrI by (metis (no_types, lifting) D.arrI) moreover have "C.par g' ?g" using 2 3 by fastforce ultimately show "g' = ?g" using is_faithful [of g' ?g] by simp qed qed qed show ?thesis using e terminal_arrow_from_functor_axioms by auto qed thus ?thesis by auto qed thus "\x e. terminal_arrow_from_functor C D F x y e" by blast qed lemma extends_to_adjoint_equivalence: shows "\G \ \. adjoint_equivalence C D G F \ \" proof - interpret left_adjoint_functor C D F using is_left_adjoint_functor by blast interpret Adj: meta_adjunction D C F G \ \ using induces_meta_adjunction by simp interpret Adj: adjunction D C replete_setcat.comp Adj.\C Adj.\D F G \ \ Adj.\ Adj.\ Adj.\ Adj.\ using induces_adjunction by simp interpret equivalence_of_categories D C F G Adj.\ Adj.\ proof show 1: "\a. D.ide a \ D.iso (Adj.\ a)" proof - fix a assume a: "D.ide a" interpret \a: terminal_arrow_from_functor C D F \G a\ a \Adj.\ a\ using a Adj.has_terminal_arrows_from_functor [of a] by blast have "D.retraction (Adj.\ a)" proof - obtain b \ where \: "C.ide b \ D.iso \ \ \\: F b \\<^sub>D a\" using a essentially_surjective by blast interpret \: arrow_from_functor C D F b a \ using \ by (unfold_locales, simp) let ?g = "\a.the_coext b \" have 1: "\?g : b \\<^sub>C G a\ \ Adj.\ a \\<^sub>D F ?g = \" using \.arrow_from_functor_axioms \a.the_coext_prop [of b \] by simp have "a = (Adj.\ a \\<^sub>D F ?g) \\<^sub>D D.inv \" using a 1 \ D.comp_cod_arr Adj.\.preserves_hom D.invert_side_of_triangle(2) by auto also have "... = Adj.\ a \\<^sub>D F ?g \\<^sub>D D.inv \" using a 1 \ D.inv_in_hom Adj.\.preserves_hom [of a a a] D.comp_assoc by blast finally have "\f. D.ide (Adj.\ a \\<^sub>D f)" using a by metis thus ?thesis unfolding D.retraction_def by blast qed moreover have "D.mono (Adj.\ a)" proof show "D.arr (Adj.\ a)" using a by simp show "\f f'. D.seq (Adj.\ a) f \ D.seq (Adj.\ a) f' \ Adj.\ a \\<^sub>D f = Adj.\ a \\<^sub>D f' \ f = f'" proof - fix f f' assume ff': "D.seq (Adj.\ a) f \ D.seq (Adj.\ a) f' \ Adj.\ a \\<^sub>D f = Adj.\ a \\<^sub>D f'" have f: "\f : D.dom f \\<^sub>D F (G a)\" using a ff' Adj.\.preserves_hom [of a a a] by fastforce have f': "\f' : D.dom f' \\<^sub>D F (G a)\" using a ff' Adj.\.preserves_hom [of a a a] by fastforce have par: "D.par f f'" using f f' ff' D.dom_comp [of "Adj.\ a" f] by force obtain b' \ where \: "C.ide b' \ D.iso \ \ \\: F b' \\<^sub>D D.dom f\" using par essentially_surjective D.ide_dom [of f] by blast have 1: "Adj.\ a \\<^sub>D f \\<^sub>D \ = Adj.\ a \\<^sub>D f' \\<^sub>D \" using ff' \ par D.comp_assoc by metis obtain g where g: "\g : b' \\<^sub>C G a\ \ F g = f \\<^sub>D \" using a f \ is_full [of "G a" b' "f \\<^sub>D \"] by auto obtain g' where g': "\g' : b' \\<^sub>C G a\ \ F g' = f' \\<^sub>D \" using a f' par \ is_full [of "G a" b' "f' \\<^sub>D \"] by auto interpret f\: arrow_from_functor C D F b' a \Adj.\ a \\<^sub>D f \\<^sub>D \\ using a \ f Adj.\.preserves_hom by (unfold_locales, fastforce) interpret f'\: arrow_from_functor C D F b' a \Adj.\ a \\<^sub>D f' \\<^sub>D \\ using a \ f' par Adj.\.preserves_hom by (unfold_locales, fastforce) have "\a.is_coext b' (Adj.\ a \\<^sub>D f \\<^sub>D \) g" unfolding \a.is_coext_def using g 1 by auto moreover have "\a.is_coext b' (Adj.\ a \\<^sub>D f' \\<^sub>D \) g'" unfolding \a.is_coext_def using g' 1 by auto ultimately have "g = g'" using 1 f\.arrow_from_functor_axioms f'\.arrow_from_functor_axioms \a.the_coext_unique \a.the_coext_unique [of b' "Adj.\ a \\<^sub>D f' \\<^sub>D \" g'] by auto hence "f \\<^sub>D \ = f' \\<^sub>D \" using g g' is_faithful by argo thus "f = f'" using \ f f' par D.iso_is_retraction D.retraction_is_epi D.epiE [of \ f f'] by auto qed qed ultimately show "D.iso (Adj.\ a)" using D.iso_iff_mono_and_retraction by simp qed interpret \: natural_isomorphism D D \F o G\ D.map Adj.\ using 1 by (unfold_locales, auto) interpret \F: natural_isomorphism C D \F o G o F\ F \Adj.\ o F\ using \.components_are_iso by (unfold_locales, simp) show "\a. C.ide a \ C.iso (Adj.\ a)" proof - fix a assume a: "C.ide a" have "D.inverse_arrows ((Adj.\ o F) a) ((F o Adj.\) a)" using a \.components_are_iso Adj.\\.triangle_F Adj.\FoF\.map_simp_ide D.section_retraction_of_iso by simp hence "D.iso ((F o Adj.\) a)" by blast thus "C.iso (Adj.\ a)" using a reflects_iso [of "Adj.\ a"] by fastforce qed qed (* * Uggh, I should have started with "right_adjoint_functor D C G" so that the * following would come out right. Instead, another step is needed to dualize. * TODO: Maybe re-work this later. *) interpret adjoint_equivalence D C F G Adj.\ Adj.\ .. interpret \': inverse_transformation D D \F o G\ D.map Adj.\ .. interpret \': inverse_transformation C C C.map \G o F\ Adj.\ .. interpret E: adjoint_equivalence C D G F \'.map \'.map using adjoint_equivalence_axioms dual_equivalence by blast show ?thesis using E.adjoint_equivalence_axioms by auto qed lemma is_right_adjoint_functor: shows "right_adjoint_functor C D F" proof - obtain G \ \ where E: "adjoint_equivalence C D G F \ \" using extends_to_adjoint_equivalence by auto interpret E: adjoint_equivalence C D G F \ \ using E by simp interpret Adj: meta_adjunction C D G F E.\ E.\ using E.induces_meta_adjunction by simp show ?thesis using Adj.has_right_adjoint_functor by simp qed lemma is_equivalence_functor: shows "equivalence_functor C D F" proof obtain G \ \ where E: "adjoint_equivalence C D G F \ \" using extends_to_adjoint_equivalence by auto interpret E: adjoint_equivalence C D G F \ \ using E by simp have "equivalence_of_categories C D G F \ \" .. thus "\G \ \. equivalence_of_categories C D G F \ \" by blast qed sublocale equivalence_functor C D F using is_equivalence_functor by blast end context equivalence_of_categories begin text \ The following development shows that an equivalence of categories can be refined to an adjoint equivalence by replacing just the counit. \ abbreviation \' where "\' a \ \ a \\<^sub>C F (D.inv (\ (G a))) \\<^sub>C C.inv (\ (F (G a)))" interpretation \': transformation_by_components C C \F \ G\ C.map \' proof show "\a. C.ide a \ \\' a : (F \ G) a \\<^sub>C C.map a\" using \.components_are_iso \.components_are_iso by simp fix f assume f: "C.arr f" show "\' (C.cod f) \\<^sub>C (F \ G) f = C.map f \\<^sub>C \' (C.dom f)" proof - have "\' (C.cod f) \\<^sub>C (F \ G) f = \ (C.cod f) \\<^sub>C F (D.inv (\ (G (C.cod f)))) \\<^sub>C C.inv (\ (F (G (C.cod f)))) \\<^sub>C F (G f)" using f C.comp_assoc by simp also have "... = \ (C.cod f) \\<^sub>C (F (D.inv (\ (G (C.cod f)))) \\<^sub>C F (G (F (G f)))) \\<^sub>C C.inv (\ (F (G (C.dom f))))" using f \.inv_naturality [of "F (G f)"] C.comp_assoc by simp also have "... = (\ (C.cod f) \\<^sub>C F (G f)) \\<^sub>C F (D.inv (\ (G (C.dom f)))) \\<^sub>C C.inv (\ (F (G (C.dom f))))" proof - have "F (D.inv (\ (G (C.cod f)))) \\<^sub>C F (G (F (G f))) = F (G f) \\<^sub>C F (D.inv (\ (G (C.dom f))))" proof - have "F (D.inv (\ (G (C.cod f)))) \\<^sub>C F (G (F (G f))) = F (D.inv (\ (G (C.cod f))) \\<^sub>D G (F (G f)))" using f by simp also have "... = F (G f \\<^sub>D D.inv (\ (G (C.dom f))))" using f \.inv_naturality [of "G f"] by simp also have "... = F (G f) \\<^sub>C F (D.inv (\ (G (C.dom f))))" using f by simp finally show ?thesis by blast qed thus ?thesis using C.comp_assoc by simp qed also have "... = C.map f \\<^sub>C \ (C.dom f) \\<^sub>C F (D.inv (\ (G (C.dom f)))) \\<^sub>C C.inv (\ (F (G (C.dom f))))" using f \.naturality C.comp_assoc by simp finally show ?thesis by blast qed qed interpretation \': natural_isomorphism C C \F \ G\ C.map \'.map proof show "\a. C.ide a \ C.iso (\'.map a)" unfolding \'.map_def using \.components_are_iso \.components_are_iso apply simp by (intro C.isos_compose) auto qed lemma F\_inverse: assumes "D.ide b" shows "F (\ (G (F b))) = F (G (F (\ b)))" and "F (\ b) \\<^sub>C \ (F b) = \ (F (G (F b))) \\<^sub>C F (\ (G (F b)))" and "C.inverse_arrows (F (\ b)) (\' (F b))" and "F (\ b) = C.inv (\' (F b))" and "C.inv (F (\ b)) = \' (F b)" proof - let ?\' = "\a. \ a \\<^sub>C F (D.inv (\ (G a))) \\<^sub>C C.inv (\ (F (G a)))" show 1: "F (\ (G (F b))) = F (G (F (\ b)))" proof - have "F (\ (G (F b))) \\<^sub>C F (\ b) = F (G (F (\ b))) \\<^sub>C F (\ b)" proof - have "F (\ (G (F b))) \\<^sub>C F (\ b) = F (\ (G (F b)) \\<^sub>D \ b)" using assms by simp also have "... = F (G (F (\ b)) \\<^sub>D \ b)" using assms \.naturality [of "\ b"] by simp also have "... = F (G (F (\ b))) \\<^sub>C F (\ b)" using assms by simp finally show ?thesis by blast qed thus ?thesis using assms \.components_are_iso C.iso_cancel_right by simp qed show "F (\ b) \\<^sub>C \ (F b) = \ (F (G (F b))) \\<^sub>C F (\ (G (F b)))" using assms 1 \.naturality [of "F (\ b)"] by simp show 2: "C.inverse_arrows (F (\ b)) (?\' (F b))" proof show 3: "C.ide (?\' (F b) \\<^sub>C F (\ b))" proof - have "?\' (F b) \\<^sub>C F (\ b) = \ (F b) \\<^sub>C (F (D.inv (\ (G (F b)))) \\<^sub>C C.inv (\ (F (G (F b))))) \\<^sub>C F (\ b)" using C.comp_assoc by simp also have "... = \ (F b) \\<^sub>C (F (D.inv (\ (G (F b)))) \\<^sub>C F (G (F (\ b)))) \\<^sub>C C.inv (\ (F b))" using assms \.naturality [of "F (\ b)"] \.components_are_iso C.comp_assoc C.invert_opposite_sides_of_square [of "\ (F (G (F b)))" "F (G (F (\ b)))" "F (\ b)" "\ (F b)"] by simp also have "... = \ (F b) \\<^sub>C C.inv (\ (F b))" proof - have "F (D.inv (\ (G (F b)))) \\<^sub>C F (G (F (\ b))) = F (G (F b))" using assms 1 D.comp_inv_arr' \.components_are_iso by (metis D.ideD(1) D.ideD(2) F.preserves_comp F.preserves_ide G.preserves_ide \.preserves_dom D.map_simp) moreover have "F (G (F b)) \\<^sub>C C.inv (\ (F b)) = C.inv (\ (F b))" using assms D.comp_cod_arr \.components_are_iso C.inv_in_hom [of "\ (F b)"] by (metis C.comp_ide_arr C_arr_expansion(1) D.ide_char F.preserves_arr F.preserves_dom F.preserves_ide G.preserves_ide C.seqE) ultimately show ?thesis by simp qed also have "... = F b" using assms \.components_are_iso C.comp_arr_inv' by simp finally have "(\ (F b) \\<^sub>C F (D.inv (\ (G (F b)))) \\<^sub>C C.inv (\ (F (G (F b))))) \\<^sub>C F (\ b) = F b" by blast thus ?thesis using assms by simp qed show "C.ide (F (\ b) \\<^sub>C ?\' (F b))" proof - have "(F (\ b) \\<^sub>C ?\' (F b)) \\<^sub>C F (\ b) = F (G (F b)) \\<^sub>C F (\ b)" proof - have "(F (\ b) \\<^sub>C ?\' (F b)) \\<^sub>C F (\ b) = F (\ b) \\<^sub>C (\ (F b) \\<^sub>C F (D.inv (\ (G (F b)))) \\<^sub>C C.inv (\ (F (G (F b))))) \\<^sub>C F (\ b)" using C.comp_assoc by simp also have "... = F (\ b)" using assms 3 C.comp_arr_dom [of "F (\ b)" "(\ (F b) \\<^sub>C F (D.inv (\ (G (F b)))) \\<^sub>C C.inv (\ (F (G (F b))))) \\<^sub>C F (\ b)"] by auto also have "... = F (G (F b)) \\<^sub>C F (\ b)" using assms C.comp_cod_arr by simp finally show ?thesis by blast qed hence "F (\ b) \\<^sub>C ?\' (F b) = F (G (F b))" using assms C.iso_cancel_right by simp thus ?thesis using assms by simp qed qed show "C.inv (F (\ b)) = ?\' (F b)" using assms 2 C.inverse_unique by simp show "F (\ b) = C.inv (?\' (F b))" proof - have "C.inverse_arrows (?\' (F b)) (F (\ b))" using assms 2 by auto thus ?thesis using assms C.inverse_unique by simp qed qed interpretation FoGoF: composite_functor D C C F \F o G\ .. interpretation GoFoG: composite_functor C D D G \G o F\ .. interpretation natural_transformation D C F FoGoF.map \F \ \\ proof - have "F \ D.map = F" using hcomp_ide_dom F.as_nat_trans.natural_transformation_axioms by blast moreover have "F o (G o F) = FoGoF.map" by auto ultimately show "natural_transformation D C F FoGoF.map (F \ \)" using \.natural_transformation_axioms F.as_nat_trans.natural_transformation_axioms horizontal_composite [of D D D.map "G o F" \ C F F F] by simp qed interpretation natural_transformation C D G GoFoG.map \\ \ G\ using \.natural_transformation_axioms G.as_nat_trans.natural_transformation_axioms horizontal_composite [of C D G G G ] by fastforce interpretation natural_transformation D C FoGoF.map F \\'.map \ F\ using \'.natural_transformation_axioms F.as_nat_trans.natural_transformation_axioms horizontal_composite [of D C F F F] by fastforce interpretation natural_transformation C D GoFoG.map G \G \ \'.map\ proof - have "G o C.map = G" using hcomp_ide_dom G.as_nat_trans.natural_transformation_axioms by blast moreover have "G o (F o G) = GoFoG.map" by auto ultimately show "natural_transformation C D GoFoG.map G (G \ \'.map)" using G.as_nat_trans.natural_transformation_axioms \'.natural_transformation_axioms horizontal_composite [of C C "F o G" C.map \'.map D G G G] by simp qed interpretation \'F_F\: vertical_composite D C F FoGoF.map F \F \ \\ \\'.map \ F\ .. interpretation G\'_\G: vertical_composite C D G GoFoG.map G \\ o G\ \G o \'.map\ .. interpretation \\': unit_counit_adjunction C D F G \ \'.map proof show 1: "\'F_F\.map = F" proof fix g show "\'F_F\.map g = F g" proof (cases "D.arr g") show "\ D.arr g \ \'F_F\.map g = F g" using \'F_F\.is_extensional F.is_extensional by simp assume g: "D.arr g" have "\'F_F\.map g = \' (F (D.cod g)) \\<^sub>C F (\ g)" using g \'F_F\.map_def by simp also have "... = \' (F (D.cod g)) \\<^sub>C F (\ (D.cod g) \\<^sub>D g)" using g \.is_natural_2 by simp also have "... = (\' (F (D.cod g)) \\<^sub>C F (\ (D.cod g))) \\<^sub>C F g" using g C.comp_assoc by simp also have "... = F (D.cod g) \\<^sub>C F g" using g F\_inverse(3) [of "D.cod g"] by fastforce also have "... = F g" using g C.comp_cod_arr by simp finally show "\'F_F\.map g = F g" by blast qed qed show "G\'_\G.map = G" proof fix f show "G\'_\G.map f = G f" proof (cases "C.arr f") show "\ C.arr f \ G\'_\G.map f = G f" using G\'_\G.is_extensional G.is_extensional by simp assume f: "C.arr f" have "F (G\'_\G.map f) = F (G (\' (C.cod f)) \\<^sub>D \ (G f))" using f G\'_\G.map_def D.comp_assoc by simp also have "... = F (G (\' (C.cod f)) \\<^sub>D \ (G (C.cod f)) \\<^sub>D G f)" using f \.is_natural_2 [of "G f"] by simp also have "... = F (G (\' (C.cod f))) \\<^sub>C F (\ (G (C.cod f))) \\<^sub>C F (G f)" using f by simp also have "... = (F (G (\' (C.cod f))) \\<^sub>C C.inv (\' (F (G (C.cod f))))) \\<^sub>C F (G f)" using f F\_inverse(4) C.comp_assoc by simp also have "... = (C.inv (\' (C.cod f)) \\<^sub>C \' (C.cod f)) \\<^sub>C F (G f)" using f \'.inv_naturality [of "\' (C.cod f)"] by simp also have "... = F (G (C.cod f)) \\<^sub>C F (G f)" using f C.comp_inv_arr' [of "\' (C.cod f)"] \'.components_are_iso by simp also have "... = F (G f)" using f C.comp_cod_arr by simp finally have "F (G\'_\G.map f) = F (G f)" by blast moreover have "D.par (G\'_\G.map f) (G f)" using f by simp ultimately show "G\'_\G.map f = G f" using f F_is_faithful by (simp add: faithful_functor_axioms_def faithful_functor_def) qed qed qed interpretation \\': adjoint_equivalence C D F G \ \'.map .. lemma refines_to_adjoint_equivalence: shows "adjoint_equivalence C D F G \ \'.map" .. end end diff --git a/thys/Category3/SetCat.thy b/thys/Category3/SetCat.thy --- a/thys/Category3/SetCat.thy +++ b/thys/Category3/SetCat.thy @@ -1,875 +1,874 @@ (* Title: SetCat Author: Eugene W. Stark , 2016 Maintainer: Eugene W. Stark *) chapter SetCat theory SetCat imports SetCategory ConcreteCategory begin text\ This theory proves the consistency of the \set_category\ locale by giving a particular concrete construction of an interpretation for it. Applying the general construction given by @{locale concrete_category}, we define arrows to be terms \MkArr A B F\, where \A\ and \B\ are sets and \F\ is an extensional function that maps \A\ to \B\. \ text\ This locale uses an extra dummy parameter just to fix the element type for sets. Without this, a type is used for each interpretation, which makes it impossible to construct set categories whose element types are related to the context. \ locale setcat = fixes dummy :: 'e and \ :: "'a rel" assumes cardinal: "Card_order \ \ infinite (Field \)" begin lemma finite_imp_card_less: assumes "finite A" shows "|A| " proof - have "finite (Field |A| )" using assms by simp thus ?thesis using cardinal card_of_Well_order card_order_on_def finite_ordLess_infinite by blast qed type_synonym 'b arr = "('b set, 'b \ 'b) concrete_category.arr" interpretation concrete_category \{A :: 'e set. |A| }\ \\A B. extensional A \ (A \ B)\ \\A. \x \ A. x\ \\C B A g f. compose A g f\ using compose_Id Id_compose apply unfold_locales apply auto[3] apply blast by (metis IntD2 compose_assoc) abbreviation comp :: "'e setcat.arr comp" (infixr "\" 55) where "comp \ COMP" notation in_hom ("\_ : _ \ _\") lemma MkArr_expansion: assumes "arr f" shows "f = MkArr (Dom f) (Cod f) (\x \ Dom f. Map f x)" proof (intro arr_eqI) show "arr f" by fact show "arr (MkArr (Dom f) (Cod f) (\x \ Dom f. Map f x))" using assms arr_char by (metis (mono_tags, lifting) Int_iff MkArr_Map extensional_restrict) show "Dom f = Dom (MkArr (Dom f) (Cod f) (\x \ Dom f. Map f x))" by simp show "Cod f = Cod (MkArr (Dom f) (Cod f) (\x \ Dom f. Map f x))" by simp show "Map f = Map (MkArr (Dom f) (Cod f) (\x \ Dom f. Map f x))" using assms arr_char by (metis (mono_tags, lifting) Int_iff MkArr_Map extensional_restrict) qed lemma arr_char: shows "arr f \ f \ Null \ |Dom f| \ |Cod f| \ Map f \ extensional (Dom f) \ (Dom f \ Cod f)" using arr_char by auto lemma terminal_char: shows "terminal a \ (\x. a = MkIde {x})" proof show "\x. a = MkIde {x} \ terminal a" proof - assume a: "\x. a = MkIde {x}" from this obtain x where x: "a = MkIde {x}" by blast have "terminal (MkIde {x})" proof show 1: "ide (MkIde {x})" using finite_imp_card_less ide_MkIde by auto show "\a. ide a \ \!f. \f : a \ MkIde {x}\" proof fix a :: "'e setcat.arr" assume a: "ide a" show "\MkArr (Dom a) {x} (\_\Dom a. x) : a \ MkIde {x}\" proof show 2: "arr (MkArr (Dom a) {x} (\_ \ Dom a. x))" using a 1 arr_MkArr [of "Dom a" "{x}"] ide_char by force show "dom (MkArr (Dom a) {x} (\_ \ Dom a. x)) = a" using a 2 dom_MkArr by force show "cod (MkArr (Dom a) {x} (\_\Dom a. x)) = MkIde {x}" using 2 cod_MkArr by blast qed fix f :: "'e setcat.arr" assume f: "\f : a \ MkIde {x}\" show "f = MkArr (Dom a) {x} (\_ \ Dom a. x)" proof - have 1: "Dom f = Dom a \ Cod f = {x}" using a f by (metis (mono_tags, lifting) Dom.simps(1) in_hom_char) moreover have "Map f = (\_ \ Dom a. x)" proof fix z have "z \ Dom a \ Map f z = (\_ \ Dom a. x) z" using f 1 MkArr_expansion by (metis (mono_tags, lifting) Map.simps(1) in_homE restrict_apply) moreover have "z \ Dom a \ Map f z = (\_ \ Dom a. x) z" using f 1 arr_char [of f] by fastforce ultimately show "Map f z = (\_ \ Dom a. x) z" by auto qed ultimately show ?thesis using f MkArr_expansion [of f] by fastforce qed qed qed thus "terminal a" using x by simp qed show "terminal a \ \x. a = MkIde {x}" proof - assume a: "terminal a" hence ide_a: "ide a" using terminal_def by auto have 1: "\!x. x \ Dom a" proof - have "Dom a = {} \ \terminal a" proof - assume "Dom a = {}" hence 1: "a = MkIde {}" using MkIde_Dom' \ide a\ by fastforce have "\f. f \ hom (MkIde {undefined}) (MkIde ({} :: 'e set)) \ Map f \ {undefined} \ {}" proof - fix f assume f: "f \ hom (MkIde {undefined}) (MkIde ({} :: 'e set))" show "Map f \ {undefined} \ {}" using f MkArr_expansion arr_char [of f] in_hom_char by auto qed hence "hom (MkIde {undefined}) a = {}" using 1 by auto moreover have "ide (MkIde {undefined})" using finite_imp_card_less by (metis (mono_tags, lifting) finite.intros(1-2) ide_MkIde mem_Collect_eq) ultimately show "\terminal a" by blast qed moreover have "\x x'. x \ Dom a \ x' \ Dom a \ x \ x' \ \terminal a" proof - fix x x' assume 1: "x \ Dom a \ x' \ Dom a \ x \ x'" let ?f = "MkArr {undefined} (Dom a) (\_ \ {undefined}. x)" let ?f' = "MkArr {undefined} (Dom a) (\_ \ {undefined}. x')" have "\?f : MkIde {undefined} \ a\" proof show 2: "arr ?f" proof (intro arr_MkArr) show "{undefined} \ {A. |A| }" by (simp add: finite_imp_card_less) show "Dom a \ {A. |A| }" using ide_a ide_char by blast show "(\_ \ {undefined}. x) \ extensional {undefined} \ ({undefined} \ Dom a)" using 1 by blast qed show "dom ?f = MkIde {undefined}" using 2 dom_MkArr by auto show "cod ?f = a" using 2 cod_MkArr ide_a by force qed moreover have "\?f' : MkIde {undefined} \ a\" proof show 2: "arr ?f'" using 1 ide_a ide_char arr_MkArr [of "{undefined}" "Dom a"] finite_imp_card_less proof (intro arr_MkArr) show "{undefined} \ {A. |A| }" by (simp add: finite_imp_card_less) show "Dom a \ {A. |A| }" using ide_a ide_char by blast show "(\_ \ {undefined}. x') \ extensional {undefined} \ ({undefined} \ Dom a)" using 1 by blast qed show "dom ?f' = MkIde {undefined}" using 2 dom_MkArr by auto show "cod ?f' = a" using 2 cod_MkArr ide_a by force qed moreover have "?f \ ?f'" using 1 by (metis arr.inject restrict_apply' singletonI) ultimately show "\terminal a" using terminal_arr_unique by (metis (mono_tags, lifting) in_homE) qed ultimately show ?thesis using a by auto qed hence "Dom a = {THE x. x \ Dom a}" using theI [of "\x. x \ Dom a"] by auto hence "a = MkIde {THE x. x \ Dom a}" using a terminal_def by (metis (mono_tags, lifting) MkIde_Dom') thus "\x. a = MkIde {x}" by auto qed qed definition IMG :: "'e setcat.arr \ 'e setcat.arr" where "IMG f = MkIde (Map f ` Dom f)" interpretation set_category_data comp IMG .. lemma terminal_unity: shows "terminal unity" using terminal_char unity_def someI_ex [of terminal] by (metis (mono_tags, lifting)) text\ The inverse maps @{term UP} and @{term DOWN} are used to pass back and forth between the inhabitants of type @{typ 'a} and the corresponding terminal objects. These are exported so that a client of the theory can relate the concrete element type @{typ 'a} to the otherwise abstract arrow type. \ definition UP :: "'e \ 'e setcat.arr" where "UP x \ MkIde {x}" definition DOWN :: "'e setcat.arr \ 'e" where "DOWN t \ the_elem (Dom t)" abbreviation U where "U \ DOWN unity" lemma UP_mapsto: shows "UP \ UNIV \ Univ" using terminal_char UP_def by fast lemma DOWN_mapsto: shows "DOWN \ Univ \ UNIV" by auto lemma DOWN_UP [simp]: shows "DOWN (UP x) = x" by (simp add: DOWN_def UP_def) lemma UP_DOWN [simp]: assumes "t \ Univ" shows "UP (DOWN t) = t" using assms terminal_char UP_def DOWN_def by (metis (mono_tags, lifting) mem_Collect_eq DOWN_UP) lemma inj_UP: shows "inj UP" by (metis DOWN_UP injI) lemma bij_UP: shows "bij_betw UP UNIV Univ" proof (intro bij_betwI) interpret category comp using is_category by auto show DOWN_UP: "\x :: 'e. DOWN (UP x) = x" by simp show UP_DOWN: "\t. t \ Univ \ UP (DOWN t) = t" by simp show "UP \ UNIV \ Univ" using UP_mapsto by auto show "DOWN \ Collect terminal \ UNIV" by auto qed lemma Dom_terminal: assumes "terminal t" shows "Dom t = {DOWN t}" using assms UP_def by (metis (mono_tags, lifting) Dom.simps(1) DOWN_def terminal_char the_elem_eq) text\ The image of a point @{term "p \ hom unity a"} is a terminal object, which is given by the formula @{term "(UP o Fun p o DOWN) unity"}. \ lemma IMG_point: assumes "\p : unity \ a\" shows "IMG \ hom unity a \ Univ" and "IMG p = (UP o Map p o DOWN) unity" proof - show "IMG \ hom unity a \ Univ" proof fix f assume f: "f \ hom unity a" have "terminal (MkIde (Map f ` Dom unity))" proof - obtain u :: 'e where u: "unity = MkIde {u}" using terminal_unity terminal_char by (metis (mono_tags, lifting)) have "Map f ` Dom unity = {Map f u}" using u by simp thus ?thesis using terminal_char by auto qed hence "MkIde (Map f ` Dom unity) \ Univ" by simp moreover have "MkIde (Map f ` Dom unity) = IMG f" using f IMG_def in_hom_char by (metis (mono_tags, lifting) mem_Collect_eq) ultimately show "IMG f \ Univ" by auto qed have "IMG p = MkIde (Map p ` Dom p)" using IMG_def by blast also have "... = MkIde (Map p ` {U})" using assms in_hom_char terminal_unity Dom_terminal by (metis (mono_tags, lifting)) also have "... = (UP o Map p o DOWN) unity" by (simp add: UP_def) finally show "IMG p = (UP o Map p o DOWN) unity" using assms by auto qed text\ The function @{term IMG} is injective on @{term "hom unity a"} and its inverse takes a terminal object @{term t} to the arrow in @{term "hom unity a"} corresponding to the constant-@{term t} function. \ abbreviation MkElem :: "'e setcat.arr => 'e setcat.arr => 'e setcat.arr" where "MkElem t a \ MkArr {U} (Dom a) (\_ \ {U}. DOWN t)" lemma MkElem_in_hom: assumes "arr f" and "x \ Dom f" shows "\MkElem (UP x) (dom f) : unity \ dom f\" proof - have "(\_ \ {U}. DOWN (UP x)) \ {U} \ Dom (dom f)" using assms dom_char [of f] by simp moreover have "MkIde {U} = unity" using terminal_char terminal_unity by (metis (mono_tags, lifting) DOWN_UP UP_def) moreover have "MkIde (Dom (dom f)) = dom f" using assms dom_char MkIde_Dom' ide_dom by blast ultimately show ?thesis using assms MkArr_in_hom [of "{U}" "Dom (dom f)" "\_ \ {U}. DOWN (UP x)"] by (metis (no_types, lifting) Dom.simps(1) Dom_in_Obj IntI arr_dom ideD(1) restrict_extensional terminal_def terminal_unity) qed lemma MkElem_IMG: assumes "p \ hom unity a" shows "MkElem (IMG p) a = p" proof - have 0: "IMG p = UP (Map p U)" using assms IMG_point(2) by auto have 1: "Dom p = {U}" using assms terminal_unity Dom_terminal by (metis (mono_tags, lifting) in_hom_char mem_Collect_eq) moreover have "Cod p = Dom a" using assms by (metis (mono_tags, lifting) in_hom_char mem_Collect_eq) moreover have "Map p = (\_ \ {U}. DOWN (IMG p))" proof fix e show "Map p e = (\_ \ {U}. DOWN (IMG p)) e" proof - have "Map p e = (\x \ Dom p. Map p x) e" using assms MkArr_expansion [of p] by (metis (mono_tags, lifting) CollectD Map.simps(1) in_homE) also have "... = (\_ \ {U}. DOWN (IMG p)) e" using assms 0 1 by simp finally show ?thesis by blast qed qed ultimately show "MkElem (IMG p) a = p" using assms MkArr_Map CollectD by (metis (mono_tags, lifting) in_homE mem_Collect_eq) qed lemma inj_IMG: assumes "ide a" shows "inj_on IMG (hom unity a)" proof (intro inj_onI) fix x y assume x: "x \ hom unity a" assume y: "y \ hom unity a" assume eq: "IMG x = IMG y" show "x = y" proof (intro arr_eqI) show "arr x" using x by blast show "arr y" using y by blast show "Dom x = Dom y" using x y in_hom_char by (metis (mono_tags, lifting) CollectD) show "Cod x = Cod y" using x y in_hom_char by (metis (mono_tags, lifting) CollectD) show "Map x = Map y" proof - have "\a. y \ hom unity a \ MkArr {U} (Dom a) (\e\{U}. DOWN (IMG x)) = y" using MkElem_IMG eq by presburger hence "y = x" using MkElem_IMG x y by blast thus ?thesis by meson qed qed qed lemma set_char: assumes "ide a" shows "set a = UP ` Dom a" proof show "set a \ UP ` Dom a" proof fix t assume "t \ set a" from this obtain p where p: "\p : unity \ a\ \ t = IMG p" using set_def by blast have "t = (UP o Map p o DOWN) unity" using p IMG_point(2) by blast moreover have "(Map p o DOWN) unity \ Dom a" using p arr_char in_hom_char Dom_terminal terminal_unity by (metis (mono_tags, lifting) IntD2 Pi_split_insert_domain o_apply) ultimately show "t \ UP ` Dom a" by simp qed show "UP ` Dom a \ set a" proof fix t assume "t \ UP ` Dom a" from this obtain x where x: "x \ Dom a \ t = UP x" by blast let ?p = "MkElem (UP x) a" have p: "?p \ hom unity a" using assms x MkElem_in_hom [of "dom a"] ideD(1-2) by force moreover have "IMG ?p = t" using p x DOWN_UP IMG_def UP_def by (metis (no_types, lifting) Dom.simps(1) Map.simps(1) image_empty image_insert image_restrict_eq) ultimately show "t \ set a" using set_def by blast qed qed lemma Map_via_comp: assumes "arr f" shows "Map f = (\x \ Dom f. Map (f \ MkElem (UP x) (dom f)) U)" proof fix x have "x \ Dom f \ Map f x = (\x \ Dom f. Map (f \ MkElem (UP x) (dom f)) U) x" using assms arr_char [of f] IntD1 extensional_arb restrict_apply by fastforce moreover have "x \ Dom f \ Map f x = (\x \ Dom f. Map (f \ MkElem (UP x) (dom f)) U) x" proof - assume x: "x \ Dom f" let ?X = "MkElem (UP x) (dom f)" have "\?X : unity \ dom f\" using assms x MkElem_in_hom by auto moreover have "Dom ?X = {U} \ Map ?X = (\_ \ {U}. x)" using x by simp ultimately have "Map (f \ MkElem (UP x) (dom f)) = compose {U} (Map f) (\_ \ {U}. x)" using assms x Map_comp [of "MkElem (UP x) (dom f)" f] by (metis (mono_tags, lifting) Cod.simps(1) Dom_dom arr_iff_in_hom seqE seqI') thus ?thesis using x by (simp add: compose_eq restrict_apply' singletonI) qed ultimately show "Map f x = (\x \ Dom f. Map (f \ MkElem (UP x) (dom f)) U) x" by auto qed lemma arr_eqI': assumes "par f f'" and "\t. \t : unity \ dom f\ \ f \ t = f' \ t" shows "f = f'" proof (intro arr_eqI) show "arr f" using assms by simp show "arr f'" using assms by simp show "Dom f = Dom f'" using assms by (metis (mono_tags, lifting) Dom_dom) show "Cod f = Cod f'" using assms by (metis (mono_tags, lifting) Cod_cod) show "Map f = Map f'" proof have 1: "\x. x \ Dom f \ \MkElem (UP x) (dom f) : unity \ dom f\" using MkElem_in_hom by (metis (mono_tags, lifting) assms(1)) fix x show "Map f x = Map f' x" using assms 1 \Dom f = Dom f'\ by (simp add: Map_via_comp) qed qed text \ We need to show that the cardinality constraint on the sets that determine objects implies a corresponding constraint on the sets of global elements of those objects. \ lemma card_points_less: assumes "ide a" shows "|hom unity a| " proof - have "bij_betw (\f. Map f U) (hom unity a) (Dom a)" proof (intro bij_betwI') show "\x. x \ hom unity a \ Map x (DOWN unity) \ Dom a" using arr_char Dom_terminal terminal_unity in_hom_char by auto show "\x y. \x \ hom unity a; y \ hom unity a\ \ Map x U = Map y U \ x = y" proof - fix x y assume x: "x \ hom unity a" and y: "y \ hom unity a" have 1: "Map x \ extensional {U} \ Map y \ extensional {U}" using x y in_hom_char Dom_terminal terminal_unity by (metis (mono_tags, lifting) Map_via_comp mem_Collect_eq restrict_extensional) show "Map x U = Map y U \ x = y" proof show "x = y \ Map x U = Map y U" by simp show "Map x U = Map y U \ x = y" proof - assume 2: "Map x U = Map y U" have "Map x = Map y" proof fix z show "Map x z = Map y z" using 1 2 extensional_arb [of "Map x"] extensional_arb [of "Map y"] by (cases "z = U") auto qed thus "x = y" using x y 1 in_hom_char by (intro arr_eqI) auto qed qed qed show "\y. y \ Dom a \ \x \ hom unity a. y = Map x (DOWN unity)" proof - fix y assume y: "y \ Dom a" let ?x = "MkArr {DOWN unity} (Dom a) (\_ \ {U}. y)" have "arr ?x" proof (intro arr_MkArr) show "{U} \ {A. |A| }" by (metis (mono_tags, lifting) Dom_terminal ide_char terminal_def terminal_unity) show "Dom a \ {A. |A| }" using assms ide_char by blast show "(\_ \ {U}. y) \ extensional {U} \ ({U} \ Dom a)" using assms y by blast qed hence "?x \ hom unity a" using UP_DOWN UP_def assms cod_MkArr dom_char in_homI terminal_unity by simp moreover have "y = Map ?x (DOWN unity)" by simp ultimately show "\x \ hom unity a. y = Map x (DOWN unity)" by auto qed qed hence "|hom unity a| =o |Dom a|" using card_of_ordIsoI by auto moreover have "|Dom a| " using assms ide_char by auto ultimately show "|hom unity a| " using ordIso_ordLess_trans by auto qed text\ The main result, which establishes the consistency of the \set_category\ locale and provides us with a way of obtaining ``set categories'' at arbitrary types. \ theorem is_set_category: shows "set_category comp \" proof show "\img :: 'e setcat.arr \ 'e setcat.arr. set_category_given_img comp img \" proof show "set_category_given_img (comp :: 'e setcat.arr comp) IMG \" proof show "Card_order \ \ infinite (Field \ )" using cardinal by simp show "Univ \ {}" using terminal_char by blast fix a :: "'e setcat.arr" assume a: "ide a" show "IMG \ hom unity a \ Univ" using IMG_point terminal_unity by blast show "|hom unity a| " using a card_points_less by simp show "inj_on IMG (hom unity a)" using a inj_IMG terminal_unity by blast next fix t :: "'e setcat.arr" assume t: "terminal t" show "t \ IMG ` hom unity t" proof - have "t \ set t" using t set_char [of t] by (metis (mono_tags, lifting) Dom.simps(1) image_insert insertI1 UP_def terminal_char terminal_def) thus ?thesis using t set_def [of t] by simp qed next fix A :: "'e setcat.arr set" assume A: "A \ Univ" and 0: "|A| " show "\a. ide a \ set a = A" proof let ?a = "MkArr (DOWN ` A) (DOWN ` A) (\x \ (DOWN ` A). x)" show "ide ?a \ set ?a = A" proof have "|DOWN ` A| " using 0 card_of_image ordLeq_ordLess_trans by blast thus 1: "ide ?a" using ide_char [of ?a] by simp show "set ?a = A" proof - have 2: "\x. x \ A \ x = UP (DOWN x)" using A UP_DOWN by force hence "UP ` DOWN ` A = A" using A UP_DOWN by auto thus ?thesis using 1 A set_char [of ?a] by simp qed qed qed next fix a b :: "'e setcat.arr" assume a: "ide a" and b: "ide b" and ab: "set a = set b" show "a = b" using a b ab set_char inj_UP inj_image_eq_iff dom_char in_homE ide_in_hom by (metis (mono_tags, lifting)) next fix f f' :: "'e setcat.arr" assume par: "par f f'" and ff': "\x. \x : unity \ dom f\ \ f \ x = f' \ x" show "f = f'" using par ff' arr_eqI' by blast next fix a b :: "'e setcat.arr" and F :: "'e setcat.arr \ 'e setcat.arr" assume a: "ide a" and b: "ide b" and F: "F \ hom unity a \ hom unity b" show "\f. \f : a \ b\ \ (\x. \x : unity \ dom f\ \ f \ x = F x)" proof let ?f = "MkArr (Dom a) (Dom b) (\x \ Dom a. Map (F (MkElem (UP x) a)) U)" have 1: "\?f : a \ b\" proof - have "(\x \ Dom a. Map (F (MkElem (UP x) a)) U) \ extensional (Dom a) \ (Dom a \ Dom b)" proof show "(\x \ Dom a. Map (F (MkElem (UP x) a)) U) \ extensional (Dom a)" using a F by simp show "(\x \ Dom a. Map (F (MkElem (UP x) a)) U) \ Dom a \ Dom b" proof fix x assume x: "x \ Dom a" have "MkElem (UP x) a \ hom unity a" using x a MkElem_in_hom [of a x] ide_char ideD(1-2) by force hence 1: "F (MkElem (UP x) a) \ hom unity b" using F by auto moreover have "Dom (F (MkElem (UP x) a)) = {U}" using 1 MkElem_IMG by (metis (mono_tags, lifting) Dom.simps(1)) moreover have "Cod (F (MkElem (UP x) a)) = Dom b" using 1 by (metis (mono_tags, lifting) CollectD in_hom_char) ultimately have "Map (F (MkElem (UP x) a)) \ {U} \ Dom b" using arr_char [of "F (MkElem (UP x) a)"] by blast thus "Map (F (MkElem (UP x) a)) U \ Dom b" by blast qed qed hence "\?f : MkIde (Dom a) \ MkIde (Dom b)\" using a b MkArr_in_hom ide_char by blast thus ?thesis using a b by simp qed moreover have "\x. \x : unity \ dom ?f\ \ ?f \ x = F x" proof - fix x assume x: "\x : unity \ dom ?f\" have 2: "x = MkElem (IMG x) a" using a x 1 MkElem_IMG [of x a] by (metis (mono_tags, lifting) in_homE mem_Collect_eq) moreover have 5: "Dom x = {U} \ Cod x = Dom a \ Map x = (\_ \ {U}. DOWN (IMG x))" using x 2 by (metis (no_types, lifting) Cod.simps(1) Dom.simps(1) Map.simps(1)) moreover have "Cod ?f = Dom b" using 1 by simp ultimately have 3: "?f \ x = MkArr {U} (Dom b) (compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)))" - using 1 x comp_char [of ?f "MkElem (IMG x) a"] - by (metis (mono_tags, lifting) in_homE seqI) + by (metis (no_types, lifting) "1" Map.simps(1) comp_MkArr in_homE x) have 4: "compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) = Map (F x)" proof fix y have "y \ {U} \ compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = Map (F x) y" proof - assume y: "y \ {U}" have "compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = undefined" using y compose_def extensional_arb by simp also have "... = Map (F x) y" proof - have 5: "F x \ hom unity b" using x F 1 by fastforce hence "Dom (F x) = {U}" by (metis (mono_tags, lifting) "2" CollectD Dom.simps(1) in_hom_char x) thus ?thesis using x y F 5 arr_char [of "F x"] extensional_arb [of "Map (F x)" "{U}" y] by (metis (mono_tags, lifting) CollectD Int_iff in_hom_char) qed ultimately show ?thesis by auto qed moreover have "y \ {U} \ compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = Map (F x) y" proof - assume y: "y \ {U}" have "compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = Map ?f (DOWN (IMG x))" using y by (simp add: compose_eq restrict_apply') also have "... = (\x. Map (F (MkElem (UP x) a)) U) (DOWN (IMG x))" proof - have "DOWN (IMG x) \ Dom a" using x y a 5 arr_char in_homE restrict_apply by force thus ?thesis using restrict_apply by simp qed also have "... = Map (F x) y" using x y 1 2 MkElem_IMG [of x a] by simp finally show "compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = Map (F x) y" by auto qed ultimately show "compose {U} (Map ?f) (\_ \ {U}. DOWN (IMG x)) y = Map (F x) y" by auto qed show "?f \ x = F x" proof (intro arr_eqI) have 5: "?f \ x \ hom unity b" using 1 x by blast have 6: "F x \ hom unity b" using x F 1 by (metis (mono_tags, lifting) PiE in_homE mem_Collect_eq) show "arr (comp ?f x)" using 5 by blast show "arr (F x)" using 6 by blast show "Dom (comp ?f x) = Dom (F x)" using 5 6 by (metis (mono_tags, lifting) CollectD in_hom_char) show "Cod (comp ?f x) = Cod (F x)" using 5 6 by (metis (mono_tags, lifting) CollectD in_hom_char) show "Map (comp ?f x) = Map (F x)" using 3 4 by simp qed qed thus "\?f : a \ b\ \ (\x. \x : unity \ dom ?f\ \ comp ?f x = F x)" using 1 by blast qed qed qed qed text\ \SetCat\ can be viewed as a concrete set category over its own element type @{typ 'a}, using @{term UP} as the required injection from @{typ 'a} to the universe of \SetCat\. \ corollary is_concrete_set_category: shows "concrete_set_category comp \ UNIV UP" proof - interpret S: set_category comp \ using is_set_category by auto show ?thesis proof show 1: "UP \ UNIV \ S.Univ" using UP_def terminal_char by force show "inj_on UP UNIV" using inj_UP by blast qed qed text\ As a consequence of the categoricity of the \set_category\ axioms, if @{term S} interprets \set_category\, and if @{term \} is a bijection between the universe of @{term S} and the elements of type @{typ 'a}, then @{term S} is isomorphic to the category \setcat\ of @{typ 'a} sets and functions between them constructed here. \ corollary set_category_iso_SetCat: fixes S :: "'s comp" and \ :: "'s \ 'e" assumes "set_category S \" and "bij_betw \ (Collect (category.terminal S)) UNIV" shows "\\. invertible_functor S (comp :: 'e setcat.arr comp) \ \ (\m. set_category.incl S \ m \ set_category.incl comp \ (\ m))" proof - interpret S: set_category S using assms by auto let ?\ = "inv_into S.Univ \" have "bij_betw (UP o \) S.Univ (Collect terminal)" proof (intro bij_betwI) show "UP o \ \ S.Univ \ Collect terminal" using assms(2) UP_mapsto by auto show "?\ o DOWN \ Collect terminal \ S.Univ" proof fix x :: "'e setcat.arr" assume x: "x \ Univ" show "(inv_into S.Univ \ \ DOWN) x \ S.Univ" using x assms(2) bij_betw_def comp_apply inv_into_into by (metis UNIV_I) qed fix t assume "t \ S.Univ" thus "(?\ o DOWN) ((UP o \) t) = t" using assms(2) bij_betw_inv_into_left by (metis comp_apply DOWN_UP) next fix t' :: "'e setcat.arr" assume "t' \ Collect terminal" thus "(UP o \) ((?\ o DOWN) t') = t'" using assms(2) by (simp add: bij_betw_def f_inv_into_f) qed thus ?thesis using assms(1) set_category_is_categorical [of S \ comp "UP o \"] is_set_category by auto qed end sublocale setcat \ set_category comp \ using is_set_category by simp sublocale setcat \ concrete_set_category comp \ UNIV UP using is_concrete_set_category by simp text\ By using a large enough cardinal, we can effectively eliminate the cardinality constraint on the sets that determine objects and thereby obtain a set category that is replete. This is the normal use case, which we want to streamline as much as possible, so it is useful to introduce a special locale for this purpose. \ locale replete_setcat = fixes dummy :: 'e begin interpretation SC: setcat dummy \cardSuc (cmax (card_of (UNIV :: 'e setcat.arr set)) natLeq)\ proof show "Card_order (cardSuc (cmax (card_of (UNIV :: 'e setcat.arr set)) natLeq)) \ infinite (Field (cardSuc (cmax (card_of (UNIV :: 'e setcat.arr set)) natLeq)))" by (metis Card_order_cmax Field_natLeq cardSuc_Card_order cardSuc_finite card_of_Card_order finite_cmax infinite_UNIV_char_0 natLeq_Card_order) qed text\ We don't want to expose the concrete details of the construction used to obtain the interpretation \SC\; instead, we want any facts proved about it to be derived solely from the assumptions of the @{locale set_category} locales. So we create another level of definitions here. \ definition comp where "comp \ SC.comp" definition UP where "UP \ SC.UP" definition DOWN where "DOWN \ SC.DOWN" sublocale set_category comp \cardSuc (cmax (card_of (UNIV :: 'e setcat.arr set)) natLeq)\ using SC.is_set_category comp_def by simp sublocale concrete_set_category comp \cardSuc (cmax (card_of (UNIV :: 'e setcat.arr set)) natLeq)\ UNIV UP using SC.is_concrete_set_category comp_def UP_def by simp sublocale replete_set_category comp .. lemma UP_mapsto: shows "UP \ UNIV \ Univ" using SC.UP_mapsto by (simp add: UP_def comp_def) lemma DOWN_mapsto: shows "DOWN \ Univ \ UNIV" by auto lemma DOWN_UP [simp]: shows "DOWN (UP x) = x" by (simp add: DOWN_def UP_def) lemma UP_DOWN [simp]: assumes "t \ Univ" shows "UP (DOWN t) = t" using assms DOWN_def UP_def by (simp add: DOWN_def UP_def comp_def) lemma inj_UP: shows "inj UP" by (metis DOWN_UP injI) lemma bij_UP: shows "bij_betw UP UNIV Univ" by (metis SC.bij_UP UP_def comp_def) end end diff --git a/thys/Category3/Yoneda.thy b/thys/Category3/Yoneda.thy --- a/thys/Category3/Yoneda.thy +++ b/thys/Category3/Yoneda.thy @@ -1,1090 +1,1088 @@ (* Title: Yoneda Author: Eugene W. Stark , 2016 Maintainer: Eugene W. Stark *) chapter Yoneda theory Yoneda imports DualCategory SetCat FunctorCategory begin text\ This theory defines the notion of a ``hom-functor'' and gives a proof of the Yoneda Lemma. In traditional developments of category theory based on set theories such as ZFC, hom-functors are normally defined to be functors into the large category \textbf{Set} whose objects are of \emph{all} sets and whose arrows are functions between sets. However, in HOL there does not exist a single ``type of all sets'', so the notion of the category of \emph{all} sets and functions does not make sense. To work around this, we consider a more general setting consisting of a category @{term C} together with a set category @{term S} and a function @{term "\ :: 'c * 'c \ 'c \ 's"} such that whenever @{term b} and @{term a} are objects of C then @{term "\ (b, a)"} maps \C.hom b a\ injectively to \S.Univ\. We show that these data induce a binary functor \Hom\ from \Cop\C\ to @{term S} in such a way that @{term \} is rendered natural in @{term "(b, a)"}. The Yoneda lemma is then proved for the Yoneda functor determined by \Hom\. \ section "Hom-Functors" text\ A hom-functor for a category @{term C} allows us to regard the hom-sets of @{term C} as objects of a category @{term S} of sets and functions. Any description of a hom-functor for @{term C} must therefore specify the category @{term S} and provide some sort of correspondence between arrows of @{term C} and elements of objects of @{term S}. If we are to think of each hom-set \C.hom b a\ of \C\ as corresponding to an object \Hom (b, a)\ of @{term S} then at a minimum it ought to be the case that the correspondence between arrows and elements is bijective between \C.hom b a\ and \Hom (b, a)\. The \hom_functor\ locale defined below captures this idea by assuming a set category @{term S} and a function @{term \} taking arrows of @{term C} to elements of \S.Univ\, such that @{term \} is injective on each set \C.hom b a\. We show that these data induce a functor \Hom\ from \Cop\C\ to \S\ in such a way that @{term \} becomes a natural bijection between \C.hom b a\ and \Hom (b, a)\. \ locale hom_functor = C: category C + Cop: dual_category C + CopxC: product_category Cop.comp C + S: replete_set_category S for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) and \ :: "'c * 'c \ 'c \ 's" + assumes maps_arr_to_Univ: "C.arr f \ \ (C.dom f, C.cod f) f \ S.Univ" and local_inj: "\ C.ide b; C.ide a \ \ inj_on (\ (b, a)) (C.hom b a)" begin notation S.in_hom ("\_ : _ \\<^sub>S _\") notation CopxC.comp (infixr "\" 55) notation CopxC.in_hom ("\_ : _ \ _\") definition set where "set ba \ \ (fst ba, snd ba) ` C.hom (fst ba) (snd ba)" lemma set_subset_Univ: assumes "C.ide b" and "C.ide a" shows "set (b, a) \ S.Univ" using assms set_def maps_arr_to_Univ CopxC.ide_char by auto definition \ :: "'c * 'c \ 's \ 'c" where "\ ba = inv_into (C.hom (fst ba) (snd ba)) (\ ba)" lemma \_mapsto: assumes "C.ide b" and "C.ide a" shows "\ (b, a) \ C.hom b a \ set (b, a)" using assms set_def maps_arr_to_Univ by auto lemma \_mapsto: assumes "C.ide b" and "C.ide a" shows "\ (b, a) \ set (b, a) \ C.hom b a" using assms set_def \_def local_inj by auto lemma \_\ [simp]: assumes "\f : b \ a\" shows "\ (b, a) (\ (b, a) f) = f" using assms local_inj [of b a] \_def by fastforce lemma \_\ [simp]: assumes "C.ide b" and "C.ide a" and "x \ set (b, a)" shows "\ (b, a) (\ (b, a) x) = x" using assms set_def local_inj \_def by auto lemma \_img_set: assumes "C.ide b" and "C.ide a" shows "\ (b, a) ` set (b, a) = C.hom b a" using assms \_def set_def local_inj by auto text\ A hom-functor maps each arrow @{term "(g, f)"} of @{term "CopxC"} to the arrow of the set category @{term[source=true] S} corresponding to the function that takes an arrow @{term h} of @{term C} to the arrow @{term "C f (C h g)"} of @{term C} obtained by precomposing with @{term g} and postcomposing with @{term f}. \ definition map where "map gf = (if CopxC.arr gf then S.mkArr (set (CopxC.dom gf)) (set (CopxC.cod gf)) (\ (CopxC.cod gf) o (\h. snd gf \ h \ fst gf) o \ (CopxC.dom gf)) else S.null)" lemma arr_map: assumes "CopxC.arr gf" shows "S.arr (map gf)" proof - have "\ (CopxC.cod gf) o (\h. snd gf \ h \ fst gf) o \ (CopxC.dom gf) \ set (CopxC.dom gf) \ set (CopxC.cod gf)" using assms \_mapsto [of "fst (CopxC.cod gf)" "snd (CopxC.cod gf)"] \_mapsto [of "fst (CopxC.dom gf)" "snd (CopxC.dom gf)"] by fastforce thus ?thesis using assms map_def S.arr_mkArr set_subset_Univ S.card_of_leq by simp qed lemma map_ide [simp]: assumes "C.ide b" and "C.ide a" shows "map (b, a) = S.mkIde (set (b, a))" proof - have "map (b, a) = S.mkArr (set (b, a)) (set (b, a)) (\ (b, a) o (\h. a \ h \ b) o \ (b, a))" using assms map_def by auto also have "... = S.mkArr (set (b, a)) (set (b, a)) (\h. h)" proof - have "S.mkArr (set (b, a)) (set (b, a)) (\h. h) = ..." using assms S.arr_mkArr set_subset_Univ set_def C.comp_arr_dom C.comp_cod_arr S.card_of_leq S.arr_mkIde by (intro S.mkArr_eqI', simp, fastforce) thus ?thesis by auto qed also have "... = S.mkIde (set (b, a))" using assms S.mkIde_as_mkArr set_subset_Univ by simp finally show ?thesis by auto qed lemma set_map: assumes "C.ide a" and "C.ide b" shows "S.set (map (b, a)) = set (b, a)" using assms map_ide S.set_mkIde set_subset_Univ by simp text\ The definition does in fact yield a functor. \ interpretation "functor" CopxC.comp S map proof fix gf assume "\CopxC.arr gf" thus "map gf = S.null" using map_def by auto next fix gf assume gf: "CopxC.arr gf" thus arr: "S.arr (map gf)" using gf arr_map by blast show "S.dom (map gf) = map (CopxC.dom gf)" - using gf set_subset_Univ \_mapsto map_def set_def S.card_of_leq S.arr_mkIde S.arr_mkArr - arr map_ide - by auto + using arr gf local.map_def map_ide by auto show "S.cod (map gf) = map (CopxC.cod gf)" using gf set_subset_Univ \_mapsto map_def set_def S.card_of_leq S.arr_mkIde S.arr_mkArr arr map_ide by auto next fix gf gf' assume gf': "CopxC.seq gf' gf" hence seq: "C.arr (fst gf) \ C.arr (snd gf) \ C.dom (snd gf') = C.cod (snd gf) \ C.arr (fst gf') \ C.arr (snd gf') \ C.dom (fst gf) = C.cod (fst gf')" by (elim CopxC.seqE C.seqE, auto) have 0: "S.arr (map (CopxC.comp gf' gf))" using gf' arr_map by blast have 1: "map (gf' \ gf) = S.mkArr (set (CopxC.dom gf)) (set (CopxC.cod gf')) (\ (CopxC.cod gf') o (\h. snd (gf' \ gf) \ h \ fst (gf' \ gf)) o \ (CopxC.dom gf))" using gf' map_def using CopxC.cod_comp CopxC.dom_comp by auto also have "... = S.mkArr (set (CopxC.dom gf)) (set (CopxC.cod gf')) (\ (CopxC.cod gf') \ (\h. snd gf' \ h \ fst gf') \ \ (CopxC.dom gf') \ (\ (CopxC.cod gf) \ (\h. snd gf \ h \ fst gf) \ \ (CopxC.dom gf)))" proof (intro S.mkArr_eqI') show "S.arr (S.mkArr (set (CopxC.dom gf)) (set (CopxC.cod gf')) (\ (CopxC.cod gf') \ (\h. snd (gf' \ gf) \ h \ fst (gf' \ gf)) \ \ (CopxC.dom gf)))" using 0 1 by simp show "\x. x \ set (CopxC.dom gf) \ (\ (CopxC.cod gf') \ (\h. snd (gf' \ gf) \ h \ fst (gf' \ gf)) \ \ (CopxC.dom gf)) x = (\ (CopxC.cod gf') \ (\h. snd gf' \ h \ fst gf') \ \ (CopxC.dom gf') \ (\ (CopxC.cod gf) \ (\h. snd gf \ h \ fst gf) \ \ (CopxC.dom gf))) x" proof - fix x assume "x \ set (CopxC.dom gf)" hence x: "x \ set (C.cod (fst gf), C.dom (snd gf))" - using gf' CopxC.seqE by (elim CopxC.seqE, fastforce) + by (simp add: seq) show "(\ (CopxC.cod gf') \ (\h. snd (gf' \ gf) \ h \ fst (gf' \ gf)) \ \ (CopxC.dom gf)) x = (\ (CopxC.cod gf') \ (\h. snd gf' \ h \ fst gf') \ \ (CopxC.dom gf') \ (\ (CopxC.cod gf) \ (\h. snd gf \ h \ fst gf) \ \ (CopxC.dom gf))) x" proof - have "(\ (CopxC.cod gf') o (\h. snd (gf' \ gf) \ h \ fst (gf' \ gf)) o \ (CopxC.dom gf)) x = \ (CopxC.cod gf') (snd (gf' \ gf) \ \ (CopxC.dom gf) x \ fst (gf' \ gf))" by simp also have "... = \ (CopxC.cod gf') (((\h. snd gf' \ h \ fst gf') \ \ (CopxC.dom gf') \ (\ (CopxC.dom gf') \ (\h. snd gf \ h \ fst gf))) (\ (CopxC.dom gf) x))" proof - have "C.ide (C.cod (fst gf)) \ C.ide (C.dom (snd gf))" using gf' by (elim CopxC.seqE, auto) hence "\\ (C.cod (fst gf), C.dom (snd gf)) x : C.cod (fst gf) \ C.dom (snd gf)\" using x \_mapsto by auto hence "\snd gf \ \ (C.cod (fst gf), C.dom (snd gf)) x \ fst gf : C.cod (fst gf') \ C.dom (snd gf')\" using x seq by auto thus ?thesis using seq \_\ C.comp_assoc by auto qed also have "... = (\ (CopxC.cod gf') \ (\h. snd gf' \ h \ fst gf') \ \ (CopxC.dom gf') \ (\ (CopxC.dom gf') \ (\h. snd gf \ h \ fst gf) \ \ (CopxC.dom gf))) x" by auto finally show ?thesis using seq by simp qed qed qed also have "... = map gf' \\<^sub>S map gf" using seq gf' map_def arr_map [of gf] arr_map [of gf'] S.comp_mkArr by auto finally show "map (gf' \ gf) = map gf' \\<^sub>S map gf" using seq gf' by auto qed interpretation binary_functor Cop.comp C S map .. lemma is_binary_functor: shows "binary_functor Cop.comp C S map" .. end sublocale hom_functor \ binary_functor Cop.comp C S map using is_binary_functor by auto context hom_functor begin text\ The map @{term \} determines a bijection between @{term "C.hom b a"} and @{term "set (b, a)"} which is natural in @{term "(b, a)"}. \ lemma \_local_bij: assumes "C.ide b" and "C.ide a" shows "bij_betw (\ (b, a)) (C.hom b a) (set (b, a))" using assms local_inj inj_on_imp_bij_betw set_def by auto lemma \_natural: assumes "C.arr g" and "C.arr f" and "h \ C.hom (C.cod g) (C.dom f)" shows "\ (C.dom g, C.cod f) (f \ h \ g) = S.Fun (map (g, f)) (\ (C.cod g, C.dom f) h)" proof - let ?\h = "\ (C.cod g, C.dom f) h" have \h: "?\h \ set (C.cod g, C.dom f)" using assms \_mapsto set_def by simp have gf: "CopxC.arr (g, f)" using assms by simp have "map (g, f) = S.mkArr (set (C.cod g, C.dom f)) (set (C.dom g, C.cod f)) (\ (C.dom g, C.cod f) \ (\h. f \ h \ g) \ \ (C.cod g, C.dom f))" using assms map_def by simp moreover have "S.arr (map (g, f))" using gf by simp ultimately have "S.Fun (map (g, f)) = restrict (\ (C.dom g, C.cod f) \ (\h. f \ h \ g) \ \ (C.cod g, C.dom f)) (set (C.cod g, C.dom f))" using S.Fun_mkArr by simp hence "S.Fun (map (g, f)) ?\h = (\ (C.dom g, C.cod f) \ (\h. f \ h \ g) \ \ (C.cod g, C.dom f)) ?\h" using \h by simp also have "... = \ (C.dom g, C.cod f) (f \ h \ g)" using assms(3) by simp finally show ?thesis by auto qed lemma Dom_map: assumes "C.arr g" and "C.arr f" shows "S.Dom (map (g, f)) = set (C.cod g, C.dom f)" using assms map_def preserves_arr S.set_mkIde S.arr_mkArr by auto lemma Cod_map: assumes "C.arr g" and "C.arr f" shows "S.Cod (map (g, f)) = set (C.dom g, C.cod f)" using assms map_def preserves_arr S.set_mkIde S.arr_mkArr by auto lemma Fun_map: assumes "C.arr g" and "C.arr f" shows "S.Fun (map (g, f)) = restrict (\ (C.dom g, C.cod f) o (\h. f \ h \ g) o \ (C.cod g, C.dom f)) (set (C.cod g, C.dom f))" using assms map_def preserves_arr by force lemma map_simp_1: assumes "C.arr g" and "C.ide a" shows "map (g, a) = S.mkArr (set (C.cod g, a)) (set (C.dom g, a)) (\ (C.dom g, a) o Cop.comp g o \ (C.cod g, a))" proof - have 1: "map (g, a) = S.mkArr (set (C.cod g, a)) (set (C.dom g, a)) (\ (C.dom g, a) o (\h. a \ h \ g) o \ (C.cod g, a))" using assms map_def by force also have "... = S.mkArr (set (C.cod g, a)) (set (C.dom g, a)) (\ (C.dom g, a) o Cop.comp g o \ (C.cod g, a))" using assms 1 preserves_arr [of "(g, a)"] set_def C.in_homI C.comp_cod_arr by (intro S.mkArr_eqI) auto finally show ?thesis by blast qed lemma map_simp_2: assumes "C.ide b" and "C.arr f" shows "map (b, f) = S.mkArr (set (b, C.dom f)) (set (b, C.cod f)) (\ (b, C.cod f) o C f o \ (b, C.dom f))" proof - have 1: "map (b, f) = S.mkArr (set (b, C.dom f)) (set (b, C.cod f)) (\ (b, C.cod f) o (\h. f \ h \ b) o \ (b, C.dom f))" using assms map_def by force also have "... = S.mkArr (set (b, C.dom f)) (set (b, C.cod f)) (\ (b, C.cod f) o C f o \ (b, C.dom f))" using assms 1 preserves_arr [of "(b, f)"] set_def C.in_homI C.comp_arr_dom by (intro S.mkArr_eqI) auto finally show ?thesis by blast qed end text\ Every category @{term C} has a hom-functor: take @{term S} to be the replete set category generated by the arrow type \'a\ of @{term C} and take @{term "\ (b, a)"} to be the map \S.UP :: 'a \ 'a SC.arr\. \ context category begin interpretation Cop: dual_category C .. interpretation CopxC: product_category Cop.comp C .. interpretation S: replete_setcat \undefined :: 'a\ . lemma has_hom_functor: shows "hom_functor C (S.comp :: 'a setcat.arr comp) (\_. S.UP)" proof show "\f. arr f \ S.UP f \ S.Univ" using S.UP_mapsto by auto show "\b a. \ide b; ide a\ \ inj_on S.UP (hom b a)" by (meson S.inj_UP injD inj_onI) qed end text\ The locales \set_valued_functor\ and \set_valued_transformation\ provide some abbreviations that are convenient when working with functors and natural transformations into a set category. \ locale set_valued_functor = C: category C + S: replete_set_category S + "functor" C S F for C :: "'c comp" and S :: "'s comp" and F :: "'c \ 's" begin abbreviation SET :: "'c \ 's set" where "SET a \ S.set (F a)" abbreviation DOM :: "'c \ 's set" where "DOM f \ S.Dom (F f)" abbreviation COD :: "'c \ 's set" where "COD f \ S.Cod (F f)" abbreviation FUN :: "'c \ 's \ 's" where "FUN f \ S.Fun (F f)" end locale set_valued_transformation = C: category C + S: replete_set_category S + F: set_valued_functor C S F + G: set_valued_functor C S G + natural_transformation C S F G \ for C :: "'c comp" and S :: "'s comp" and F :: "'c \ 's" and G :: "'c \ 's" and \ :: "'c \ 's" begin abbreviation DOM :: "'c \ 's set" where "DOM f \ S.Dom (\ f)" abbreviation COD :: "'c \ 's set" where "COD f \ S.Cod (\ f)" abbreviation FUN :: "'c \ 's \ 's" where "FUN f \ S.Fun (\ f)" end section "Yoneda Functors" text\ A Yoneda functor is the functor from @{term C} to \[Cop, S]\ obtained by ``currying'' a hom-functor in its first argument. \ locale yoneda_functor = C: category C + Cop: dual_category C + CopxC: product_category Cop.comp C + S: replete_set_category S + Hom: hom_functor C S \ + Cop_S: functor_category Cop.comp S + curried_functor' Cop.comp C S Hom.map for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) and \ :: "'c * 'c \ 'c \ 's" begin notation Cop_S.in_hom ("\_ : _ \\<^sub>[\<^sub>C\<^sub>o\<^sub>p\<^sub>,\<^sub>S\<^sub>] _\") abbreviation \ where "\ \ Hom.\" text\ An arrow of the functor category \[Cop, S]\ consists of a natural transformation bundled together with its domain and codomain functors. However, when considering a Yoneda functor from @{term[source=true] C} to \[Cop, S]\ we generally are only interested in the mapping @{term Y} that takes each arrow @{term f} of @{term[source=true] C} to the corresponding natural transformation @{term "Y f"}. The domain and codomain functors are then the identity transformations @{term "Y (C.dom f)"} and @{term "Y (C.cod f)"}. \ definition Y where "Y f \ Cop_S.Map (map f)" lemma Y_simp [simp]: assumes "C.arr f" shows "Y f = (\g. Hom.map (g, f))" using assms preserves_arr Y_def by simp lemma Y_ide_is_functor: assumes "C.ide a" shows "functor Cop.comp S (Y a)" using assms Y_def Hom.fixing_ide_gives_functor_2 by force lemma Y_arr_is_transformation: assumes "C.arr f" shows "natural_transformation Cop.comp S (Y (C.dom f)) (Y (C.cod f)) (Y f)" using assms Y_def [of f] map_def Hom.fixing_arr_gives_natural_transformation_2 preserves_dom preserves_cod by fastforce lemma Y_ide_arr [simp]: assumes a: "C.ide a" and "\g : b' \ b\" shows "\Y a g : Hom.map (b, a) \\<^sub>S Hom.map (b', a)\" and "Y a g = S.mkArr (Hom.set (b, a)) (Hom.set (b', a)) (\ (b', a) o Cop.comp g o \ (b, a))" using assms Hom.map_simp_1 by (fastforce, auto) lemma Y_arr_ide [simp]: assumes "C.ide b" and "\f : a \ a'\" shows "\Y f b : Hom.map (b, a) \\<^sub>S Hom.map (b, a')\" and "Y f b = S.mkArr (Hom.set (b, a)) (Hom.set (b, a')) (\ (b, a') o C f o \ (b, a))" using assms apply fastforce using assms Hom.map_simp_2 by auto end locale yoneda_functor_fixed_object = yoneda_functor C S \ for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) and \ :: "'c * 'c \ 'c \ 's" and a :: 'c + assumes ide_a: "C.ide a" sublocale yoneda_functor_fixed_object \ "functor" Cop.comp S "(Y a)" using ide_a Y_ide_is_functor by auto sublocale yoneda_functor_fixed_object \ set_valued_functor Cop.comp S "(Y a)" .. text\ The Yoneda lemma states that, given a category @{term C} and a functor @{term F} from @{term Cop} to a set category @{term S}, for each object @{term a} of @{term C}, the set of natural transformations from the contravariant functor @{term "Y a"} to @{term F} is in bijective correspondence with the set \F.SET a\ of elements of @{term "F a"}. Explicitly, if @{term e} is an arbitrary element of the set \F.SET a\, then the functions \\x. F.FUN (\ (b, a) x) e\ are the components of a natural transformation from @{term "Y a"} to @{term F}. Conversely, if @{term \} is a natural transformation from @{term "Y a"} to @{term F}, then the component @{term "\ b"} of @{term \} at an arbitrary object @{term b} is completely determined by the single arrow \\.FUN a (\ (a, a) a)))\, which is the the element of \F.SET a\ that corresponds to the image of the identity @{term a} under the function \\.FUN a\. Then @{term "\ b"} is the arrow from @{term "Y a b"} to @{term "F b"} corresponding to the function \\x. (F.FUN (\ (b, a) x) (\.FUN a (\ (a, a) a)))\ from \S.set (Y a b)\ to \F.SET b\. The above expressions look somewhat more complicated than the usual versions due to the need to account for the coercions @{term \} and @{term \}. \ locale yoneda_lemma = C: category C + Cop: dual_category C + S: replete_set_category S + F: set_valued_functor Cop.comp S F + yoneda_functor_fixed_object C S \ a for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) and \ :: "'c * 'c \ 'c \ 's" and F :: "'c \ 's" and a :: 'c begin text\ The mapping that evaluates the component @{term "\ a"} at @{term a} of a natural transformation @{term \} from @{term Y} to @{term F} on the element @{term "\ (a, a) a"} of @{term "SET a"}, yielding an element of @{term "F.SET a"}. \ definition \ :: "('c \ 's) \ 's" where "\ \ = S.Fun (\ a) (\ (a, a) a)" text\ The mapping that takes an element @{term e} of @{term "F.SET a"} and produces a map on objects of @{term[source=true] C} whose value at @{term b} is the arrow of @{term[source=true] S} corresponding to the function @{term "(\x. F.FUN (\ (b, a) x) e) \ Hom.set (b, a) \ F.SET b"}. \ definition \o :: "'s \ 'c \ 's" where "\o e b = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. F.FUN (\ (b, a) x) e)" lemma \o_e_ide: assumes e: "e \ S.set (F a)" and b: "C.ide b" shows "\\o e b : Y a b \\<^sub>S F b\" and "\o e b = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. F.FUN (\ (b, a) x) e)" proof - show "\o e b = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. F.FUN (\ (b, a) x) e)" using \o_def by auto moreover have "(\x. F.FUN (\ (b, a) x) e) \ Hom.set (b, a) \ F.SET b" proof fix x assume x: "x \ Hom.set (b, a)" thus "F.FUN (\ (b, a) x) e \ F.SET b" using assms e ide_a Hom.\_mapsto S.Fun_mapsto [of "F (\ (b, a) x)"] by force qed ultimately show "\\o e b : Y a b \\<^sub>S F b\" using ide_a b S.mkArr_in_hom [of "Hom.set (b, a)" "F.SET b"] Hom.set_subset_Univ S.mkIde_set by force qed text\ For each @{term "e \ F.SET a"}, the mapping @{term "\o e"} gives the components of a natural transformation @{term \} from @{term "Y a"} to @{term F}. \ lemma \o_e_induces_transformation: assumes e: "e \ S.set (F a)" shows "transformation_by_components Cop.comp S (Y a) F (\o e)" proof fix b :: 'c assume b: "Cop.ide b" show "\\o e b : Y a b \\<^sub>S F b\" using ide_a b e \o_e_ide by simp next fix g :: 'c assume g: "Cop.arr g" let ?b = "Cop.dom g" let ?b' = "Cop.cod g" show "\o e (Cop.cod g) \\<^sub>S Y a g = F g \\<^sub>S \o e (Cop.dom g)" proof - have 1: "\o e (Cop.cod g) \\<^sub>S Y a g = S.mkArr (Hom.set (?b, a)) (F.SET ?b') ((\x. F.FUN (\ (?b', a) x) e) o (\ (?b', a) o Cop.comp g o \ (?b, a)))" proof - have "S.arr (S.mkArr (Hom.set (Cop.cod g, a)) (F.SET (Cop.cod g)) (\s. F.FUN (\ (Cop.cod g, a) s) e)) \ S.dom (S.mkArr (Hom.set (Cop.cod g, a)) (F.SET (Cop.cod g)) (\s. F.FUN (\ (Cop.cod g, a) s) e)) = Y a (Cop.cod g) \ S.cod (S.mkArr (Hom.set (Cop.cod g, a)) (F.SET (Cop.cod g)) (\s. F.FUN (\ (Cop.cod g, a) s) e)) = F (Cop.cod g)" using Cop.cod_char \o_e_ide [of e ?b'] \o_e_ide [of e ?b'] e g by (metis Cop.ide_char Cop.ide_cod S.in_homE) moreover have "Y a g = S.mkArr (Hom.set (Cop.dom g, a)) (Hom.set (Cop.cod g, a)) (\ (Cop.cod g, a) \ Cop.comp g \ \ (Cop.dom g, a))" using Y_ide_arr [of a g ?b' ?b] ide_a g by auto ultimately show ?thesis using ide_a e g Y_ide_arr Cop.cod_char \o_e_ide S.comp_mkArr [of "Hom.set (?b, a)" "Hom.set (?b', a)" "\ (?b', a) o Cop.comp g o \ (?b, a)" "F.SET ?b'" "\x. F.FUN (\ (?b', a) x) e"] by (metis C.ide_dom Cop.arr_char preserves_arr) qed also have "... = S.mkArr (Hom.set (?b, a)) (F.SET ?b') (F.FUN g o (\x. F.FUN (\ (?b, a) x) e))" proof (intro S.mkArr_eqI') have "(\x. F.FUN (\ (?b', a) x) e) o (\ (?b', a) o Cop.comp g o \ (?b, a)) \ Hom.set (?b, a) \ F.SET ?b'" proof - have "S.arr (S (\o e ?b') (Y a g))" using ide_a e g \o_e_ide [of e ?b'] Y_ide_arr(1) [of a "C.dom g" "C.cod g" g] Cop.ide_char Cop.ide_cod by blast thus ?thesis using 1 S.arr_mkArr by simp qed thus "S.arr (S.mkArr (Hom.set (?b, a)) (F.SET ?b') ((\x. F.FUN (\ (?b', a) x) e) o (\ (?b', a) o Cop.comp g o \ (?b, a))))" using ide_a e g Hom.set_subset_Univ S.arr_mkArr by force show "\x. x \ Hom.set (?b, a) \ ((\x. F.FUN (\ (?b', a) x) e) o (\ (?b', a) o Cop.comp g o \ (?b, a))) x = (F.FUN g o (\x. F.FUN (\ (?b, a) x) e)) x" proof - fix x assume x: "x \ Hom.set (?b, a)" have "((\x. (F.FUN o \ (?b', a)) x e) o (\ (?b', a) o Cop.comp g o \ (?b, a))) x = F.FUN (\ (?b', a) (\ (?b', a) (C (\ (?b, a) x) g))) e" by simp also have "... = (F.FUN g o (F.FUN o \ (?b, a)) x) e" proof - have 1: "\\ (Cop.dom g, a) x : Cop.dom g \ a\" using ide_a x g Hom.\_mapsto [of ?b a] by auto moreover have "S.seq (F g) (F (\ (C.cod g, a) x))" using 1 g by (intro S.seqI', auto) moreover have "\ (C.dom g, a) (\ (C.dom g, a) (C (\ (C.cod g, a) x) g)) = C (\ (C.cod g, a) x) g" using g 1 Hom.\_\ [of "C (\ (?b, a) x) g" ?b' a] by fastforce ultimately show ?thesis using assms F.preserves_comp by fastforce qed also have "... = (F.FUN g o (\x. F.FUN (\ (?b, a) x) e)) x" by fastforce finally show "((\x. F.FUN (\ (?b', a) x) e) o (\ (?b', a) o Cop.comp g o \ (?b, a))) x = (F.FUN g o (\x. F.FUN (\ (?b, a) x) e)) x" by simp qed qed also have "... = F g \\<^sub>S \o e (Cop.dom g)" proof - have "S.arr (F g) \ F g = S.mkArr (F.SET ?b) (F.SET ?b') (F.FUN g)" using g S.mkArr_Fun [of "F g"] by simp moreover have "S.arr (\o e ?b) \ \o e ?b = S.mkArr (Hom.set (?b, a)) (F.SET ?b) (\x. F.FUN (\ (?b, a) x) e)" using e g \o_e_ide by (metis C.ide_cod Cop.arr_char Cop.dom_char S.in_homE) ultimately show ?thesis using S.comp_mkArr [of "Hom.set (?b, a)" "F.SET ?b" "\x. F.FUN (\ (?b, a) x) e" "F.SET ?b'" "F.FUN g"] by metis qed finally show ?thesis by blast qed qed abbreviation \ :: "'s \ 'c \ 's" where "\ e \ transformation_by_components.map Cop.comp S (Y a) (\o e)" end locale yoneda_lemma_fixed_e = yoneda_lemma C S \ F a for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) and \ :: "'c * 'c \ 'c \ 's" and F :: "'c \ 's" and a :: 'c and e :: 's + assumes E: "e \ F.SET a" begin interpretation \e: transformation_by_components Cop.comp S \Y a\ F \\o e\ using E \o_e_induces_transformation by auto lemma natural_transformation_\e: shows "natural_transformation Cop.comp S (Y a) F (\ e)" .. lemma \e_ide: assumes "Cop.ide b" shows "S.arr (\ e b)" and "\ e b = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. F.FUN (\ (b, a) x) e)" using assms apply fastforce using assms \o_def by auto end locale yoneda_lemma_fixed_\ = yoneda_lemma C S \ F a + \: set_valued_transformation Cop.comp S "Y a" F \ for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) and \ :: "'c * 'c \ 'c \ 's" and F :: "'c \ 's" and a :: 'c and \ :: "'c \ 's" begin text\ The key lemma: The component @{term "\ b"} of @{term \} at an arbitrary object @{term b} is completely determined by the single element @{term "\.FUN a (\ (a, a) a) \ F.SET a"}. \ lemma \_ide: assumes b: "Cop.ide b" shows "\ b = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. (F.FUN (\ (b, a) x) (\.FUN a (\ (a, a) a))))" proof - let ?\a = "\ (a, a) a" have \a: "\ (a, a) a \ Hom.set (a, a)" using ide_a Hom.\_mapsto [of a a] by fastforce have 1: "\ b = S.mkArr (Hom.set (b, a)) (F.SET b) (\.FUN b)" using ide_a b S.mkArr_Fun [of "\ b"] Hom.set_map by auto also have "... = S.mkArr (Hom.set (b, a)) (F.SET b) (\x. (F.FUN (\ (b, a) x) (\.FUN a ?\a)))" proof (intro S.mkArr_eqI') show "S.arr (S.mkArr (Hom.set (b, a)) (F.SET b) (\.FUN b))" using ide_a b 1 S.mkArr_Fun [of "\ b"] Hom.set_map by auto show "\x. x \ Hom.set (b, a) \ \.FUN b x = (F.FUN (\ (b, a) x) (\.FUN a ?\a))" proof - fix x assume x: "x \ Hom.set (b, a)" let ?\x = "\ (b, a) x" have \x: "\?\x : b \ a\" using ide_a b x Hom.\_mapsto [of b a] by auto show "\.FUN b x = (F.FUN (\ (b, a) x) (\.FUN a ?\a))" proof - have "\.FUN b x = S.Fun (\ b \\<^sub>S Y a ?\x) ?\a" proof - have "\.FUN b x = \.FUN b ((\ (b, a) o Cop.comp ?\x) a)" using ide_a b x \x Hom.\_\ by (metis C.comp_cod_arr C.in_homE C.ide_dom Cop.comp_def comp_apply) also have "\.FUN b ((\ (b, a) o Cop.comp ?\x) a) = (\.FUN b o (\ (b, a) o Cop.comp ?\x o \ (a, a))) ?\a" using ide_a b C.ide_in_hom by simp also have "... = S.Fun (\ b \\<^sub>S Y a ?\x) ?\a" proof - have "S.arr (Y a ?\x)" using ide_a \x preserves_arr by (elim C.in_homE, auto) moreover have "Y a ?\x = S.mkArr (Hom.set (a, a)) (SET b) (\ (b, a) \ Cop.comp ?\x \ \ (a, a))" using ide_a b \x preserves_hom Y_ide_arr Hom.set_map C.arrI by auto moreover have "S.arr (\ b) \ \ b = S.mkArr (SET b) (F.SET b) (\.FUN b)" using ide_a b S.mkArr_Fun [of "\ b"] by simp ultimately have "S.seq (\ b) (Y a ?\x) \ \ b \\<^sub>S Y a ?\x = S.mkArr (Hom.set (a, a)) (F.SET b) (\.FUN b o (\ (b, a) \ Cop.comp ?\x \ \ (a, a)))" using 1 S.comp_mkArr S.seqI by (metis S.cod_mkArr S.dom_mkArr) thus ?thesis using ide_a b x Hom.\_mapsto S.Fun_mkArr by force qed finally show ?thesis by auto qed also have "... = S.Fun (F ?\x \\<^sub>S \ a) ?\a" using ide_a b \x \.naturality [of ?\x] by force also have "... = F.FUN ?\x (\.FUN a ?\a)" proof - have "restrict (S.Fun (F ?\x \\<^sub>S \ a)) (Hom.set (a, a)) = restrict (F.FUN (\ (b, a) x) o \.FUN a) (Hom.set (a, a))" proof - have "S.arr (F ?\x \\<^sub>S \ a) \ F ?\x \\<^sub>S \ a = S.mkArr (Hom.set (a, a)) (F.SET b) (F.FUN ?\x o \.FUN a)" proof show 1: "S.seq (F ?\x) (\ a)" using \x ide_a \.preserves_cod F.preserves_dom by (elim C.in_homE, auto) have "\ a = S.mkArr (Hom.set (a, a)) (F.SET a) (\.FUN a)" using ide_a 1 S.mkArr_Fun [of "\ a"] Hom.set_map by auto moreover have "F ?\x = S.mkArr (F.SET a) (F.SET b) (F.FUN ?\x)" using x \x 1 S.mkArr_Fun [of "F ?\x"] by fastforce ultimately show "F ?\x \\<^sub>S \ a = S.mkArr (Hom.set (a, a)) (F.SET b) (F.FUN ?\x o \.FUN a)" using 1 S.comp_mkArr [of "Hom.set (a, a)" "F.SET a" "\.FUN a" "F.SET b" "F.FUN ?\x"] by (elim S.seqE, auto) qed thus ?thesis by force qed thus "S.Fun (F (\ (b, a) x) \\<^sub>S \ a) ?\a = F.FUN ?\x (\.FUN a ?\a)" using ide_a \a restr_eqE [of "S.Fun (F ?\x \\<^sub>S \ a)" "Hom.set (a, a)" "F.FUN ?\x o \.FUN a"] by simp qed finally show ?thesis by simp qed qed qed finally show ?thesis by auto qed text\ Consequently, if @{term \'} is any natural transformation from @{term "Y a"} to @{term F} that agrees with @{term \} at @{term a}, then @{term "\' = \"}. \ lemma eqI: assumes "natural_transformation Cop.comp S (Y a) F \'" and "\' a = \ a" shows "\' = \" proof (intro NaturalTransformation.eqI) interpret \': natural_transformation Cop.comp S \Y a\ F \' using assms by auto interpret T': yoneda_lemma_fixed_\ C S \ F a \' .. show "natural_transformation Cop.comp S (Y a) F \" .. show "natural_transformation Cop.comp S (Y a) F \'" .. show "\b. Cop.ide b \ \' b = \ b" using assms(2) \_ide T'.\_ide by simp qed end context yoneda_lemma begin text\ One half of the Yoneda lemma: The mapping @{term \} is an injection, with left inverse @{term \}, from the set @{term "F.SET a"} to the set of natural transformations from @{term "Y a"} to @{term F}. \ lemma \_is_injection: assumes "e \ F.SET a" shows "natural_transformation Cop.comp S (Y a) F (\ e)" and "\ (\ e) = e" proof - interpret yoneda_lemma_fixed_e C S \ F a e using assms by (unfold_locales, auto) interpret \e: natural_transformation Cop.comp S \Y a\ F \\ e\ using natural_transformation_\e by auto show "natural_transformation Cop.comp S (Y a) F (\ e)" .. show "\ (\ e) = e" unfolding \_def using assms \e_ide S.Fun_mkArr Hom.\_mapsto Hom.\_\ ide_a F.preserves_ide S.Fun_ide restrict_apply C.ide_in_hom by (auto simp add: Pi_iff) qed lemma \\_in_Fa: assumes "natural_transformation Cop.comp S (Y a) F \" shows "\ \ \ F.SET a" proof - interpret \: natural_transformation Cop.comp S \Y a\ F \ using assms by auto interpret yoneda_lemma_fixed_\ C S \ F a \ .. show ?thesis proof (unfold \_def) have "S.arr (\ a) \ S.Dom (\ a) = Hom.set (a, a) \ S.Cod (\ a) = F.SET a" using ide_a Hom.set_map by auto hence "\.FUN a \ Hom.set (a, a) \ F.SET a" using S.Fun_mapsto by blast thus "\.FUN a (\ (a, a) a) \ F.SET a" using ide_a Hom.\_mapsto by fastforce qed qed text\ The other half of the Yoneda lemma: The mapping @{term \} is a surjection, with right inverse @{term \}, taking natural transformations from @{term "Y a"} to @{term F} to elements of @{term "F.SET a"}. \ lemma \_is_surjection: assumes "natural_transformation Cop.comp S (Y a) F \" shows "\ \ \ F.SET a" and "\ (\ \) = \" proof - interpret natural_transformation Cop.comp S \Y a\ F \ using assms by auto interpret yoneda_lemma_fixed_\ C S \ F a \ .. show 1: "\ \ \ F.SET a" using assms \\_in_Fa by auto interpret yoneda_lemma_fixed_e C S \ F a \\ \\ using 1 by (unfold_locales, auto) interpret \e: natural_transformation Cop.comp S \Y a\ F \\ (\ \)\ using natural_transformation_\e by auto show "\ (\ \) = \" proof (intro eqI) show "natural_transformation Cop.comp S (Y a) F (\ (\ \))" .. show "\ (\ \) a = \ a" using ide_a \_ide [of a] \e_ide \_def by simp qed qed text\ The main result. \ theorem yoneda_lemma: shows "bij_betw \ (F.SET a) {\. natural_transformation Cop.comp S (Y a) F \}" using \_is_injection \_is_surjection by (intro bij_betwI, auto) end text\ We now consider the special case in which @{term F} is the contravariant functor @{term "Y a'"}. Then for any @{term e} in \Hom.set (a, a')\ we have @{term "\ e = Y (\ (a, a') e)"}, and @{term \} is a bijection from \Hom.set (a, a')\ to the set of natural transformations from @{term "Y a"} to @{term "Y a'"}. It then follows that that the Yoneda functor @{term Y} is a fully faithful functor from @{term C} to the functor category \[Cop, S]\. \ locale yoneda_lemma_for_hom = C: category C + Cop: dual_category C + S: replete_set_category S + yoneda_functor_fixed_object C S \ a + Ya': yoneda_functor_fixed_object C S \ a' + yoneda_lemma C S \ "Y a'" a for C :: "'c comp" (infixr "\" 55) and S :: "'s comp" (infixr "\\<^sub>S" 55) and \ :: "'c * 'c \ 'c \ 's" and F :: "'c \ 's" and a :: 'c and a' :: 'c + assumes ide_a': "C.ide a'" begin text\ In case @{term F} is the functor @{term "Y a'"}, for any @{term "e \ Hom.set (a, a')"} the induced natural transformation @{term "\ e"} from @{term "Y a"} to @{term "Y a'"} is just @{term "Y (\ (a, a') e)"}. \ lemma \_equals_Yo\: assumes e: "e \ Hom.set (a, a')" shows "\ e = Y (\ (a, a') e)" proof - let ?\e = "\ (a, a') e" have \e: "\?\e : a \ a'\" using ide_a ide_a' e Hom.\_mapsto [of a a'] by auto interpret Ye: natural_transformation Cop.comp S \Y a\ \Y a'\ \Y ?\e\ using Y_arr_is_transformation [of ?\e] \e by (elim C.in_homE, auto) interpret yoneda_lemma_fixed_e C S \ \Y a'\ a e using ide_a ide_a' e S.set_mkIde Hom.set_map by (unfold_locales, simp_all) interpret \e: natural_transformation Cop.comp S \Y a\ \Y a'\ \\ e\ using natural_transformation_\e by auto interpret yoneda_lemma_fixed_\ C S \ \Y a'\ a \\ e\ .. have "natural_transformation Cop.comp S (Y a) (Y a') (Y ?\e)" .. moreover have "natural_transformation Cop.comp S (Y a) (Y a') (\ e)" .. moreover have "\ e a = Y ?\e a" proof - have 1: "S.arr (\ e a)" using ide_a e \e.preserves_reflects_arr by simp have 2: "\ e a = S.mkArr (Hom.set (a, a)) (Ya'.SET a) (\x. Ya'.FUN (\ (a, a) x) e)" using ide_a \o_def \e_ide by simp also have "... = S.mkArr (Hom.set (a, a)) (Hom.set (a, a')) (\ (a, a') o C ?\e o \ (a, a))" proof (intro S.mkArr_eqI) show "S.arr (S.mkArr (Hom.set (a, a)) (Ya'.SET a) (\x. Ya'.FUN (\ (a, a) x) e))" using ide_a e 1 2 by simp show "Hom.set (a, a) = Hom.set (a, a)" .. show 3: "Ya'.SET a = Hom.set (a, a')" using ide_a ide_a' Y_simp Hom.set_map by simp show "\x. x \ Hom.set (a, a) \ Ya'.FUN (\ (a, a) x) e = (\ (a, a') o C ?\e o \ (a, a)) x" proof - fix x assume x: "x \ Hom.set (a, a)" have \x: "\\ (a, a) x : a \ a\" using ide_a x Hom.\_mapsto [of a a] by auto have "S.arr (Y a' (\ (a, a) x)) \ Y a' (\ (a, a) x) = S.mkArr (Hom.set (a, a')) (Hom.set (a, a')) (\ (a, a') \ Cop.comp (\ (a, a) x) \ \ (a, a'))" using Y_ide_arr ide_a ide_a' \x by blast hence "Ya'.FUN (\ (a, a) x) e = (\ (a, a') \ Cop.comp (\ (a, a) x) \ \ (a, a')) e" using e 3 S.Fun_mkArr Ya'.preserves_reflects_arr [of "\ (a, a) x"] by simp also have "... = (\ (a, a') o C ?\e o \ (a, a)) x" by simp finally show "Ya'.FUN (\ (a, a) x) e = (\ (a, a') o C ?\e o \ (a, a)) x" by auto qed qed also have "... = Y ?\e a" using ide_a ide_a' Y_arr_ide \e by simp finally show "\ e a = Y ?\e a" by auto qed ultimately show ?thesis using eqI by auto qed lemma Y_injective_on_homs: assumes "\f : a \ a'\" and "\f' : a \ a'\" and "map f = map f'" shows "f = f'" proof - have "f = \ (a, a') (\ (a, a') f)" using assms ide_a Hom.\_\ by simp also have "... = \ (a, a') (\ (\ (\ (a, a') f)))" using ide_a ide_a' assms(1) \_is_injection Hom.\_mapsto Hom.set_map by (elim C.in_homE, simp add: Pi_iff) also have "... = \ (a, a') (\ (Y (\ (a, a') (\ (a, a') f))))" using assms Hom.\_mapsto [of a a'] \_equals_Yo\ [of "\ (a, a') f"] by force also have "... = \ (a, a') (\ (\ (\ (a, a') f')))" using assms Hom.\_mapsto [of a a'] ide_a Hom.\_\ Y_def \_equals_Yo\ [of "\ (a, a') f'"] by fastforce also have "... = \ (a, a') (\ (a, a') f')" using ide_a ide_a' assms(2) \_is_injection Hom.\_mapsto Hom.set_map by (elim C.in_homE, simp add: Pi_iff) also have "... = f'" using assms ide_a Hom.\_\ by simp finally show "f = f'" by auto qed lemma Y_surjective_on_homs: assumes \: "natural_transformation Cop.comp S (Y a) (Y a') \" shows "Y (\ (a, a') (\ \)) = \" using ide_a ide_a' \ \_is_surjection \_equals_Yo\ \\_in_Fa Hom.set_map by simp end context yoneda_functor begin lemma is_faithful_functor: shows "faithful_functor C Cop_S.comp map" proof fix f :: 'c and f' :: 'c assume par: "C.par f f'" and ff': "map f = map f'" show "f = f'" proof - interpret Ya': yoneda_functor_fixed_object C S \ \C.cod f\ using par by (unfold_locales, auto) interpret yoneda_lemma_for_hom C S \ \Y (C.cod f)\ \C.dom f\ \C.cod f\ using par by (unfold_locales, auto) show "f = f'" using par ff' Y_injective_on_homs [of f f'] by fastforce qed qed lemma is_full_functor: shows "full_functor C Cop_S.comp map" proof fix a :: 'c and a' :: 'c and t assume a: "C.ide a" and a': "C.ide a'" assume t: "\t : map a \\<^sub>[\<^sub>C\<^sub>o\<^sub>p\<^sub>,\<^sub>S\<^sub>] map a'\" show "\e. \e : a \ a'\ \ map e = t" proof interpret Ya': yoneda_functor_fixed_object C S \ a' using a' by (unfold_locales, auto) interpret yoneda_lemma_for_hom C S \ \Y a'\ a a' using a a' by (unfold_locales, auto) have NT: "natural_transformation Cop.comp S (Y a) (Y a') (Cop_S.Map t)" using t a' Y_def Cop_S.Map_dom Cop_S.Map_cod Cop_S.dom_char Cop_S.cod_char Cop_S.in_homE Cop_S.arrE by metis hence 1: "\ (Cop_S.Map t) \ Hom.set (a, a')" using \\_in_Fa ide_a ide_a' Hom.set_map by simp moreover have "map (\ (a, a') (\ (Cop_S.Map t))) = t" proof (intro Cop_S.arr_eqI) have 2: "\map (\ (a, a') (\ (Cop_S.Map t))) : map a \\<^sub>[\<^sub>C\<^sub>o\<^sub>p\<^sub>,\<^sub>S\<^sub>] map a'\" using 1 ide_a ide_a' Hom.\_mapsto [of a a'] by blast show "Cop_S.arr t" using t by blast show "Cop_S.arr (map (\ (a, a') (\ (Cop_S.Map t))))" using 2 by blast show 3: "Cop_S.Map (map (\ (a, a') (\ (Cop_S.Map t)))) = Cop_S.Map t" using NT Y_surjective_on_homs Y_def by simp show 4: "Cop_S.Dom (map (\ (a, a') (\ (Cop_S.Map t)))) = Cop_S.Dom t" using t 2 functor_axioms Cop_S.Map_dom by (metis Cop_S.in_homE) show "Cop_S.Cod (map (\ (a, a') (\ (Cop_S.Map t)))) = Cop_S.Cod t" using 2 3 4 t Cop_S.Map_cod by (metis Cop_S.in_homE) qed ultimately show "\\ (a, a') (\ (Cop_S.Map t)) : a \ a'\ \ map (\ (a, a') (\ (Cop_S.Map t))) = t" using ide_a ide_a' Hom.\_mapsto by auto qed qed end sublocale yoneda_functor \ faithful_functor C Cop_S.comp map using is_faithful_functor by auto sublocale yoneda_functor \ full_functor C Cop_S.comp map using is_full_functor by auto sublocale yoneda_functor \ fully_faithful_functor C Cop_S.comp map .. end diff --git a/thys/MonoidalCategory/MonoidalCategory.thy b/thys/MonoidalCategory/MonoidalCategory.thy --- a/thys/MonoidalCategory/MonoidalCategory.thy +++ b/thys/MonoidalCategory/MonoidalCategory.thy @@ -1,4467 +1,4462 @@ (* Title: MonoidalCategory Author: Eugene W. Stark , 2017 Maintainer: Eugene W. Stark *) chapter "Monoidal Category" text_raw\ \label{monoidal-category-chap} \ text \ In this theory, we define the notion ``monoidal category,'' and develop consequences of the definition. The main result is a proof of MacLane's coherence theorem. \ theory MonoidalCategory imports Category3.EquivalenceOfCategories begin section "Monoidal Category" text \ A typical textbook presentation defines a monoidal category to be a category @{term C} equipped with (among other things) a binary ``tensor product'' functor \\: C \ C \ C\ and an ``associativity'' natural isomorphism \\\, whose components are isomorphisms \\ (a, b, c): (a \ b) \ c \ a \ (b \ c)\ for objects \a\, \b\, and \c\ of \C\. This way of saying things avoids an explicit definition of the functors that are the domain and codomain of \\\ and, in particular, what category serves as the domain of these functors. The domain category is in fact the product category \C \ C \ C\ and the domain and codomain of \\\ are the functors \T o (T \ C): C \ C \ C \ C\ and \T o (C \ T): C \ C \ C \ C\. In a formal development, though, we can't gloss over the fact that \C \ C \ C\ has to mean either \C \ (C \ C)\ or \(C \ C) \ C\, which are not formally identical, and that associativities are somehow involved in the definitions of the functors \T o (T \ C)\ and \T o (C \ T)\. Here we define the \binary_endofunctor\ locale to codify our choices about what \C \ C \ C\, \T o (T \ C)\, and \T o (C \ T)\ actually mean. In particular, we choose \C \ C \ C\ to be \C \ (C \ C)\ and define the functors \T o (T \ C)\, and \T o (C \ T)\ accordingly. \ locale binary_endofunctor = C: category C + CC: product_category C C + CCC: product_category C CC.comp + binary_functor C C C T for C :: "'a comp" (infixr "\" 55) and T :: "'a * 'a \ 'a" begin definition ToTC where "ToTC f \ if CCC.arr f then T (T (fst f, fst (snd f)), snd (snd f)) else C.null" lemma functor_ToTC: shows "functor CCC.comp C ToTC" using ToTC_def apply unfold_locales apply auto[4] proof - fix f g assume gf: "CCC.seq g f" show "ToTC (CCC.comp g f) = ToTC g \ ToTC f" - using gf unfolding CCC.seq_char CC.seq_char ToTC_def apply auto - apply (elim C.seqE, auto) - by (metis C.seqI CC.comp_simp CC.seqI fst_conv preserves_comp preserves_seq snd_conv) + using gf unfolding CCC.seq_char CC.seq_char ToTC_def + apply auto + by (metis CC.comp_simp CC.seqI fst_conv preserves_comp preserves_seq snd_conv) qed lemma ToTC_simp [simp]: assumes "C.arr f" and "C.arr g" and "C.arr h" shows "ToTC (f, g, h) = T (T (f, g), h)" using assms ToTC_def CCC.arr_char by simp definition ToCT where "ToCT f \ if CCC.arr f then T (fst f, T (fst (snd f), snd (snd f))) else C.null" lemma functor_ToCT: shows "functor CCC.comp C ToCT" using ToCT_def apply unfold_locales apply auto[4] proof - fix f g assume gf: "CCC.seq g f" show "ToCT (CCC.comp g f) = ToCT g \ ToCT f" - using gf unfolding CCC.seq_char CC.seq_char ToCT_def apply auto - apply (elim C.seqE, auto) - by (metis C.seqI CC.comp_simp CC.seqI fst_conv preserves_comp preserves_seq snd_conv) + using gf unfolding CCC.seq_char CC.seq_char ToCT_def + apply auto + by (metis CC.comp_simp CC.seq_char as_nat_trans.preserves_comp_2 fst_conv + preserves_reflects_arr snd_conv) qed lemma ToCT_simp [simp]: assumes "C.arr f" and "C.arr g" and "C.arr h" shows "ToCT (f, g, h) = T (f, T (g, h))" using assms ToCT_def CCC.arr_char by simp end text \ Our primary definition for ``monoidal category'' follows the somewhat non-traditional development in \cite{Etingof15}. There a monoidal category is defined to be a category \C\ equipped with a binary \emph{tensor product} functor \T: C \ C \ C\, an \emph{associativity isomorphism}, which is a natural isomorphism \\: T o (T \ C) \ T o (C \ T)\, a \emph{unit object} \\\ of \C\, and an isomorphism \\: T (\, \) \ \\, subject to two axioms: the \emph{pentagon axiom}, which expresses the commutativity of certain pentagonal diagrams involving components of \\\, and the left and right \emph{unit axioms}, which state that the endofunctors \T (\, -)\ and \T (-, \)\ are equivalences of categories. This definition is formalized in the \monoidal_category\ locale. In more traditional developments, the definition of monoidal category involves additional left and right \emph{unitor} isomorphisms \\\ and \\\ and associated axioms involving their components. However, as is shown in \cite{Etingof15} and formalized here, the unitors are uniquely determined by \\\ and their values \\(\)\ and \\(\)\ at \\\, which coincide. Treating \\\ and \\\ as defined notions results in a more economical basic definition of monoidal category that requires less data to be given, and has a similar effect on the definition of ``monoidal functor.'' Moreover, in the context of the formalization of categories that we use here, the unit object \\\ also need not be given separately, as it can be obtained as the codomain of the isomorphism \\\. \ locale monoidal_category = category C + CC: product_category C C + CCC: product_category C CC.comp + T: binary_endofunctor C T + \: natural_isomorphism CCC.comp C T.ToTC T.ToCT \ + L: equivalence_functor C C "\f. T (cod \, f)" + R: equivalence_functor C C "\f. T (f, cod \)" for C :: "'a comp" (infixr "\" 55) and T :: "'a * 'a \ 'a" and \ :: "'a * 'a * 'a \ 'a" and \ :: 'a + assumes \_in_hom': "\\ : T (cod \, cod \) \ cod \\" and \_is_iso: "iso \" and pentagon: "\ ide a; ide b; ide c; ide d \ \ T (a, \ (b, c, d)) \ \ (a, T (b, c), d) \ T (\ (a, b, c), d) = \ (a, b, T (c, d)) \ \ (T (a, b), c, d)" begin text\ We now define helpful notation and abbreviations to improve readability. We did not define and use the notation \\\ for the tensor product in the definition of the locale because to define \\\ as a binary operator requires that it be in curried form, whereas for \T\ to be a binary functor requires that it take a pair as its argument. \ definition unity :: 'a ("\") where "unity \ cod \" abbreviation L :: "'a \ 'a" where "L f \ T (\, f)" abbreviation R :: "'a \ 'a" where "R f \ T (f, \)" abbreviation tensor (infixr "\" 53) where "f \ g \ T (f, g)" abbreviation assoc ("\[_, _, _]") where "\[a, b, c] \ \ (a, b, c)" text\ In HOL we can just give the definitions of the left and right unitors ``up front'' without any preliminary work. Later we will have to show that these definitions have the right properties. The next two definitions define the values of the unitors when applied to identities; that is, their components as natural transformations. \ definition lunit ("\[_]") where "lunit a \ THE f. \f : \ \ a \ a\ \ \ \ f = (\ \ a) \ inv \[\, \, a]" definition runit ("\[_]") where "runit a \ THE f. \f : a \ \ \ a\ \ f \ \ = (a \ \) \ \[a, \, \]" text\ The next two definitions extend the unitors to all arrows, not just identities. Unfortunately, the traditional symbol \\\ for the left unitor is already reserved for a higher purpose, so we have to make do with a poor substitute. \ abbreviation \ where "\ f \ if arr f then f \ \[dom f] else null" abbreviation \ where "\ f \ if arr f then f \ \[dom f] else null" text\ We now embark upon a development of the consequences of the monoidal category axioms. One of our objectives is to be able to show that an interpretation of the \monoidal_category\ locale induces an interpretation of a locale corresponding to a more traditional definition of monoidal category. Another is to obtain the facts we need to prove the coherence theorem. \ lemma \_in_hom [intro]: shows "\\ : \ \ \ \ \\" using unity_def \_in_hom' by force lemma ide_unity [simp]: shows "ide \" using \_in_hom unity_def by auto lemma tensor_in_hom [simp]: assumes "\f : a \ b\" and "\g : c \ d\" shows "\f \ g : a \ c \ b \ d\" using assms T.preserves_hom CC.arr_char by simp lemma arr_tensor [simp]: assumes "arr f" and "arr g" shows "arr (f \ g)" using assms by simp lemma dom_tensor [simp]: assumes "\f : a \ b\" and "\g : c \ d\" shows "dom (f \ g) = a \ c" using assms by fastforce lemma cod_tensor [simp]: assumes "\f : a \ b\" and "\g : c \ d\" shows "cod (f \ g) = b \ d" using assms by fastforce lemma tensor_preserves_ide [simp]: assumes "ide a" and "ide b" shows "ide (a \ b)" using assms T.preserves_ide CC.ide_char by simp lemma tensor_preserves_iso [simp]: assumes "iso f" and "iso g" shows "iso (f \ g)" using assms by simp lemma inv_tensor [simp]: assumes "iso f" and "iso g" shows "inv (f \ g) = inv f \ inv g" using assms T.preserves_inv by auto lemma interchange: assumes "seq h g" and "seq h' g'" shows "(h \ h') \ (g \ g') = h \ g \ h' \ g'" using assms T.preserves_comp [of "(h, h')" "(g, g')"] by simp lemma \_simp: assumes "arr f" and "arr g" and "arr h" shows "\ (f, g, h) = (f \ g \ h) \ \[dom f, dom g, dom h]" using assms \.is_natural_1 [of "(f, g, h)"] by simp 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 CCC.in_homE by auto lemma arr_assoc [simp]: assumes "ide a" and "ide b" and "ide c" shows "arr \[a, b, c]" using assms assoc_in_hom by simp lemma dom_assoc [simp]: assumes "ide a" and "ide b" and "ide c" shows "dom \[a, b, c] = (a \ b) \ c" using assms assoc_in_hom by simp lemma cod_assoc [simp]: assumes "ide a" and "ide b" and "ide c" shows "cod \[a, b, c] = a \ b \ c" using assms assoc_in_hom by simp lemma assoc_naturality: assumes "arr f0" and "arr f1" and "arr f2" shows "\[cod f0, cod f1, cod f2] \ ((f0 \ f1) \ f2) = (f0 \ f1 \ f2) \ \[dom f0, dom f1, dom f2]" using assms \.naturality by auto lemma iso_assoc [simp]: assumes "ide a" and "ide b" and "ide c" shows "iso \[a, b, c]" using assms \.preserves_iso by simp text\ The next result uses the fact that the functor \L\ is an equivalence (and hence faithful) to show the existence of a unique solution to the characteristic equation used in the definition of a component @{term "\[a]"} of the left unitor. It follows that @{term "\[a]"}, as given by our definition using definite description, satisfies this characteristic equation and is therefore uniquely determined by by \\\, @{term \}, and \\\. \ lemma lunit_char: assumes "ide a" shows "\\[a] : \ \ a \ a\" and "\ \ \[a] = (\ \ a) \ inv \[\, \, a]" and "\!f. \f : \ \ a \ a\ \ \ \ f = (\ \ a) \ inv \[\, \, a]" proof - obtain F \ \ where L: "equivalence_of_categories C C F (\f. \ \ f) \ \" using L.induces_equivalence unity_def by auto interpret L: equivalence_of_categories C C F \\f. \ \ f\ \ \ using L by auto let ?P = "\f. \f : \ \ a \ a\ \ \ \ f = (\ \ a) \ inv \[\, \, a]" have "\(\ \ a) \ inv \[\, \, a] : \ \ \ \ a \ \ \ a\" proof show "\\ \ a : (\ \ \) \ a \ \ \ a\" using assms ide_in_hom by (intro tensor_in_hom, auto) show "\inv \[\, \, a] : \ \ \ \ a \ (\ \ \) \ a\" using assms by auto qed moreover have "ide (\ \ a)" using assms by simp ultimately have "\f. ?P f" using assms L.is_full by blast moreover have "\f f'. ?P f \ ?P f' \ f = f'" proof - fix f f' assume f: "?P f" and f': "?P f'" have "par f f'" using f f' by force show "f = f'" using f f' L.is_faithful [of f f'] by force qed ultimately show "\!f. ?P f" by blast hence 1: "?P \[a]" unfolding lunit_def using theI' [of ?P] by auto show "\\[a] : \ \ a \ a\" using 1 by fast show "\ \ \[a] = (\ \ a) \ inv \[\, \, a]" using 1 by fast qed lemma \_ide_simp: assumes "ide a" shows "\ a = \[a]" using assms lunit_char comp_cod_arr ide_in_hom by (metis in_homE) lemma lunit_in_hom [intro]: assumes "ide a" shows "\\[a] : \ \ a \ a\" using assms lunit_char(1) by blast lemma arr_lunit [simp]: assumes "ide a" shows "arr \[a]" using assms lunit_in_hom by auto lemma dom_lunit [simp]: assumes "ide a" shows "dom \[a] = \ \ a" using assms lunit_in_hom by auto lemma cod_lunit [simp]: assumes "ide a" shows "cod \[a] = a" using assms lunit_in_hom by auto text\ As the right-hand side of the characteristic equation for @{term "\ \ \[a]"} is an isomorphism, and the equivalence functor \L\ reflects isomorphisms, it follows that @{term "\[a]"} is an isomorphism. \ lemma iso_lunit [simp]: assumes "ide a" shows "iso \[a]" using assms lunit_char(2) \_is_iso ide_unity isos_compose iso_assoc iso_inv_iso \_in_hom L.reflects_iso arr_lunit arr_tensor ideD(1) ide_is_iso lunit_in_hom tensor_preserves_iso unity_def by metis text\ To prove that an arrow @{term f} is equal to @{term "\[a]"} we need only show that it is parallel to @{term "\[a]"} and that @{term "\ \ f"} satisfies the same characteristic equation as @{term "\ \ \[a]"} does. \ lemma lunit_eqI: assumes "\f : \ \ a \ a\" and "\ \ f = (\ \ a) \ inv \[\, \, a]" shows "f = \[a]" proof - have "ide a" using assms(1) by auto thus ?thesis using assms lunit_char the1_equality by blast qed text\ The next facts establish the corresponding results for the components of the right unitor. \ lemma runit_char: assumes "ide a" shows "\\[a] : a \ \ \ a\" and "\[a] \ \ = (a \ \) \ \[a, \, \]" and "\!f. \f : a \ \ \ a\ \ f \ \ = (a \ \) \ \[a, \, \]" proof - obtain F \ \ where R: "equivalence_of_categories C C F (\f. f \ \) \ \" using R.induces_equivalence \_in_hom by auto interpret R: equivalence_of_categories C C F \\f. f \ \\ \ \ using R by auto let ?P = "\f. \f : a \ \ \ a\ \ f \ \ = (a \ \) \ \[a, \, \]" have "\(a \ \) \ \[a, \, \] : (a \ \) \ \ \ a \ \\" using assms by fastforce moreover have "ide (a \ \)" using assms by simp ultimately have "\f. ?P f" using assms R.is_full [of a "a \ \" "(a \ \) \ \[a, \, \]"] by blast moreover have "\f f'. ?P f \ ?P f' \ f = f'" proof - fix f f' assume f: "?P f" and f': "?P f'" have "par f f'" using f f' by force show "f = f'" using f f' R.is_faithful [of f f'] by force qed ultimately show "\!f. ?P f" by blast hence 1: "?P \[a]" unfolding runit_def using theI' [of ?P] by fast show "\\[a] : a \ \ \ a\" using 1 by fast show "\[a] \ \ = (a \ \) \ \[a, \, \]" using 1 by fast qed lemma \_ide_simp: assumes "ide a" shows "\ a = \[a]" using assms runit_char [of a] comp_cod_arr by auto lemma runit_in_hom [intro]: assumes "ide a" shows "\\[a] : a \ \ \ a\" using assms runit_char(1) by blast lemma arr_runit [simp]: assumes "ide a" shows "arr \[a]" using assms runit_in_hom by blast lemma dom_runit [simp]: assumes "ide a" shows "dom \[a] = a \ \" using assms runit_in_hom by blast lemma cod_runit [simp]: assumes "ide a" shows "cod \[a] = a" using assms runit_in_hom by blast lemma runit_eqI: assumes "\f : a \ \ \ a\" and "f \ \ = (a \ \) \ \[a, \, \]" shows "f = \[a]" proof - have "ide a" using assms(1) by auto thus ?thesis using assms runit_char the1_equality by blast qed lemma iso_runit [simp]: assumes "ide a" shows "iso \[a]" using assms \_is_iso iso_inv_iso isos_compose ide_is_iso R.preserves_reflects_arr arrI ide_unity iso_assoc runit_char tensor_preserves_iso R.reflects_iso unity_def by metis text\ We can now show that the components of the left and right unitors have the naturality properties required of a natural transformation. \ lemma lunit_naturality: assumes "arr f" shows "\[cod f] \ (\ \ f) = f \ \[dom f]" proof - interpret \': inverse_transformation CCC.comp C T.ToTC T.ToCT \ .. have par: "par (\[cod f] \ (\ \ f)) (f \ \[dom f])" using assms by simp moreover have "\ \ \[cod f] \ (\ \ f) = \ \ f \ \[dom f]" proof - have "\ \ \[cod f] \ (\ \ f) = ((\ \ cod f) \ ((\ \ \) \ f)) \ inv \[\, \, dom f]" using assms interchange [of \ \ "\ \ f" "\[cod f]"] lunit_char(2) \_in_hom unity_def \'.naturality [of "(\, \, f)"] comp_assoc by auto also have "... = ((\ \ f) \ (\ \ dom f)) \ inv \[\, \, dom f]" using assms interchange comp_arr_dom comp_cod_arr \_in_hom by auto also have "... = (\ \ f) \ (\ \ \[dom f])" using assms \_in_hom lunit_char(2) comp_assoc by auto also have "... = \ \ f \ \[dom f]" using assms interchange by simp finally show ?thesis by blast qed ultimately show "\[cod f] \ (\ \ f) = f \ \[dom f]" using L.is_faithful unity_def by metis qed lemma runit_naturality: assumes "arr f" shows "\[cod f] \ (f \ \) = f \ \[dom f]" proof - have "par (\[cod f] \ (f \ \)) (f \ \[dom f])" using assms by force moreover have "\[cod f] \ (f \ \) \ \ = f \ \[dom f] \ \" proof - have "\[cod f] \ (f \ \) \ \ = (cod f \ \) \ \[cod f, \, \] \ ((f \ \) \ \)" using assms interchange [of \ \ "\ \ f" "\[cod f]"] runit_char(2) \_in_hom unity_def comp_assoc by auto also have "... = (cod f \ \) \ (f \ \ \ \) \ \[dom f, \, \]" using assms \.naturality [of "(f, \, \)"] by auto also have "... = ((cod f \ \) \ (f \ \ \ \)) \ \[dom f, \, \]" using comp_assoc by simp also have "... = ((f \ \) \ (dom f \ \)) \ \[dom f, \, \]" using assms \_in_hom interchange comp_arr_dom comp_cod_arr by auto also have "... = (f \ \) \ (\[dom f] \ \)" using assms runit_char comp_assoc by auto also have "... = f \ \[dom f] \ \" using assms interchange by simp finally show ?thesis by blast qed ultimately show "\[cod f] \ (f \ \) = f \ \[dom f]" using R.is_faithful unity_def by metis qed end text\ The following locale is an extension of @{locale monoidal_category} that has the unitors and their inverses, as well as the inverse to the associator, conveniently pre-interpreted. \ locale extended_monoidal_category = monoidal_category + \': inverse_transformation CCC.comp C T.ToTC T.ToCT \ + \: natural_isomorphism C C L map \ + \': inverse_transformation C C L map \ + \: natural_isomorphism C C R map \ + \': inverse_transformation C C R map \ text\ Next we show that an interpretation of @{locale monoidal_category} extends to an interpretation of @{locale extended_monoidal_category} and we arrange for the former locale to inherit from the latter. \ context monoidal_category begin interpretation \': inverse_transformation CCC.comp C T.ToTC T.ToCT \ .. interpretation \: natural_transformation C C L map \ proof - interpret \: transformation_by_components C C L map \\a. \[a]\ using lunit_in_hom lunit_naturality unity_def \_in_hom' L.is_extensional by (unfold_locales, auto) have "\.map = \" using \.is_natural_1 \.is_extensional by auto thus "natural_transformation C C L map \" using \.natural_transformation_axioms by auto qed interpretation \: natural_isomorphism C C L map \ apply unfold_locales using iso_lunit \_ide_simp by simp interpretation \: natural_transformation C C R map \ proof - interpret \: transformation_by_components C C R map \\a. \[a]\ using runit_naturality unity_def \_in_hom' R.is_extensional by (unfold_locales, auto) have "\.map = \" using \.is_natural_1 \.is_extensional by auto thus "natural_transformation C C R map \" using \.natural_transformation_axioms by auto qed interpretation \: natural_isomorphism C C R map \ apply unfold_locales using \_ide_simp by simp lemma induces_extended_monoidal_category: shows "extended_monoidal_category C T \ \" .. end sublocale monoidal_category \ extended_monoidal_category using induces_extended_monoidal_category by auto context monoidal_category begin abbreviation \' where "\' \ \'.map" lemma natural_isomorphism_\': shows "natural_isomorphism CCC.comp C T.ToCT T.ToTC \'" .. abbreviation assoc' ("\\<^sup>-\<^sup>1[_, _, _]") where "\\<^sup>-\<^sup>1[a, b, c] \ inv \[a, b, c]" lemma \'_ide_simp: assumes "ide a" and "ide b" and "ide c" shows "\' (a, b, c) = \\<^sup>-\<^sup>1[a, b, c]" using assms \'.inverts_components inverse_unique by force lemma \'_simp: assumes "arr f" and "arr g" and "arr h" shows "\' (f, g, h) = ((f \ g) \ h) \ \\<^sup>-\<^sup>1[dom f, dom g, dom h]" using assms T.ToTC_simp \'.is_natural_1 \'_ide_simp by force lemma assoc_inv: assumes "ide a" and "ide b" and "ide c" shows "inverse_arrows \[a, b, c] \\<^sup>-\<^sup>1[a, b, c]" using assms inv_is_inverse by simp 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 by auto lemma arr_assoc' [simp]: assumes "ide a" and "ide b" and "ide c" shows "arr \\<^sup>-\<^sup>1[a, b, c]" using assms by simp lemma dom_assoc' [simp]: assumes "ide a" and "ide b" and "ide c" shows "dom \\<^sup>-\<^sup>1[a, b, c] = a \ b \ c" using assms by simp lemma cod_assoc' [simp]: assumes "ide a" and "ide b" and "ide c" shows "cod \\<^sup>-\<^sup>1[a, b, c] = (a \ b) \ c" using assms by simp lemma comp_assoc_assoc' [simp]: assumes "ide a" and "ide b" and "ide c" shows "\[a, b, c] \ \\<^sup>-\<^sup>1[a, b, c] = a \ (b \ c)" and "\\<^sup>-\<^sup>1[a, b, c] \ \[a, b, c] = (a \ b) \ c" using assms assoc_inv comp_arr_inv comp_inv_arr by auto lemma assoc'_naturality: assumes "arr f0" and "arr f1" and "arr f2" shows "((f0 \ f1) \ f2) \ \\<^sup>-\<^sup>1[dom f0, dom f1, dom f2] = \\<^sup>-\<^sup>1[cod f0, cod f1, cod f2] \ (f0 \ f1 \ f2)" using assms \'.naturality by auto abbreviation \' where "\' \ \'.map" abbreviation lunit' ("\\<^sup>-\<^sup>1[_]") where "\\<^sup>-\<^sup>1[a] \ inv \[a]" lemma \'_ide_simp: assumes "ide a" shows "\'.map a = \\<^sup>-\<^sup>1[a]" using assms \'.inverts_components \_ide_simp inverse_unique by force lemma lunit_inv: assumes "ide a" shows "inverse_arrows \[a] \\<^sup>-\<^sup>1[a]" using assms inv_is_inverse by simp lemma lunit'_in_hom [intro]: assumes "ide a" shows "\\\<^sup>-\<^sup>1[a] : a \ \ \ a\" using assms by auto lemma comp_lunit_lunit' [simp]: assumes "ide a" shows "\[a] \ \\<^sup>-\<^sup>1[a] = a" and "\\<^sup>-\<^sup>1[a] \ \[a] = \ \ a" proof - show "\[a] \ \\<^sup>-\<^sup>1[a] = a" using assms comp_arr_inv lunit_inv by fastforce show "\\<^sup>-\<^sup>1[a] \ \[a] = \ \ a" using assms comp_arr_inv lunit_inv by fastforce qed lemma lunit'_naturality: assumes "arr f" shows "(\ \ f) \ \\<^sup>-\<^sup>1[dom f] = \\<^sup>-\<^sup>1[cod f] \ f" using assms \'.naturality \'_ide_simp by simp abbreviation \' where "\' \ \'.map" abbreviation runit' ("\\<^sup>-\<^sup>1[_]") where "\\<^sup>-\<^sup>1[a] \ inv \[a]" lemma \'_ide_simp: assumes "ide a" shows "\'.map a = \\<^sup>-\<^sup>1[a]" using assms \'.inverts_components \_ide_simp inverse_unique by auto lemma runit_inv: assumes "ide a" shows "inverse_arrows \[a] \\<^sup>-\<^sup>1[a]" using assms inv_is_inverse by simp lemma runit'_in_hom [intro]: assumes "ide a" shows "\\\<^sup>-\<^sup>1[a] : a \ a \ \\" using assms by auto lemma comp_runit_runit' [simp]: assumes "ide a" shows "\[a] \ \\<^sup>-\<^sup>1[a] = a" and "\\<^sup>-\<^sup>1[a] \ \[a] = a \ \" proof - show "\[a] \ \\<^sup>-\<^sup>1[a] = a" using assms runit_inv by fastforce show "\\<^sup>-\<^sup>1[a] \ \[a] = a \ \" using assms runit_inv by fastforce qed lemma runit'_naturality: assumes "arr f" shows "(f \ \) \ \\<^sup>-\<^sup>1[dom f] = \\<^sup>-\<^sup>1[cod f] \ f" using assms \'.naturality \'_ide_simp by simp lemma lunit_commutes_with_L: assumes "ide a" shows "\[\ \ a] = \ \ \[a]" using assms lunit_naturality lunit_in_hom iso_lunit iso_is_section section_is_mono monoE L.preserves_ide arrI cod_lunit dom_lunit seqI unity_def by metis lemma runit_commutes_with_R: assumes "ide a" shows "\[a \ \] = \[a] \ \" using assms runit_naturality runit_in_hom iso_runit iso_is_section section_is_mono monoE R.preserves_ide arrI cod_runit dom_runit seqI unity_def by metis text\ The components of the left and right unitors are related via a ``triangle'' diagram that also involves the associator. The proof follows \cite{Etingof15}, Proposition 2.2.3. \ lemma triangle: assumes "ide a" and "ide b" shows "(a \ \[b]) \ \[a, \, b] = \[a] \ b" proof - text\ We show that the lower left triangle in the following diagram commutes. \ text\ $$\xymatrix{ {@{term "((a \ \) \ \) \ b"}} \ar[rrrr]^{\scriptsize @{term "\[a, \, \] \ b"}} \ar[ddd]_{\scriptsize @{term "\[a \ \, \, b]"}} \ar[drr]_{\scriptsize @{term "(\[a] \ \) \ b"}} && && {@{term "(a \ (\ \ \)) \ b"}} \ar[dll]^{\scriptsize @{term "(a \ \) \ b"}} \ar[ddd]^{\scriptsize @{term "\[a, \ \ \, b]"}} \\ && {@{term "(a \ \) \ b"}} \ar[d]^{\scriptsize @{term "\[a, \, b]"}} \\ && {@{term "a \ \ \ b"}} \\ {@{term "(a \ \) \ \ \ b"}} \ar[urr]^{\scriptsize @{term "\[a] \ \ \ b"}} \ar[drr]_{\scriptsize @{term "\[a, \, \ \ b]"}} && && {@{term "a \ (\ \ \) \ b"}} \ar[ull]_{\scriptsize @{term "a \ \ \ b"}} \ar[dll]^{\scriptsize @{term "a \ \[\, \, b]"}} \\ && {@{term "a \ \ \ \ \ b"}} \ar[uu]^{\scriptsize @{term "a \ \[\ \ b]"}} }$$ \ have *: "(a \ \[\ \ b]) \ \[a, \, \ \ b] = \[a] \ \ \ b" proof - have 1: "((a \ \[\ \ b]) \ \[a, \, \ \ b]) \ \[a \ \, \, b] = (\[a] \ \ \ b) \ \[a \ \, \, b]" proof - have "((a \ \[\ \ b]) \ \[a, \, \ \ b]) \ \[a \ \, \, b] = ((a \ \[\ \ b]) \ (a \ \[\, \, b])) \ \[a, \ \ \, b] \ (\[a, \, \] \ b)" using assms pentagon comp_assoc by auto also have "... = (a \ ((\ \ \[b]) \ \[\, \, b])) \ \[a, \ \ \, b] \ (\[a, \, \] \ b)" using assms interchange lunit_commutes_with_L by simp also have "... = ((a \ (\ \ b)) \ \[a, \ \ \, b]) \ (\[a, \, \] \ b)" using assms lunit_char \_in_hom comp_arr_dom comp_assoc by auto also have "... = (\[a, \, b] \ ((a \ \) \ b)) \ (\[a, \, \] \ b)" using assms \_in_hom assoc_naturality [of a \ b] by fastforce also have "... = \[a, \, b] \ ((\[a] \ \) \ b)" using assms \_in_hom interchange runit_char(2) comp_assoc by auto also have "... = (\[a] \ \ \ b) \ \[a \ \, \, b]" using assms assoc_naturality [of "\[a]" \ b] by simp finally show ?thesis by blast qed show ?thesis proof - have "epi \[a \ \, \, b]" using assms iso_assoc iso_is_retraction retraction_is_epi by simp thus ?thesis using 1 assms epiE [of "\[a \ \, \, b]" "(a \ \[\ \ b]) \ \[a, \, \ \ b]"] by fastforce qed qed text\ In \cite{Etingof15} it merely states that the preceding result suffices ``because any object of \C\ is isomorphic to one of the form @{term "\ \ b"}.'' However, it seems a little bit more involved than that to formally transport the equation \(*)\ along the isomorphism @{term "\[b]"} from @{term "\ \ b"} to @{term b}. \ have "(a \ \[b]) \ \[a, \, b] = ((a \ \[b]) \ (a \ \[\ \ b]) \ (a \ \ \ \\<^sup>-\<^sup>1[b])) \ (a \ \ \ \[b]) \ \[a, \, \ \ b] \ ((a \ \) \ \\<^sup>-\<^sup>1[b])" proof - have "\[a, \, b] = (a \ \ \ \[b]) \ \[a, \, \ \ b] \ ((a \ \) \ \\<^sup>-\<^sup>1[b])" proof - have "(a \ \ \ \[b]) \ \[a, \, \ \ b] \ ((a \ \) \ \\<^sup>-\<^sup>1[b]) = ((a \ \ \ \[b]) \ (a \ \ \ \\<^sup>-\<^sup>1[b])) \ \[a, \, b]" using assms assoc_naturality [of a \ "\\<^sup>-\<^sup>1[b]"] comp_assoc by simp also have "... = \[a, \, b]" using assms inv_is_inverse interchange comp_cod_arr by simp finally show ?thesis by auto qed moreover have "a \ \[b] = (a \ \[b]) \ (a \ \[\ \ b]) \ (a \ \ \ \\<^sup>-\<^sup>1[b])" using assms lunit_commutes_with_L comp_arr_dom interchange by auto ultimately show ?thesis by argo qed also have "... = (a \ \[b]) \ (a \ \[\ \ b]) \ ((a \ \ \ \\<^sup>-\<^sup>1[b]) \ (a \ \ \ \[b])) \ \[a, \, \ \ b] \ ((a \ \) \ \\<^sup>-\<^sup>1[b])" using assms comp_assoc by auto also have "... = (a \ \[b]) \ ((a \ \[\ \ b]) \ \[a, \, \ \ b]) \ ((a \ \) \ \\<^sup>-\<^sup>1[b])" using assms interchange comp_cod_arr comp_assoc by auto also have "... = \[a] \ b" using assms * interchange runit_char(1) comp_arr_dom comp_cod_arr by auto finally show ?thesis by blast qed lemma lunit_tensor_gen: assumes "ide a" and "ide b" and "ide c" shows "(a \ \[b \ c]) \ (a \ \[\, b, c]) = a \ \[b] \ c" proof - text\ We show that the lower right triangle in the following diagram commutes. \ text\ $$\xymatrix{ {@{term "((a \ \) \ b) \ c"}} \ar[rrrr]^{\scriptsize @{term "\[a, \, b] \ c"}} \ar[ddd]_{\scriptsize @{term "\[a \ \, b, c]"}} \ar[drr]_{\scriptsize @{term "\[a] \ b \ c"}} && && {@{term "(a \ (\ \ b)) \ c"}} \ar[dll]^{\scriptsize @{term "(a \ \[b]) \ c"}} \ar[ddd]^{\scriptsize @{term "\[a, \ \ b, c]"}} \\ && {@{term "(a \ b) \ c"}} \ar[d]^{\scriptsize @{term "\[a, b, c]"}} \\ && {@{term "a \ b \ c"}} \\ {@{term "(a \ \) \ b \ c"}} \ar[urr]^{\scriptsize @{term "\[a] \ b \ c"}} \ar[drr]_{\scriptsize @{term "\[a, \, b \ c]"}} && && {@{term "a \ (\ \ b) \ c"}} \ar[ull]_{\scriptsize @{term "a \ \[b] \ c"}} \ar[dll]^{\scriptsize @{term "a \ \[\, b, c]"}} \\ && {@{term "a \ \ \ b \ c"}} \ar[uu]^{\scriptsize @{term "a \ \[b \ c]"}} }$$ \ have "((a \ \[b \ c]) \ (a \ \[\, b, c])) \ (\[a, \ \ b, c] \ (\[a, \, b] \ c)) = ((a \ \[b \ c]) \ \[a, \, b \ c]) \ \[a \ \, b, c]" using assms pentagon comp_assoc by simp also have "... = (\[a] \ (b \ c)) \ \[a \ \, b, c]" using assms triangle by auto also have "... = \[a, b, c] \ ((\[a] \ b) \ c)" using assms assoc_naturality [of "\[a]" b c] by auto also have "... = (\[a, b, c] \ ((a \ \[b]) \ c)) \ (\[a, \, b] \ c)" using assms triangle interchange comp_assoc by auto also have "... = (a \ (\[b] \ c)) \ (\[a, \ \ b, c] \ (\[a, \, b] \ c))" using assms assoc_naturality [of a "\[b]" c] comp_assoc by auto finally have 1: "((a \ \[b \ c]) \ (a \ \[\, b, c])) \ \[a, \ \ b, c] \ (\[a, \, b] \ c) = (a \ (\[b] \ c)) \ \[a, \ \ b, c] \ (\[a, \, b] \ c)" by blast text\ The result follows by cancelling the isomorphism @{term "\[a, \ \ b, c] \ (\[a, \, b] \ c)"} \ have 2: "iso (\[a, \ \ b, c] \ (\[a, \, b] \ c))" using assms isos_compose by simp moreover have "seq ((a \ \[b \ c]) \ (a \ \[\, b, c])) (\[a, \ \ b, c] \ (\[a, \, b] \ c))" using assms by auto moreover have "seq (a \ (\[b] \ c)) (\[a, \ \ b, c] \ (\[a, \, b] \ c))" using assms by auto ultimately show ?thesis using 1 2 assms iso_is_retraction retraction_is_epi epiE [of "\[a, \ \ b, c] \ (\[a, \, b] \ c)" "(a \ \[b \ c]) \ (a \ \[\, b, c])" "a \ \[b] \ c"] by auto qed text\ The following result is quoted without proof as Theorem 7 of \cite{Kelly64} where it is attributed to MacLane \cite{MacLane63}. It also appears as \cite{MacLane71}, Exercise 1, page 161. I did not succeed within a few hours to construct a proof following MacLane's hint. The proof below is based on \cite{Etingof15}, Proposition 2.2.4. \ lemma lunit_tensor': assumes "ide a" and "ide b" shows "\[a \ b] \ \[\, a, b] = \[a] \ b" proof - have "\ \ (\[a \ b] \ \[\, a, b]) = \ \ (\[a] \ b)" using assms interchange [of \ \] lunit_tensor_gen by simp moreover have "par (\[a \ b] \ \[\, a, b]) (\[a] \ b)" using assms by simp ultimately show ?thesis using assms L.is_faithful [of "\[a \ b] \ \[\, a, b]" "\[a] \ b"] unity_def by simp qed lemma lunit_tensor: assumes "ide a" and "ide b" shows "\[a \ b] = (\[a] \ b) \ \\<^sup>-\<^sup>1[\, a, b]" using assms lunit_tensor' invert_side_of_triangle by simp text\ We next show the corresponding result for the right unitor. \ lemma runit_tensor_gen: assumes "ide a" and "ide b" and "ide c" shows "\[a \ b] \ c = ((a \ \[b]) \ c) \ (\[a, b, \] \ c)" proof - text\ We show that the upper right triangle in the following diagram commutes. \ text\ $$\xymatrix{ && {@{term "((a \ b) \ \) \ c"}} \ar[dll]_{\scriptsize @{term "\[a \ b, \, c]"}} \ar[dd]^{\scriptsize @{term "\[a \ b] \ c"}} \ar[drr]^{\scriptsize @{term "\[a, b, \] \ c"}} \\ {@{term "(a \ b) \ \ \ c"}} \ar[ddd]_{\scriptsize @{term "\[a, b, \ \ c]"}} \ar[drr]_{\scriptsize @{term "(a \ b) \ \[c]"}} && && {@{term "(a \ b \ \) \ c"}} \ar[dll]^{\scriptsize @{term "(a \ \[b]) \ c"}} \ar[ddd]^{\scriptsize @{term "\[a, b \ \, c]"}} \\ && {@{term "(a \ b) \ c"}} \ar[d]^{\scriptsize @{term "\[a, b, c]"}} \\ && {@{term "a \ b \ c"}} \\ {@{term "a \ b \ \ \ c"}} \ar[urr]^{\scriptsize @{term "a \ b \ \[c]"}} && && {@{term "a \ (b \ \) \ c"}} \ar[llll]^{\scriptsize @{term "a \ \[b, \, c]"}} \ar[ull]_{\scriptsize @{term "a \ \[b] \ c"}} }$$ \ have "\[a \ b] \ c = ((a \ b) \ \[c]) \ \[a \ b, \, c]" using assms triangle by simp also have "... = (\\<^sup>-\<^sup>1[a, b, c] \ (a \ b \ \[c]) \ \[a, b, \ \ c]) \ \[a \ b, \, c]" using assms assoc_naturality [of a b "\[c]"] comp_arr_dom comp_cod_arr invert_side_of_triangle(1) by force also have "... = \\<^sup>-\<^sup>1[a, b, c] \ (a \ b \ \[c]) \ \[a, b, \ \ c] \ \[a \ b, \, c]" using comp_assoc by force also have "... = \\<^sup>-\<^sup>1[a, b, c] \ ((a \ (\[b] \ c)) \ (a \ \\<^sup>-\<^sup>1[b, \, c])) \ \[a, b, \ \ c] \ \[a \ b, \, c]" using assms triangle [of b c] interchange invert_side_of_triangle(2) by force also have "... = (((a \ \[b]) \ c) \ \\<^sup>-\<^sup>1[a, b \ \, c]) \ (a \ \\<^sup>-\<^sup>1[b, \, c]) \ \[a, b, \ \ c] \ \[a \ b, \, c]" using assms assoc'_naturality [of a "\[b]" c] comp_assoc by force also have "... = ((a \ \[b]) \ c) \ \\<^sup>-\<^sup>1[a, b \ \, c] \ (a \ \\<^sup>-\<^sup>1[b, \, c]) \ \[a, b, \ \ c] \ \[a \ b, \, c]" using comp_assoc by simp also have "... = ((a \ \[b]) \ c) \ (\[a, b, \] \ c)" using assms pentagon invert_side_of_triangle(1) invert_side_of_triangle(1) [of "\[a, b, \ \ c] \ \[a \ b, \, c]" "a \ \[b, \, c]" "\[a, b \ \, c] \ (\[a, b, \] \ c)"] by force finally show ?thesis by blast qed lemma runit_tensor: assumes "ide a" and "ide b" shows "\[a \ b] = (a \ \[b]) \ \[a, b, \]" proof - have "((a \ \[b]) \ \[a, b, \]) \ \ = \[a \ b] \ \" using assms interchange [of \ \] runit_tensor_gen by simp moreover have "par ((a \ \[b]) \ \[a, b, \]) \[a \ b]" using assms by simp ultimately show ?thesis using assms R.is_faithful [of "(a \ \[b]) \ (\[a, b, \])" "\[a \ b]"] unity_def by argo qed lemma runit_tensor': assumes "ide a" and "ide b" shows "\[a \ b] \ \\<^sup>-\<^sup>1[a, b, \] = a \ \[b]" using assms runit_tensor invert_side_of_triangle by force text \ Sometimes inverted forms of the triangle and pentagon axioms are useful. \ lemma triangle': assumes "ide a" and "ide b" shows "(a \ \[b]) = (\[a] \ b) \ \\<^sup>-\<^sup>1[a, \, b]" proof - have "(\[a] \ b) \ \\<^sup>-\<^sup>1[a, \, b] = ((a \ \[b]) \ \[a, \, b]) \ \\<^sup>-\<^sup>1[a, \, b]" using assms triangle by auto also have "... = (a \ \[b])" using assms comp_arr_dom comp_assoc by auto finally show ?thesis by auto qed lemma pentagon': assumes "ide a" and "ide b" and "ide c" and "ide d" shows "((\\<^sup>-\<^sup>1[a, b, c] \ d) \ \\<^sup>-\<^sup>1[a, b \ c, d]) \ (a \ \\<^sup>-\<^sup>1[b, c, d]) = \\<^sup>-\<^sup>1[a \ b, c, d] \ \\<^sup>-\<^sup>1[a, b, c \ d]" proof - have "((\\<^sup>-\<^sup>1[a, b, c] \ d) \ \\<^sup>-\<^sup>1[a, b \ c, d]) \ (a \ \\<^sup>-\<^sup>1[b, c, d]) = inv ((a \ \[b, c, d]) \ (\[a, b \ c, d] \ (\[a, b, c] \ d)))" using assms isos_compose inv_comp by simp also have "... = inv (\[a, b, c \ d] \ \[a \ b, c, d])" using assms pentagon by auto also have "... = \\<^sup>-\<^sup>1[a \ b, c, d] \ \\<^sup>-\<^sup>1[a, b, c \ d]" using assms inv_comp by simp finally show ?thesis by auto qed text\ The following non-obvious fact is Corollary 2.2.5 from \cite{Etingof15}. The statement that @{term "\[\] = \[\]"} is Theorem 6 from \cite{Kelly64}. MacLane \cite{MacLane71} does not show this, but assumes it as an axiom. \ lemma unitor_coincidence: shows "\[\] = \" and "\[\] = \" proof - have "\[\] \ \ = (\ \ \[\]) \ \[\, \, \]" using lunit_tensor' [of \ \] lunit_commutes_with_L [of \] by simp moreover have "\[\] \ \ = (\ \ \[\]) \ \[\, \, \]" using triangle [of \ \] by simp moreover have "\ \ \ = (\ \ \[\]) \ \[\, \, \]" using lunit_char comp_arr_dom \_in_hom comp_assoc by auto ultimately have "\[\] \ \ = \ \ \ \ \[\] \ \ = \ \ \" by argo moreover have "par \[\] \ \ par \[\] \" using \_in_hom by force ultimately have 1: "\[\] = \ \ \[\] = \" using R.is_faithful unity_def by metis show "\[\] = \" using 1 by auto show "\[\] = \" using 1 by auto qed lemma \_triangle: shows "\ \ \ = (\ \ \) \ \[\, \, \]" and "(\ \ \) \ \\<^sup>-\<^sup>1[\, \, \] = \ \ \" using triangle [of \ \] triangle' [of \ \] unitor_coincidence by auto text\ The only isomorphism that commutes with @{term \} is @{term \}. \ lemma iso_commuting_with_\_equals_\: assumes "\f : \ \ \\" and "iso f" and "f \ \ = \ \ (f \ f)" shows "f = \" proof - have 1: "f \ f = f \ \" proof - have "f \ f = (inv \ \ \) \ (f \ f)" using assms \_in_hom \_is_iso inv_is_inverse comp_inv_arr comp_cod_arr [of "f \ f"] by force also have "... = (inv \ \ f) \ \" using assms \_is_iso inv_is_inverse comp_assoc by force also have "... = ((f \ \) \ inv \) \ \" using assms unitor_coincidence runit'_naturality [of f] by auto also have "... = (f \ \) \ (inv \ \ \)" using comp_assoc by force also have "... = f \ \" using assms \_in_hom \_is_iso inv_is_inverse comp_inv_arr comp_arr_dom [of "f \ \" "\ \ \"] by force finally show ?thesis by blast qed moreover have "(f \ f) \ (inv f \ \) = \ \ f \ (f \ \) \ (inv f \ \) = \ \ \" using assms interchange comp_arr_inv inv_is_inverse comp_arr_dom by auto ultimately have "\ \ f = \ \ \" by metis moreover have "par f \" using assms by auto ultimately have "f = \" using L.is_faithful unity_def by metis thus ?thesis using 1 by blast qed end text\ We now show that the unit \\\ of a monoidal category is unique up to a unique isomorphism (Proposition 2.2.6 of \cite{Etingof15}). \ locale monoidal_category_with_alternate_unit = monoidal_category C T \ \ + C\<^sub>1: monoidal_category C T \ \\<^sub>1 for C :: "'a comp" (infixr "\" 55) and T :: "'a * 'a \ 'a" and \ :: "'a * 'a * 'a \ 'a" and \ :: 'a and \\<^sub>1 :: 'a begin no_notation C\<^sub>1.tensor (infixr "\" 53) no_notation C\<^sub>1.unity ("\") no_notation C\<^sub>1.lunit ("\[_]") no_notation C\<^sub>1.runit ("\[_]") no_notation C\<^sub>1.assoc ("\[_, _, _]") no_notation C\<^sub>1.assoc' ("\\<^sup>-\<^sup>1[_, _, _]") notation C\<^sub>1.tensor (infixr "\\<^sub>1" 53) notation C\<^sub>1.unity ("\\<^sub>1") notation C\<^sub>1.lunit ("\\<^sub>1[_]") notation C\<^sub>1.runit ("\\<^sub>1[_]") notation C\<^sub>1.assoc ("\\<^sub>1[_, _, _]") notation C\<^sub>1.assoc' ("\\<^sub>1\<^sup>-\<^sup>1[_, _, _]") definition i where "i \ \[\\<^sub>1] \ inv \\<^sub>1[\]" lemma iso_i: shows "\i : \ \ \\<^sub>1\" and "iso i" proof - show "\i : \ \ \\<^sub>1\" using C\<^sub>1.iso_runit inv_in_hom i_def by auto show "iso i" using iso_lunit C\<^sub>1.iso_runit isos_compose i_def by simp qed text\ The following is Exercise 2.2.7 of \cite{Etingof15}. \ lemma i_maps_\_to_\\<^sub>1: shows "i \ \ = \\<^sub>1 \ (i \ i)" proof - have 1: "inv \\<^sub>1[\] \ \ = (\\<^sub>1[\] \ \[\\<^sub>1]) \ (inv \\<^sub>1[\] \ inv \\<^sub>1[\])" proof - have "\ \ (\\<^sub>1[\] \ \\<^sub>1[\]) = \\<^sub>1[\] \ (\\<^sub>1[\] \ \[\\<^sub>1])" proof - text \ $$\xymatrix{ && {@{term[source=true] "(\ \ \\<^sub>1) \ \ \ \\<^sub>1"}} \ar[dddll]_{\scriptsize @{term[source=true] "\\<^sub>1[\] \ \\<^sub>1[\]"}} \ar[dd]^{\scriptsize @{term[source=true] "\\<^sub>1[\] \ \ \ \\<^sub>1"}} \ar[dddrr]^{\scriptsize @{term[source=true] "\\<^sub>1[\] \ \[\\<^sub>1]"}} \\ \\ && {@{term[source=true] "\ \ \ \ \\<^sub>1"}} \ar[dll]^{\scriptsize @{term[source=true] "\ \ \\<^sub>1[\]"}} \ar[drr]_{\scriptsize @{term[source=true] "\ \ \[\\<^sub>1]"}} \ar[dd]^{\scriptsize @{term[source=true] "\\<^sup>-\<^sup>1[\, \, \\<^sub>1]"}} \\ {@{term[source=true] "\ \ \"}} \ar[dddrr]_{\scriptsize @{term[source=true] "\"}} && && {@{term[source=true] "\ \ \\<^sub>1"}} \ar[dddll]^{\scriptsize @{term[source=true] "\\<^sub>1[\]"}} \\ && {@{ term[source=true] "(\ \ \) \ \\<^sub>1"}} \ar[ull]_{\scriptsize @{term[source=true] "\\<^sub>1[\ \ \]"}} \ar[urr]^{\scriptsize @{term[source=true] "\ \ \"}} \\ \\ && {@{term[source=true] "\"}} }$$ \ have "\ \ (\\<^sub>1[\] \ \\<^sub>1[\]) = \ \ (\ \ \\<^sub>1[\]) \ (\\<^sub>1[\] \ \ \ \\<^sub>1)" using interchange comp_cod_arr comp_arr_dom by simp also have "... = \ \ (\\<^sub>1[\ \ \] \ \\<^sup>-\<^sup>1[\, \, \\<^sub>1]) \ (\\<^sub>1[\] \ \ \ \\<^sub>1)" using C\<^sub>1.runit_tensor' by auto also have "... = (\ \ \\<^sub>1[\ \ \]) \ \\<^sup>-\<^sup>1[\, \, \\<^sub>1] \ (\\<^sub>1[\] \ \ \ \\<^sub>1)" using comp_assoc by auto also have "... = (\\<^sub>1[\] \ (\ \ \\<^sub>1)) \ \\<^sup>-\<^sup>1[\, \, \\<^sub>1] \ (\\<^sub>1[\] \ \ \ \\<^sub>1)" using C\<^sub>1.runit_naturality [of \] \_in_hom by fastforce also have "... = \\<^sub>1[\] \ ((\ \ \\<^sub>1) \ \\<^sup>-\<^sup>1[\, \, \\<^sub>1]) \ (\\<^sub>1[\] \ \ \ \\<^sub>1)" using comp_assoc by auto also have "... = \\<^sub>1[\] \ (\ \ \[\\<^sub>1]) \ (\\<^sub>1[\] \ \ \ \\<^sub>1)" using lunit_tensor lunit_commutes_with_L unitor_coincidence by simp also have "... = \\<^sub>1[\] \ (\\<^sub>1[\] \ \[\\<^sub>1])" using interchange comp_arr_dom comp_cod_arr by simp finally show ?thesis by blast qed moreover have "seq \ (\\<^sub>1[\] \ \\<^sub>1[\]) \ seq \\<^sub>1[\] (\\<^sub>1[\] \ \[\\<^sub>1])" using \_in_hom by fastforce moreover have "iso \\<^sub>1[\] \ iso (\\<^sub>1[\] \ \\<^sub>1[\])" using C\<^sub>1.iso_runit tensor_preserves_iso by force ultimately show ?thesis using invert_opposite_sides_of_square inv_tensor by metis qed have 2: "\[\\<^sub>1] \ (\\<^sub>1[\] \ \[\\<^sub>1]) = \\<^sub>1 \ (\[\\<^sub>1] \ \[\\<^sub>1])" proof - text \ $$\xymatrix{ && {@{term[source=true] "(\ \ \\<^sub>1) \ (\ \ \\<^sub>1)"}} \ar[dddll]_{\scriptsize @{term[source=true] "\[\\<^sub>1] \ \[\\<^sub>1]"}} \ar[dd]^{\scriptsize @{term[source=true] "(\ \ \\<^sub>1) \ \[\\<^sub>1]"}} \ar[dddrr]^{\scriptsize @{term[source=true] "\\<^sub>1[\] \ \[\\<^sub>1]"}} \\ \\ && {@{term[source=true] "(\ \ \\<^sub>1) \ \\<^sub>1"}} \ar[dll]^{\scriptsize @{term[source=true] "\[\\<^sub>1] \ \\<^sub>1"}} \ar[drr]_{\scriptsize @{term[source=true] "\\<^sub>1[\] \ \\<^sub>1"}} \ar[dd]^{\scriptsize @{term[source=true] "\[\, \\<^sub>1, \\<^sub>1]"}} \\ {@{term[source=true] "\\<^sub>1 \ \\<^sub>1"}} \ar[dddrr]_{\scriptsize @{term[source=true] "\\<^sub>1"}} && && {@{term[source=true] "\ \ \\<^sub>1"}} \ar[dddll]^{\scriptsize @{term[source=true] "\[\\<^sub>1]"}} \\ && {@{term[source=true] "\ \ \\<^sub>1 \ \\<^sub>1"}} \ar[ull]_{\scriptsize @{term[source=true] "\[\\<^sub>1 \ \\<^sub>1]"}} \ar[urr]^{\scriptsize @{term[source=true] "\ \ \\<^sub>1"}} \\ \\ && {@{term[source=true] "\\<^sub>1"}} }$$ \ have "\[\\<^sub>1] \ (\\<^sub>1[\] \ \[\\<^sub>1]) = \[\\<^sub>1] \ (\\<^sub>1[\] \ \\<^sub>1) \ ((\ \ \\<^sub>1) \ \[\\<^sub>1])" using interchange comp_arr_dom comp_cod_arr by force also have "... = \[\\<^sub>1] \ ((\ \ \\<^sub>1) \ \[\, \\<^sub>1, \\<^sub>1]) \ ((\ \ \\<^sub>1) \ \[\\<^sub>1])" using C\<^sub>1.runit_tensor C\<^sub>1.unitor_coincidence C\<^sub>1.runit_commutes_with_R by simp also have "... = (\[\\<^sub>1] \ (\ \ \\<^sub>1)) \ \[\, \\<^sub>1, \\<^sub>1] \ ((\ \ \\<^sub>1) \ \[\\<^sub>1])" using comp_assoc by fastforce also have "... = (\\<^sub>1 \ \[\\<^sub>1 \ \\<^sub>1]) \ \[\, \\<^sub>1, \\<^sub>1] \ ((\ \ \\<^sub>1) \ \[\\<^sub>1])" using lunit_naturality [of \\<^sub>1] C\<^sub>1.\_in_hom lunit_commutes_with_L by fastforce also have "... = \\<^sub>1 \ (\[\\<^sub>1 \ \\<^sub>1] \ \[\, \\<^sub>1, \\<^sub>1]) \ ((\ \ \\<^sub>1) \ \[\\<^sub>1])" using comp_assoc by force also have "... = \\<^sub>1 \ (\[\\<^sub>1] \ \\<^sub>1) \ ((\ \ \\<^sub>1) \ \[\\<^sub>1])" using lunit_tensor' by auto also have "... = \\<^sub>1 \ (\[\\<^sub>1] \ \[\\<^sub>1])" using interchange comp_arr_dom comp_cod_arr by simp finally show ?thesis by blast qed show ?thesis proof - text \ $$\xymatrix{ {@{term[source=true] "\\<^sub>1 \ \\<^sub>1"}} \ar[dd]_{\scriptsize @{term "\\<^sub>1"}} && {@{term[source=true] "(\ \ \\<^sub>1) \ (\ \ \\<^sub>1)"}} \ar[ll]_{\scriptsize @{term[source=true] "\[\\<^sub>1] \ \[\\<^sub>1]"}} \ar[dd]^{\scriptsize @{term[source=true] "\\<^sub>1[\] \ \[\\<^sub>1]"}} \ar[rr]^{\scriptsize @{term[source=true] "\\<^sub>1[\] \ \\<^sub>1[\]"}} && {@{term[source=true] "\\<^sub>1 \ \\<^sub>1"}} \ar[dd]^{\scriptsize @{term[source=true] "\"}} \\ \\ {@{term[source=true] "\\<^sub>1"}} && {@{term[source=true] "\ \ \\<^sub>1"}} \ar[ll]_{\scriptsize @{term[source=true] "\[\\<^sub>1]"}} \ar[rr]^{\scriptsize @{term[source=true] "\\<^sub>1[\]"}} && {@{term[source=true] "\"}} }$$ \ have "i \ \ = \[\\<^sub>1] \ inv \\<^sub>1[\] \ \" using i_def comp_assoc by auto also have "... = (\[\\<^sub>1] \ (\\<^sub>1[\] \ \[\\<^sub>1])) \ (inv \\<^sub>1[\] \ inv \\<^sub>1[\])" using 1 comp_assoc by simp also have "... = \\<^sub>1 \ (\[\\<^sub>1] \ \[\\<^sub>1]) \ (inv \\<^sub>1[\] \ inv \\<^sub>1[\])" using 2 comp_assoc by fastforce also have "... = \\<^sub>1 \ (i \ i)" using interchange i_def by simp finally show ?thesis by blast qed qed lemma inv_i_iso_\: assumes "\f : \ \ \\<^sub>1\" and "iso f" and "f \ \ = \\<^sub>1 \ (f \ f)" shows "\inv i \ f : \ \ \\" and "iso (inv i \ f)" and "(inv i \ f) \ \ = \ \ (inv i \ f \ inv i \ f)" proof - show 1: "\inv i \ f : \ \ \\" using assms iso_i inv_in_hom by blast show "iso (inv i \ f)" using assms 1 iso_i inv_in_hom by (intro isos_compose, auto) show "(inv i \ f) \ \ = \ \ (inv i \ f \ inv i \ f)" proof - have "(inv i \ f) \ \ = (inv i \ \\<^sub>1) \ (f \ f)" using assms iso_i comp_assoc by auto also have "... = (\ \ (inv i \ inv i)) \ (f \ f)" using assms iso_i invert_opposite_sides_of_square inv_tensor \_in_hom C\<^sub>1.\_in_hom tensor_in_hom tensor_preserves_iso inv_in_hom i_maps_\_to_\\<^sub>1 unity_def seqI' by metis also have "... = \ \ (inv i \ f \ inv i \ f)" using assms 1 iso_i interchange comp_assoc by fastforce finally show ?thesis by blast qed qed lemma unit_unique_upto_unique_iso: shows "\!f. \f : \ \ \\<^sub>1\ \ iso f \ f \ \ = \\<^sub>1 \ (f \ f)" proof show "\i : \ \ \\<^sub>1\ \ iso i \ i \ \ = \\<^sub>1 \ (i \ i)" using iso_i i_maps_\_to_\\<^sub>1 by auto show "\f. \f : \ \ \\<^sub>1\ \ iso f \ f \ \ = \\<^sub>1 \ (f \ f) \ f = i" proof - fix f assume f: "\f : \ \ \\<^sub>1\ \ iso f \ f \ \ = \\<^sub>1 \ (f \ f)" have "inv i \ f = \" using f inv_i_iso_\ iso_commuting_with_\_equals_\ by blast hence "ide (C (inv i) f)" using iso_i by simp thus "f = i" using section_retraction_of_iso(2) [of "inv i" f] inverse_arrow_unique inv_is_inverse iso_i by blast qed qed end section "Elementary Monoidal Category" text\ Although the economy of data assumed by @{locale monoidal_category} is useful for general results, to establish interpretations it is more convenient to work with a traditional definition of monoidal category. The following locale provides such a definition. It permits a monoidal category to be specified by giving the tensor product and the components of the associator and unitors, which are required only to satisfy elementary conditions that imply functoriality and naturality, without having to worry about extensionality or formal interpretations for the various functors and natural transformations. \ locale elementary_monoidal_category = category C 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" ("\[_, _, _]") + assumes ide_unity [simp]: "ide \" and iso_lunit: "ide a \ iso \[a]" and iso_runit: "ide a \ iso \[a]" and iso_assoc: "\ ide a; ide b; ide c \ \ iso \[a, b, c]" and tensor_in_hom [simp]: "\ \f : a \ b\; \g : c \ d\ \ \ \f \ g : a \ c \ b \ d\" and tensor_preserves_ide: "\ ide a; ide b \ \ ide (a \ b)" and interchange: "\ seq g f; seq g' f' \ \ (g \ g') \ (f \ f') = g \ f \ g' \ f'" and lunit_in_hom [simp]: "ide a \ \\[a] : \ \ a \ a\" and lunit_naturality: "arr f \ \[cod f] \ (\ \ f) = f \ \[dom f]" and runit_in_hom [simp]: "ide a \ \\[a] : a \ \ \ a\" and runit_naturality: "arr f \ \[cod f] \ (f \ \) = f \ \[dom f]" and assoc_in_hom [simp]: "\ ide a; ide b; ide c \ \ \\[a, b, c] : (a \ b) \ c \ a \ b \ c\" and assoc_naturality: "\ arr f0; arr f1; arr f2 \ \ \[cod f0, cod f1, cod f2] \ ((f0 \ f1) \ f2) = (f0 \ (f1 \ f2)) \ \[dom f0, dom f1, dom f2]" and triangle: "\ ide a; ide b \ \ (a \ \[b]) \ \[a, \, b] = \[a] \ b" and pentagon: "\ ide a; ide b; ide c; ide d \ \ (a \ \[b, c, d]) \ \[a, b \ c, d] \ (\[a, b, c] \ d) = \[a, b, c \ d] \ \[a \ b, c, d]" text\ An interpretation for the \monoidal_category\ locale readily induces an interpretation for the \elementary_monoidal_category\ locale. \ context monoidal_category begin lemma induces_elementary_monoidal_category: shows "elementary_monoidal_category C tensor \ lunit runit assoc" using \_in_hom iso_assoc tensor_preserves_ide assoc_in_hom tensor_in_hom assoc_naturality lunit_naturality runit_naturality lunit_in_hom runit_in_hom iso_lunit iso_runit interchange pentagon triangle by unfold_locales auto end context elementary_monoidal_category begin interpretation CC: product_category C C .. interpretation CCC: product_category C CC.comp .. definition T :: "'a * 'a \ 'a" where "T f \ if CC.arr f then (fst f \ snd f) else null" lemma T_simp [simp]: assumes "arr f" and "arr g" shows "T (f, g) = f \ g" using assms T_def by simp lemma arr_tensor [simp]: assumes "arr f" and "arr g" shows "arr (f \ g)" using assms tensor_in_hom by blast lemma dom_tensor [simp]: assumes "arr f" and "arr g" shows "dom (f \ g) = dom f \ dom g" using assms tensor_in_hom by blast lemma cod_tensor [simp]: assumes "arr f" and "arr g" shows "cod (f \ g) = cod f \ cod g" using assms tensor_in_hom by blast interpretation T: binary_endofunctor C T using tensor_in_hom interchange T_def apply unfold_locales apply auto[4] by (elim CC.seqE, auto) lemma binary_endofunctor_T: shows "binary_endofunctor C T" .. 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 definition \ where "\ f \ if CCC.arr f then (fst f \ (fst (snd f) \ snd (snd f))) \ \[dom (fst f), dom (fst (snd f)), dom (snd (snd f))] else null" lemma \_ide_simp [simp]: assumes "ide a" and "ide b" and "ide c" shows "\ (a, b, c) = \[a, b, c]" unfolding \_def using assms assoc_in_hom comp_cod_arr by (metis CC.arrI CCC.arrI fst_conv ide_char in_homE snd_conv) lemma arr_assoc [simp]: assumes "ide a" and "ide b" and "ide c" shows "arr \[a, b, c]" using assms assoc_in_hom by blast lemma dom_assoc [simp]: assumes "ide a" and "ide b" and "ide c" shows "dom \[a, b, c] = (a \ b) \ c" using assms assoc_in_hom by blast lemma cod_assoc [simp]: assumes "ide a" and "ide b" and "ide c" shows "cod \[a, b, c] = a \ b \ c" using assms assoc_in_hom by blast interpretation \: natural_isomorphism CCC.comp C T.ToTC T.ToCT \ proof - interpret \: transformation_by_components CCC.comp C T.ToTC T.ToCT \ apply unfold_locales unfolding \_def T.ToTC_def T.ToCT_def T_def using comp_arr_dom comp_cod_arr assoc_naturality by simp_all interpret \: natural_isomorphism CCC.comp C T.ToTC T.ToCT \.map using iso_assoc \.map_simp_ide assoc_in_hom tensor_preserves_ide \_def by (unfold_locales, auto) have "\ = \.map" using assoc_naturality \_def comp_cod_arr T.ToTC_def T_def \.map_def by auto thus "natural_isomorphism CCC.comp C T.ToTC T.ToCT \" using \.natural_isomorphism_axioms by simp qed interpretation \': inverse_transformation CCC.comp C T.ToTC T.ToCT \ .. interpretation L: "functor" C C \\f. T (\, f)\ using T.fixing_ide_gives_functor_1 by auto interpretation R: "functor" C C \\f. T (f, \)\ using T.fixing_ide_gives_functor_2 by auto interpretation \: natural_isomorphism C C \\f. T (\, f)\ map \\f. if arr f then f \ \[dom f] else null\ proof - interpret \: transformation_by_components C C \\f. T (\, f)\ map \\a. \[a]\ using lunit_naturality by (unfold_locales, auto) interpret \: natural_isomorphism C C \\f. T (\, f)\ map \.map using iso_lunit by (unfold_locales, simp) have "\.map = (\f. if arr f then f \ \[dom f] else null)" using \.map_def lunit_naturality by fastforce thus "natural_isomorphism C C (\f. T (\, f)) map (\f. if arr f then f \ \[dom f] else null)" using \.natural_isomorphism_axioms by force qed interpretation \: natural_isomorphism C C \\f. T (f, \)\ map \\f. if arr f then f \ \[dom f] else null\ proof - interpret \: transformation_by_components C C \\f. T (f, \)\ map \\a. \[a]\ using runit_naturality by (unfold_locales, auto) interpret \: natural_isomorphism C C \\f. T (f, \)\ map \.map using iso_runit \.map_simp_ide by (unfold_locales, simp) have "(\f. if arr f then f \ \[dom f] else null) = \.map" using \.map_def runit_naturality T_simp by fastforce thus "natural_isomorphism C C (\f. T (f, \)) map (\f. if arr f then f \ \[dom f] else null)" using \.natural_isomorphism_axioms by force qed text\ The endofunctors \\f. T (\, f)\ and \\f. T (f, \)\ are equivalence functors, due to the existence of the unitors. \ interpretation L: equivalence_functor C C \\f. T (\, f)\ proof - interpret endofunctor C \\f. T (\, f)\ .. show "equivalence_functor C C (\f. T (\, f))" using isomorphic_to_identity_is_equivalence \.natural_isomorphism_axioms by simp qed interpretation R: equivalence_functor C C \\f. T (f, \)\ proof - interpret endofunctor C \\f. T (f, \)\ .. show "equivalence_functor C C (\f. T (f, \))" using isomorphic_to_identity_is_equivalence \.natural_isomorphism_axioms by simp qed text\ To complete an interpretation of the @{locale "monoidal_category"} locale, we define @{term "\ \ \[\]"}. We could also have chosen @{term "\ \ \[\]"} as the two are equal, though to prove that requires some work yet. \ definition \ where "\ \ \[\]" lemma \_in_hom: shows "\\ : \ \ \ \ \\" using lunit_in_hom \_def by simp lemma induces_monoidal_category: shows "monoidal_category C T \ \" proof - have 1: "\\ : \ \ \ \ \\" using lunit_in_hom \_def by simp interpret L: equivalence_functor C C \\f. T (cod \, f)\ proof - have "(\f. T (\, f)) = (\f. T (cod \, f))" using 1 by fastforce thus "equivalence_functor C C (\f. T (cod \, f))" using L.equivalence_functor_axioms T_def by simp qed interpret R: equivalence_functor C C \\f. T (f, cod \)\ proof - have "(\f. T (f, \)) = (\f. T (f, cod \))" using 1 by fastforce thus "equivalence_functor C C (\f. T (f, cod \))" using R.equivalence_functor_axioms T_def by simp qed show ?thesis proof show "\\ : T (cod \, cod \) \ cod \\" using 1 by fastforce show "iso \" using iso_lunit \_def by simp show "\a b c d. \ ide a; ide b; ide c; ide d \ \ T (a, \ (b, c, d)) \ \ (a, T (b, c), d) \ T (\ (a, b, c), d) = \ (a, b, T (c, d)) \ \ (T (a, b), c, d)" using pentagon tensor_preserves_ide by simp qed qed interpretation MC: monoidal_category C T \ \ using induces_monoidal_category by auto text\ We now show that the notions defined in the interpretation \MC\ agree with their counterparts in the present locale. These facts are needed if we define an interpretation for the @{locale elementary_monoidal_category} locale, use it to obtain the induced interpretation for @{locale monoidal_category}, and then want to transfer facts obtained in the induced interpretation back to the original one. \ lemma \_agreement: shows "MC.unity = \" by (metis MC.unity_def \_def ide_unity in_homE lunit_in_hom) lemma L_agreement: shows "MC.L = (\f. T (\, f))" using \_in_hom MC.unity_def by auto lemma R_agreement: shows "MC.R = (\f. T (f, \))" using \_in_hom MC.unity_def by auto text\ We wish to show that the components of the unitors @{term MC.\} and @{term MC.\} defined in the induced interpretation \MC\ agree with those given by the parameters @{term lunit} and @{term runit} to the present locale. To avoid a lengthy development that repeats work already done in the @{locale monoidal_category} locale, we establish the agreement in a special case and then use the properties already shown for \MC\ to prove the general case. In particular, we first show that @{term "\[\] = MC.lunit MC.unity"} and @{term "\[\] = MC.runit MC.unity"}, from which it follows by facts already proved for @{term MC} that both are equal to @{term \}. We then show that for an arbitrary identity @{term a} the arrows @{term "\[a]"} and @{term "\[a]"} satisfy the equations that uniquely characterize the components @{term "MC.lunit a"} and @{term "MC.runit a"}, respectively, and are therefore equal to those components. \ lemma unitor_coincidence: shows "\[\] = \" and "\[\] = \" proof - have "\[\] = MC.runit MC.unity" proof (intro MC.runit_eqI) show 1: "\\[\] : MC.tensor MC.unity MC.unity \ MC.unity\" using runit_in_hom \_in_hom MC.unity_def by auto show "MC.tensor \[\] MC.unity = MC.tensor MC.unity \ \ MC.assoc MC.unity MC.unity MC.unity" using T_def \_in_hom \_def triangle MC.unity_def by (metis 1 MC.ide_unity T_simp \_ide_simp ideD(1) in_homE) qed moreover have "\[\] = MC.lunit MC.unity" proof (intro MC.lunit_eqI) show "\\[\] : MC.tensor MC.unity MC.unity \ MC.unity\" using lunit_in_hom [of \] \_in_hom MC.unity_def by auto show "MC.tensor MC.unity \[\] = MC.tensor \ MC.unity \ MC.assoc' MC.unity MC.unity MC.unity" using T_def lunit_in_hom \_in_hom \_def MC.\_triangle by argo qed moreover have "MC.lunit MC.unity = \ \ MC.runit MC.unity = \" using MC.unitor_coincidence by simp ultimately have 1: "\[\] = \ \ \[\] = \" by simp show "\[\] = \" using 1 by simp show "\[\] = \" using 1 by simp qed lemma lunit_char: assumes "ide a" shows "\ \ \[a] = (\ \ a) \ MC.assoc' MC.unity MC.unity a" proof - have "(\[\] \ a) \ MC.assoc' MC.unity MC.unity a = ((\ \ \[a]) \ \[\, \, a]) \ MC.assoc' MC.unity MC.unity a" using assms triangle by simp also have "... = \ \ \[a]" using assms MC.assoc_inv comp_arr_inv \_agreement iso_assoc comp_arr_dom comp_assoc by (metis MC.\_is_iso \_ide_simp arr_tensor ideD(1) ide_unity invert_side_of_triangle(2) iso_is_arr triangle unitor_coincidence(2)) finally have "\ \ \[a] = (\[\] \ a) \ MC.assoc' MC.unity MC.unity a" by argo thus "\ \ \[a] = (\ \ a) \ MC.assoc' MC.unity MC.unity a" using unitor_coincidence by auto qed lemma runit_char: assumes "ide a" shows "\[a] \ \ = (a \ \) \ \[a, \, \]" using assms triangle \_def by simp lemma \_agreement: shows "MC.\ = (\f. if arr f then f \ \[dom f] else null)" proof fix f have "\ arr f \ MC.\ f = null" by simp moreover have "arr f \ MC.\ f = f \ \[dom f]" proof - have "\a. ide a \ \[a] = MC.lunit a" using \_agreement T_def lunit_char \_in_hom iso_lunit apply (intro MC.lunit_eqI) apply auto by blast thus ?thesis by (metis ide_dom ext seqE) qed ultimately show "MC.\ f = (if arr f then f \ \[dom f] else null)" by simp qed lemma \_agreement: shows "MC.\ = (\f. if arr f then f \ \[dom f] else null)" proof fix f have "\ arr f \ MC.\ f = null" by simp moreover have "arr f \ MC.\ f = f \ \[dom f]" proof - have "\a. ide a \ \[a] = MC.runit a" using \_agreement T_def runit_char \_in_hom iso_runit apply (intro MC.runit_eqI) apply auto by blast thus ?thesis by (metis ide_dom local.ext seqE) qed ultimately show "MC.\ f = (if arr f then f \ \[dom f] else null)" by simp qed lemma lunit_agreement: assumes "ide a" shows "MC.lunit a = \[a]" using assms comp_cod_arr \_agreement by (metis (no_types, lifting) MC.\_ide_simp ide_char in_homE lunit_in_hom) lemma runit_agreement: assumes "ide a" shows "MC.runit a = \[a]" using assms comp_cod_arr \_agreement by (metis (no_types, lifting) MC.\_ide_simp ide_char in_homE runit_in_hom) end section "Strict Monoidal Category" text\ A monoidal category is \emph{strict} if the components of the associator and unitors are all identities. \ locale strict_monoidal_category = monoidal_category + assumes strict_assoc: "\ ide a0; ide a1; ide a2 \ \ ide \[a0, a1, a2]" and strict_lunit: "ide a \ \[a] = a" and strict_runit: "ide a \ \[a] = a" begin lemma strict_unit: shows "\ = \" using strict_lunit unitor_coincidence(1) by auto lemma tensor_assoc [simp]: assumes "arr f0" and "arr f1" and "arr f2" shows "(f0 \ f1) \ f2 = f0 \ f1 \ f2" proof - have "(f0 \ f1) \ f2 = \[cod f0, cod f1, cod f2] \ ((f0 \ f1) \ f2)" using assms assoc_in_hom [of "cod f0" "cod f1" "cod f2"] strict_assoc comp_cod_arr by auto also have "... = (f0 \ f1 \ f2) \ \[dom f0, dom f1, dom f2]" using assms assoc_naturality by simp also have "... = f0 \ f1 \ f2" using assms assoc_in_hom [of "dom f0" "dom f1" "dom f2"] strict_assoc comp_arr_dom by auto finally show ?thesis by simp qed end section "Opposite Monoidal Category" text\ The \emph{opposite} of a monoidal category has the same underlying category, but the arguments to the tensor product are reversed and the associator is inverted and its arguments reversed. \ locale opposite_monoidal_category = C: monoidal_category C T\<^sub>C \\<^sub>C \ for C :: "'a comp" (infixr "\" 55) and T\<^sub>C :: "'a * 'a \ 'a" and \\<^sub>C :: "'a * 'a * 'a \ 'a" and \ :: 'a begin abbreviation T where "T f \ T\<^sub>C (snd f, fst f)" abbreviation \ where "\ f \ C.\' (snd (snd f), fst (snd f), fst f)" end sublocale opposite_monoidal_category \ monoidal_category C T \ \ proof - interpret T: binary_endofunctor C T using C.T.is_extensional C.CC.seq_char C.interchange by (unfold_locales, auto) interpret ToTC: "functor" C.CCC.comp C T.ToTC using T.functor_ToTC by auto interpret ToCT: "functor" C.CCC.comp C T.ToCT using T.functor_ToCT by auto interpret \: natural_transformation C.CCC.comp C T.ToTC T.ToCT \ using C.\'.is_extensional C.CCC.dom_char C.CCC.cod_char T.ToTC_def T.ToCT_def C.\'_simp C.\'.naturality by (unfold_locales) auto interpret \: natural_isomorphism C.CCC.comp C T.ToTC T.ToCT \ using C.\'.components_are_iso by (unfold_locales, simp) interpret L: equivalence_functor C C \\f. T (C.cod \, f)\ using C.R.equivalence_functor_axioms by simp interpret R: equivalence_functor C C \\f. T (f, C.cod \)\ using C.L.equivalence_functor_axioms by simp show "monoidal_category C T \ \" using C.\_in_hom C.\_is_iso C.unity_def C.pentagon' C.comp_assoc by (unfold_locales) auto qed context opposite_monoidal_category begin lemma lunit_simp: assumes "C.ide a" shows "lunit a = C.runit a" using assms lunit_char C.iso_assoc by (intro C.runit_eqI, auto) lemma runit_simp: assumes "C.ide a" shows "runit a = C.lunit a" using assms runit_char C.iso_assoc by (intro C.lunit_eqI, auto) end section "Monoidal Language" text\ In this section we assume that a category @{term C} is given, and we define a formal syntax of terms constructed from arrows of @{term C} using function symbols that correspond to unity, composition, tensor, the associator and its formal inverse, and the left and right unitors and their formal inverses. We will use this syntax to state and prove the coherence theorem and then to construct the free monoidal category generated by @{term C}. \ locale monoidal_language = C: category C for C :: "'a comp" (infixr "\" 55) begin datatype (discs_sels) 't "term" = Prim 't ("\<^bold>\_\<^bold>\") | Unity ("\<^bold>\") | Tensor "'t term" "'t term" (infixr "\<^bold>\" 53) | Comp "'t term" "'t term" (infixr "\<^bold>\" 55) | Lunit "'t term" ("\<^bold>\\<^bold>[_\<^bold>]") | Lunit' "'t term" ("\<^bold>\\<^sup>-\<^sup>1\<^bold>[_\<^bold>]") | Runit "'t term" ("\<^bold>\\<^bold>[_\<^bold>]") | Runit' "'t term" ("\<^bold>\\<^sup>-\<^sup>1\<^bold>[_\<^bold>]") | Assoc "'t term" "'t term" "'t term" ("\<^bold>\\<^bold>[_, _, _\<^bold>]") | Assoc' "'t term" "'t term" "'t term" ("\<^bold>\\<^sup>-\<^sup>1\<^bold>[_, _, _\<^bold>]") lemma not_is_Tensor_Unity: shows "\ is_Tensor Unity" by simp text\ We define formal domain and codomain functions on terms. \ primrec Dom :: "'a term \ 'a term" where "Dom \<^bold>\f\<^bold>\ = \<^bold>\C.dom f\<^bold>\" | "Dom \<^bold>\ = \<^bold>\" | "Dom (t \<^bold>\ u) = (Dom t \<^bold>\ Dom u)" | "Dom (t \<^bold>\ u) = Dom u" | "Dom \<^bold>\\<^bold>[t\<^bold>] = (\<^bold>\ \<^bold>\ Dom t)" | "Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Dom t" | "Dom \<^bold>\\<^bold>[t\<^bold>] = (Dom t \<^bold>\ \<^bold>\)" | "Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Dom t" | "Dom \<^bold>\\<^bold>[t, u, v\<^bold>] = ((Dom t \<^bold>\ Dom u) \<^bold>\ Dom v)" | "Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = (Dom t \<^bold>\ (Dom u \<^bold>\ Dom v))" primrec Cod :: "'a term \ 'a term" where "Cod \<^bold>\f\<^bold>\ = \<^bold>\C.cod f\<^bold>\" | "Cod \<^bold>\ = \<^bold>\" | "Cod (t \<^bold>\ u) = (Cod t \<^bold>\ Cod u)" | "Cod (t \<^bold>\ u) = Cod t" | "Cod \<^bold>\\<^bold>[t\<^bold>] = Cod t" | "Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = (\<^bold>\ \<^bold>\ Cod t)" | "Cod \<^bold>\\<^bold>[t\<^bold>] = Cod t" | "Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = (Cod t \<^bold>\ \<^bold>\)" | "Cod \<^bold>\\<^bold>[t, u, v\<^bold>] = (Cod t \<^bold>\ (Cod u \<^bold>\ Cod v))" | "Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = ((Cod t \<^bold>\ Cod u) \<^bold>\ Cod v)" text\ A term is a ``formal arrow'' if it is constructed from arrows of @{term[source=true] C} in such a way that composition is applied only to formally composable pairs of terms. \ primrec Arr :: "'a term \ bool" where "Arr \<^bold>\f\<^bold>\ = C.arr f" | "Arr \<^bold>\ = True" | "Arr (t \<^bold>\ u) = (Arr t \ Arr u)" | "Arr (t \<^bold>\ u) = (Arr t \ Arr u \ Dom t = Cod u)" | "Arr \<^bold>\\<^bold>[t\<^bold>] = Arr t" | "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Arr t" | "Arr \<^bold>\\<^bold>[t\<^bold>] = Arr t" | "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Arr t" | "Arr \<^bold>\\<^bold>[t, u, v\<^bold>] = (Arr t \ Arr u \ Arr v)" | "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = (Arr t \ Arr u \ Arr v)" abbreviation Par :: "'a term \ 'a term \ bool" where "Par t u \ Arr t \ Arr u \ Dom t = Dom u \ Cod t = Cod u" abbreviation Seq :: "'a term \ 'a term \ bool" where "Seq t u \ Arr t \ Arr u \ Dom t = Cod u" abbreviation Hom :: "'a term \ 'a term \ 'a term set" where "Hom a b \ { t. Arr t \ Dom t = a \ Cod t = b }" text\ A term is a ``formal identity'' if it is constructed from identity arrows of @{term[source=true] C} and @{term "\<^bold>\"} using only the \\<^bold>\\ operator. \ primrec Ide :: "'a term \ bool" where "Ide \<^bold>\f\<^bold>\ = C.ide f" | "Ide \<^bold>\ = True" | "Ide (t \<^bold>\ u) = (Ide t \ Ide u)" | "Ide (t \<^bold>\ u) = False" | "Ide \<^bold>\\<^bold>[t\<^bold>] = False" | "Ide \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = False" | "Ide \<^bold>\\<^bold>[t\<^bold>] = False" | "Ide \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = False" | "Ide \<^bold>\\<^bold>[t, u, v\<^bold>] = False" | "Ide \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = False" lemma Ide_implies_Arr [simp]: shows "Ide t \ Arr t" by (induct t) auto lemma Arr_implies_Ide_Dom [elim]: shows "Arr t \ Ide (Dom t)" by (induct t) auto lemma Arr_implies_Ide_Cod [elim]: shows "Arr t \ Ide (Cod t)" by (induct t) auto lemma Ide_in_Hom [simp]: shows "Ide t \ t \ Hom t t" by (induct t) auto text\ A formal arrow is ``canonical'' if the only arrows of @{term[source=true] C} used in its construction are identities. \ primrec Can :: "'a term \ bool" where "Can \<^bold>\f\<^bold>\ = C.ide f" | "Can \<^bold>\ = True" | "Can (t \<^bold>\ u) = (Can t \ Can u)" | "Can (t \<^bold>\ u) = (Can t \ Can u \ Dom t = Cod u)" | "Can \<^bold>\\<^bold>[t\<^bold>] = Can t" | "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Can t" | "Can \<^bold>\\<^bold>[t\<^bold>] = Can t" | "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = Can t" | "Can \<^bold>\\<^bold>[t, u, v\<^bold>] = (Can t \ Can u \ Can v)" | "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = (Can t \ Can u \ Can v)" lemma Ide_implies_Can: shows "Ide t \ Can t" by (induct t) auto lemma Can_implies_Arr: shows "Can t \ Arr t" by (induct t) auto text\ We next define the formal inverse of a term. This is only sensible for formal arrows built using only isomorphisms of @{term[source=true] C}; in particular, for canonical formal arrows. \ primrec Inv :: "'a term \ 'a term" where "Inv \<^bold>\f\<^bold>\ = \<^bold>\C.inv f\<^bold>\" | "Inv \<^bold>\ = \<^bold>\" | "Inv (t \<^bold>\ u) = (Inv t \<^bold>\ Inv u)" | "Inv (t \<^bold>\ u) = (Inv u \<^bold>\ Inv t)" | "Inv \<^bold>\\<^bold>[t\<^bold>] = \<^bold>\\<^sup>-\<^sup>1\<^bold>[Inv t\<^bold>]" | "Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = \<^bold>\\<^bold>[Inv t\<^bold>]" | "Inv \<^bold>\\<^bold>[t\<^bold>] = \<^bold>\\<^sup>-\<^sup>1\<^bold>[Inv t\<^bold>]" | "Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = \<^bold>\\<^bold>[Inv t\<^bold>]" | "Inv \<^bold>\\<^bold>[t, u, v\<^bold>] = \<^bold>\\<^sup>-\<^sup>1\<^bold>[Inv t, Inv u, Inv v\<^bold>]" | "Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = \<^bold>\\<^bold>[Inv t, Inv u, Inv v\<^bold>]" lemma Inv_preserves_Ide: shows "Ide t \ Ide (Inv t)" by (induct t) auto lemma Inv_preserves_Can: assumes "Can t" shows "Can (Inv t)" and "Dom (Inv t) = Cod t" and "Cod (Inv t) = Dom t" proof - have 0: "Can t \ Can (Inv t) \ Dom (Inv t) = Cod t \ Cod (Inv t) = Dom t" by (induct t) auto show "Can (Inv t)" using assms 0 by blast show "Dom (Inv t) = Cod t" using assms 0 by blast show "Cod (Inv t) = Dom t" using assms 0 by blast qed lemma Inv_in_Hom [simp]: assumes "Can t" shows "Inv t \ Hom (Cod t) (Dom t)" using assms Inv_preserves_Can Can_implies_Arr by simp lemma Inv_Ide [simp]: assumes "Ide a" shows "Inv a = a" using assms by (induct a) auto lemma Inv_Inv [simp]: assumes "Can t" shows "Inv (Inv t) = t" using assms by (induct t) auto text\ We call a term ``diagonal'' if it is either @{term "\<^bold>\"} or it is constructed from arrows of @{term[source=true] C} using only the \\<^bold>\\ operator associated to the right. Essentially, such terms are lists of arrows of @{term[source=true] C}, where @{term "\<^bold>\"} represents the empty list and \\<^bold>\\ is used as the list constructor. We call them ``diagonal'' because terms can regarded as defining ``interconnection matrices'' of arrows connecting ``inputs'' to ``outputs'', and from this point of view diagonal terms correspond to diagonal matrices. The matrix point of view is suggestive for the extension of the results presented here to the symmetric monoidal and cartesian monoidal cases. \ fun Diag :: "'a term \ bool" where "Diag \<^bold>\ = True" | "Diag \<^bold>\f\<^bold>\ = C.arr f" | "Diag (\<^bold>\f\<^bold>\ \<^bold>\ u) = (C.arr f \ Diag u \ u \ \<^bold>\)" | "Diag _ = False" lemma Diag_TensorE: assumes "Diag (Tensor t u)" shows "\<^bold>\un_Prim t\<^bold>\ = t" and "C.arr (un_Prim t)" and "Diag t" and "Diag u" and "u \ \<^bold>\" proof - have 1: "t = \<^bold>\un_Prim t\<^bold>\ \ C.arr (un_Prim t) \ Diag t \ Diag u \ u \ \<^bold>\" using assms by (cases t; simp; cases u; simp) show "\<^bold>\un_Prim t\<^bold>\ = t" using 1 by simp show "C.arr (un_Prim t)" using 1 by simp show "Diag t" using 1 by simp show "Diag u" using 1 by simp show "u \ \<^bold>\" using 1 by simp qed lemma Diag_implies_Arr [elim]: shows "Diag t \ Arr t" apply (induct t, simp_all) by (simp add: Diag_TensorE) lemma Dom_preserves_Diag [elim]: shows "Diag t \ Diag (Dom t)" apply (induct t, simp_all) proof - fix u v assume I2: "Diag v \ Diag (Dom v)" assume uv: "Diag (u \<^bold>\ v)" show "Diag (Dom u \<^bold>\ Dom v)" proof - have 1: "is_Prim (Dom u) \ C.arr (un_Prim (Dom u)) \ Dom u = \<^bold>\C.dom (un_Prim u)\<^bold>\" using uv by (cases u; simp; cases v, simp_all) have 2: "Diag v \ v \ \<^bold>\ \ \ is_Comp v \ \ is_Lunit' v \ \ is_Runit' v" using uv by (cases u; simp; cases v, simp_all) have "Diag (Dom v) \ Dom v \ \<^bold>\" using 2 I2 by (cases v, simp_all) thus ?thesis using 1 by force qed qed lemma Cod_preserves_Diag [elim]: shows "Diag t \ Diag (Cod t)" apply (induct t, simp_all) proof - fix u v assume I2: "Diag v \ Diag (Cod v)" assume uv: "Diag (u \<^bold>\ v)" show "Diag (Cod u \<^bold>\ Cod v)" proof - have 1: "is_Prim (Cod u) \ C.arr (un_Prim (Cod u)) \ Cod u = \<^bold>\C.cod (un_Prim u)\<^bold>\" using uv by (cases u; simp; cases v; simp) have 2: "Diag v \ v \ \<^bold>\ \ \ is_Comp v \ \ is_Lunit' v \ \ is_Runit' v" using uv by (cases u; simp; cases v; simp) have "Diag (Cod v) \ Cod v \ \<^bold>\" using I2 2 by (cases v, simp_all) thus ?thesis using 1 by force qed qed lemma Inv_preserves_Diag: assumes "Can t" and "Diag t" shows "Diag (Inv t)" proof - have "Can t \ Diag t \ Diag (Inv t)" apply (induct t, simp_all) by (metis (no_types, lifting) Can.simps(1) Inv.simps(1) Inv.simps(2) Diag.simps(3) Inv_Inv Diag_TensorE(1) C.inv_ide) thus ?thesis using assms by blast qed text\ The following function defines the ``dimension'' of a term, which is the number of arrows of @{term C} it contains. For diagonal terms, this is just the length of the term when regarded as a list of arrows of @{term C}. Alternatively, if a term is regarded as defining an interconnection matrix, then the dimension is the number of inputs (or outputs). \ primrec dim :: "'a term \ nat" where "dim \<^bold>\f\<^bold>\ = 1" | "dim \<^bold>\ = 0" | "dim (t \<^bold>\ u) = (dim t + dim u)" | "dim (t \<^bold>\ u) = dim t" | "dim \<^bold>\\<^bold>[t\<^bold>] = dim t" | "dim \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = dim t" | "dim \<^bold>\\<^bold>[t\<^bold>] = dim t" | "dim \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] = dim t" | "dim \<^bold>\\<^bold>[t, u, v\<^bold>] = dim t + dim u + dim v" | "dim \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] = dim t + dim u + dim v" text\ The following function defines a tensor product for diagonal terms. If terms are regarded as lists, this is just list concatenation. If terms are regarded as matrices, this corresponds to constructing a block diagonal matrix. \ fun TensorDiag (infixr "\<^bold>\\<^bold>\\<^bold>\" 53) where "\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u = u" | "t \<^bold>\\<^bold>\\<^bold>\ \<^bold>\ = t" | "\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u = \<^bold>\f\<^bold>\ \<^bold>\ u" | "(t \<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" | "t \<^bold>\\<^bold>\\<^bold>\ u = undefined" lemma TensorDiag_Prim [simp]: assumes "t \ \<^bold>\" shows "\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ t = \<^bold>\f\<^bold>\ \<^bold>\ t" using assms by (cases t, simp_all) lemma TensorDiag_term_Unity [simp]: shows "t \<^bold>\\<^bold>\\<^bold>\ \<^bold>\ = t" by (cases "t = \<^bold>\"; cases t, simp_all) lemma TensorDiag_Diag: assumes "Diag (t \<^bold>\ u)" shows "t \<^bold>\\<^bold>\\<^bold>\ u = t \<^bold>\ u" using assms TensorDiag_Prim by (cases t, simp_all) lemma TensorDiag_preserves_Diag: assumes "Diag t" and "Diag u" shows "Diag (t \<^bold>\\<^bold>\\<^bold>\ u)" and "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom t \<^bold>\\<^bold>\\<^bold>\ Dom u" and "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t \<^bold>\\<^bold>\\<^bold>\ Cod u" proof - have 0: "\u. \ Diag t; Diag u \ \ Diag (t \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom t \<^bold>\\<^bold>\\<^bold>\ Dom u \ Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t \<^bold>\\<^bold>\\<^bold>\ Cod u" apply (induct t, simp_all) proof - fix f :: 'a and u :: "'a term" assume f: "C.arr f" assume u: "Diag u" show "Diag (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = \<^bold>\C.dom f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Dom u \ Cod (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = \<^bold>\C.cod f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Cod u" using u f by (cases u, simp_all) next fix u v w assume I1: "\u. Diag v \ Diag u \ Diag (v \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (v \<^bold>\\<^bold>\\<^bold>\ u) = Dom v \<^bold>\\<^bold>\\<^bold>\ Dom u \ Cod (v \<^bold>\\<^bold>\\<^bold>\ u) = Cod v \<^bold>\\<^bold>\\<^bold>\ Cod u" assume I2: "\u. Diag w \ Diag u \ Diag (w \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (w \<^bold>\\<^bold>\\<^bold>\ u) = Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u \ Cod (w \<^bold>\\<^bold>\\<^bold>\ u) = Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u" assume vw: "Diag (v \<^bold>\ w)" assume u: "Diag u" show "Diag ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) \ Dom ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = (Dom v \<^bold>\ Dom w) \<^bold>\\<^bold>\\<^bold>\ Dom u \ Cod ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = (Cod v \<^bold>\ Cod w) \<^bold>\\<^bold>\\<^bold>\ Cod u" proof - have v: "v = \<^bold>\un_Prim v\<^bold>\ \ Diag v" using vw Diag_implies_Arr Diag_TensorE [of v w] by force have w: "Diag w" using vw by (simp add: Diag_TensorE) have "u = \<^bold>\ \ ?thesis" by (simp add: vw) moreover have "u \ \<^bold>\ \ ?thesis" using u v w I1 I2 Dom_preserves_Diag [of u] Cod_preserves_Diag [of u] by (cases u, simp_all) ultimately show ?thesis by blast qed qed show "Diag (t \<^bold>\\<^bold>\\<^bold>\ u) " using assms 0 by blast show "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom t \<^bold>\\<^bold>\\<^bold>\ Dom u" using assms 0 by blast show "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t \<^bold>\\<^bold>\\<^bold>\ Cod u" using assms 0 by blast qed lemma TensorDiag_in_Hom: assumes "Diag t" and "Diag u" shows "t \<^bold>\\<^bold>\\<^bold>\ u \ Hom (Dom t \<^bold>\\<^bold>\\<^bold>\ Dom u) (Cod t \<^bold>\\<^bold>\\<^bold>\ Cod u)" using assms TensorDiag_preserves_Diag Diag_implies_Arr by simp lemma Dom_TensorDiag: assumes "Diag t" and "Diag u" shows "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom t \<^bold>\\<^bold>\\<^bold>\ Dom u" using assms TensorDiag_preserves_Diag(2) by simp lemma Cod_TensorDiag: assumes "Diag t" and "Diag u" shows "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t \<^bold>\\<^bold>\\<^bold>\ Cod u" using assms TensorDiag_preserves_Diag(3) by simp lemma not_is_Tensor_TensorDiagE: assumes "\ is_Tensor (t \<^bold>\\<^bold>\\<^bold>\ u)" and "Diag t" and "Diag u" and "t \ \<^bold>\" and "u \ \<^bold>\" shows "False" proof - have "\ \ is_Tensor (t \<^bold>\\<^bold>\\<^bold>\ u); Diag t; Diag u; t \ \<^bold>\; u \ \<^bold>\ \ \ False" apply (induct t, simp_all) proof - fix v w assume I2: "\ is_Tensor (w \<^bold>\\<^bold>\\<^bold>\ u) \ Diag w \ w \ \<^bold>\ \ False" assume 1: "\ is_Tensor ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u)" assume vw: "Diag (v \<^bold>\ w)" assume u: "Diag u" assume 2: "u \ \<^bold>\" show "False" proof - have v: "v = \<^bold>\un_Prim v\<^bold>\" using vw Diag_TensorE [of v w] by force have w: "Diag w \ w \ \<^bold>\" using vw Diag_TensorE [of v w] by simp have "(v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u = v \<^bold>\ (w \<^bold>\\<^bold>\\<^bold>\ u)" proof - have "(v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u = v \<^bold>\\<^bold>\\<^bold>\ (w \<^bold>\\<^bold>\\<^bold>\ u)" using u 2 by (cases u, simp_all) also have "... = v \<^bold>\ (w \<^bold>\\<^bold>\\<^bold>\ u)" using u v w I2 TensorDiag_Prim not_is_Tensor_Unity by metis finally show ?thesis by simp qed thus ?thesis using 1 by simp qed qed thus ?thesis using assms by blast qed lemma TensorDiag_assoc: assumes "Diag t" and "Diag u" and "Diag v" shows "(t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - have "\n t u v. \ dim t = n; Diag t; Diag u; Diag v \ \ (t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - fix n show "\t u v. \ dim t = n; Diag t; Diag u; Diag v \ \ (t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof (induction n rule: nat_less_induct) fix n :: nat and t :: "'a term" and u v assume I: "\mt u v. dim t = m \ Diag t \ Diag u \ Diag v \ (t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" assume dim: "dim t = n" assume t: "Diag t" assume u: "Diag u" assume v: "Diag v" show "(t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - have "t = \<^bold>\ \ ?thesis" by simp moreover have "u = \<^bold>\ \ ?thesis" by simp moreover have "v = \<^bold>\ \ ?thesis" by simp moreover have "t \ \<^bold>\ \ u \ \<^bold>\ \ v \ \<^bold>\ \ is_Prim t \ ?thesis" using v by (cases t, simp_all, cases u, simp_all; cases v, simp_all) moreover have "t \ \<^bold>\ \ u \ \<^bold>\ \ v \ \<^bold>\ \ is_Tensor t \ ?thesis" proof (cases t; simp) fix w :: "'a term" and x :: "'a term" assume 1: "u \ \<^bold>\ \ v \ \<^bold>\" assume 2: "t = (w \<^bold>\ x)" show "((w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = (w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - have w: "w = \<^bold>\un_Prim w\<^bold>\" using t 1 2 Diag_TensorE [of w x] by auto have x: "Diag x" using t w 1 2 by (cases w, simp_all) have "((w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = (w \<^bold>\\<^bold>\\<^bold>\ (x \<^bold>\\<^bold>\\<^bold>\ u)) \<^bold>\\<^bold>\\<^bold>\ v" using u v w x 1 2 by (cases u, simp_all) also have "... = (w \<^bold>\ (x \<^bold>\\<^bold>\\<^bold>\ u)) \<^bold>\\<^bold>\\<^bold>\ v" using t w u 1 2 TensorDiag_Prim not_is_Tensor_TensorDiagE Diag_TensorE not_is_Tensor_Unity by metis also have "... = w \<^bold>\\<^bold>\\<^bold>\ ((x \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v)" using u v w x 1 by (cases u, simp_all; cases v, simp_all) also have "... = w \<^bold>\\<^bold>\\<^bold>\ (x \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v))" proof - have "dim x < dim t" using w 2 by (cases w, simp_all) thus ?thesis using u v x dim I by simp qed also have "... = (w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - have 3: "is_Tensor (u \<^bold>\\<^bold>\\<^bold>\ v)" using u v 1 not_is_Tensor_TensorDiagE by auto obtain u' :: "'a term" and v' where uv': "u \<^bold>\\<^bold>\\<^bold>\ v = u' \<^bold>\ v'" using 3 is_Tensor_def by blast thus ?thesis by simp qed finally show ?thesis by simp qed qed moreover have "t = \<^bold>\ \ is_Prim t \ is_Tensor t" using t by (cases t, simp_all) ultimately show ?thesis by blast qed qed qed thus ?thesis using assms by blast qed lemma TensorDiag_preserves_Ide: assumes "Ide t" and "Ide u" and "Diag t" and "Diag u" shows "Ide (t \<^bold>\\<^bold>\\<^bold>\ u)" using assms by (metis (mono_tags, lifting) Arr_implies_Ide_Dom Ide_in_Hom Diag_implies_Arr TensorDiag_preserves_Diag(1) TensorDiag_preserves_Diag(2) mem_Collect_eq) lemma TensorDiag_preserves_Can: assumes "Can t" and "Can u" and "Diag t" and "Diag u" shows "Can (t \<^bold>\\<^bold>\\<^bold>\ u)" proof - have "\u. \ Can t \ Diag t; Can u \ Diag u \ \ Can (t \<^bold>\\<^bold>\\<^bold>\ u)" proof (induct t; simp) show "\x u. C.ide x \ C.arr x \ Can u \ Diag u \ Can (\<^bold>\x\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u)" by (metis Ide.simps(1) Ide.simps(2) Ide_implies_Can Diag.simps(2) TensorDiag_Prim TensorDiag_preserves_Ide Can.simps(3)) show "\t1 t2 u. (\u. Diag t1 \ Can u \ Diag u \ Can (t1 \<^bold>\\<^bold>\\<^bold>\ u)) \ (\u. Diag t2 \ Can u \ Diag u \ Can (t2 \<^bold>\\<^bold>\\<^bold>\ u)) \ Can t1 \ Can t2 \ Diag (t1 \<^bold>\ t2) \ Can u \ Diag u \ Can ((t1 \<^bold>\ t2) \<^bold>\\<^bold>\\<^bold>\ u)" by (metis Diag_TensorE(3) Diag_TensorE(4) TensorDiag_Diag TensorDiag_assoc TensorDiag_preserves_Diag(1)) qed thus ?thesis using assms by blast qed lemma Inv_TensorDiag: assumes "Can t" and "Can u" and "Diag t" and "Diag u" shows "Inv (t \<^bold>\\<^bold>\\<^bold>\ u) = Inv t \<^bold>\\<^bold>\\<^bold>\ Inv u" proof - have "\u. \ Can t \ Diag t; Can u \ Diag u \ \ Inv (t \<^bold>\\<^bold>\\<^bold>\ u) = Inv t \<^bold>\\<^bold>\\<^bold>\ Inv u" proof (induct t, simp_all) fix f u assume f: "C.ide f \ C.arr f" assume u: "Can u \ Diag u" show "Inv (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = \<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Inv u" using f u by (cases u, simp_all) next fix t u v assume I1: "\v. \ Diag t; Can v \ Diag v \ \ Inv (t \<^bold>\\<^bold>\\<^bold>\ v) = Inv t \<^bold>\\<^bold>\\<^bold>\ Inv v" assume I2: "\v. \ Diag u; Can v \ Diag v \ \ Inv (u \<^bold>\\<^bold>\\<^bold>\ v) = Inv u \<^bold>\\<^bold>\\<^bold>\ Inv v" assume tu: "Can t \ Can u \ Diag (t \<^bold>\ u)" have t: "Can t \ Diag t" using tu Diag_TensorE [of t u] by force have u: "Can u \ Diag u" using t tu by (cases t, simp_all) assume v: "Can v \ Diag v" show "Inv ((t \<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v) = (Inv t \<^bold>\ Inv u) \<^bold>\\<^bold>\\<^bold>\ Inv v" proof - have "v = Unity \ ?thesis" by simp moreover have "v \ Unity \ ?thesis" proof - assume 1: "v \ Unity" have "Inv ((t \<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v) = Inv (t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v))" using 1 by (cases v, simp_all) also have "... = Inv t \<^bold>\\<^bold>\\<^bold>\ Inv (u \<^bold>\\<^bold>\\<^bold>\ v)" using t u v I1 TensorDiag_preserves_Diag TensorDiag_preserves_Can Inv_preserves_Diag Inv_preserves_Can by simp also have "... = Inv t \<^bold>\\<^bold>\\<^bold>\ (Inv u \<^bold>\\<^bold>\\<^bold>\ Inv v)" using t u v I2 TensorDiag_preserves_Diag TensorDiag_preserves_Can Inv_preserves_Diag Inv_preserves_Can by simp also have "... = (Inv t \<^bold>\ Inv u) \<^bold>\\<^bold>\\<^bold>\ Inv v" using v 1 by (cases v, simp_all) finally show ?thesis by blast qed ultimately show ?thesis by blast qed qed thus ?thesis using assms by blast qed text\ The following function defines composition for compatible diagonal terms, by ``pushing the composition down'' to arrows of \C\. \ fun CompDiag :: "'a term \ 'a term \ 'a term" (infixr "\<^bold>\\<^bold>\\<^bold>\" 55) where "\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u = u" | "\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\g\<^bold>\ = \<^bold>\f \ g\<^bold>\" | "(u \<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (w \<^bold>\ x) = (u \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\ v \<^bold>\\<^bold>\\<^bold>\ x)" | "t \<^bold>\\<^bold>\\<^bold>\ \<^bold>\ = t" | "t \<^bold>\\<^bold>\\<^bold>\ _ = undefined \<^bold>\ undefined" text\ Note that the last clause above is not relevant to diagonal terms. We have chosen a provably non-diagonal value in order to validate associativity. \ lemma CompDiag_preserves_Diag: assumes "Diag t" and "Diag u" and "Dom t = Cod u" shows "Diag (t \<^bold>\\<^bold>\\<^bold>\ u)" and "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom u" and "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t" proof - have 0: "\u. \ Diag t; Diag u; Dom t = Cod u \ \ Diag (t \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom u \ Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t" proof (induct t, simp_all add: Diag_TensorE) fix f u assume f: "C.arr f" assume u: "Diag u" assume 1: "\<^bold>\C.dom f\<^bold>\ = Cod u" show "Diag (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = Dom u \ Cod (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) = \<^bold>\C.cod f\<^bold>\" using f u 1 by (cases u, simp_all) next fix u v w assume I2: "\u. \ Diag u; Dom w = Cod u \ \ Diag (w \<^bold>\\<^bold>\\<^bold>\ u) \ Dom (w \<^bold>\\<^bold>\\<^bold>\ u) = Dom u \ Cod (w \<^bold>\\<^bold>\\<^bold>\ u) = Cod w" assume vw: "Diag (v \<^bold>\ w)" have v: "Diag v" using vw Diag_TensorE [of v w] by force have w: "Diag w" using vw Diag_TensorE [of v w] by force assume u: "Diag u" assume 1: "(Dom v \<^bold>\ Dom w) = Cod u" show "Diag ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) \ Dom ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = Dom u \ Cod ((v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u) = Cod v \<^bold>\ Cod w" using u v w 1 proof (cases u, simp_all) fix x y assume 2: "u = Tensor x y" have 4: "is_Prim x \ x = \<^bold>\un_Prim x\<^bold>\ \ C.arr (un_Prim x) \ Diag y \ y \ \<^bold>\" using u 2 by (cases x, cases y, simp_all) have 5: "is_Prim v \ v = \<^bold>\un_Prim v\<^bold>\ \ C.arr (un_Prim v) \ Diag w \ w \ \<^bold>\" using v w vw by (cases v, simp_all) have 6: "C.dom (un_Prim v) = C.cod (un_Prim x) \ Dom w = Cod y" using 1 2 4 5 apply (cases u, simp_all) by (metis Cod.simps(1) Dom.simps(1) term.simps(1)) have "(v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u = \<^bold>\un_Prim v \ un_Prim x\<^bold>\ \<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ y" using 2 4 5 6 CompDiag.simps(2) [of "un_Prim v" "un_Prim x"] by simp moreover have "Diag (\<^bold>\un_Prim v \ un_Prim x\<^bold>\ \<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ y)" proof - have "Diag (w \<^bold>\\<^bold>\\<^bold>\ y)" using I2 4 5 6 by simp thus ?thesis using 4 5 6 Diag.simps(3) [of "un_Prim v \ un_Prim x" "(w \<^bold>\\<^bold>\\<^bold>\ y)"] by (cases w; cases y) auto qed ultimately show "Diag (v \<^bold>\\<^bold>\\<^bold>\ x \<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ y) \ Dom (v \<^bold>\\<^bold>\\<^bold>\ x) = Dom x \ Dom (w \<^bold>\\<^bold>\\<^bold>\ y) = Dom y \ Cod (v \<^bold>\\<^bold>\\<^bold>\ x) = Cod v \ Cod (w \<^bold>\\<^bold>\\<^bold>\ y) = Cod w" using 4 5 6 I2 by (metis (full_types) C.cod_comp C.dom_comp Cod.simps(1) CompDiag.simps(2) Dom.simps(1) C.seqI) qed qed show "Diag (t \<^bold>\\<^bold>\\<^bold>\ u)" using assms 0 by blast show "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom u" using assms 0 by blast show "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t" using assms 0 by blast qed lemma CompDiag_in_Hom: assumes "Diag t" and "Diag u" and "Dom t = Cod u" shows "t \<^bold>\\<^bold>\\<^bold>\ u \ Hom (Dom u) (Cod t)" using assms CompDiag_preserves_Diag Diag_implies_Arr by simp lemma Dom_CompDiag: assumes "Diag t" and "Diag u" and "Dom t = Cod u" shows "Dom (t \<^bold>\\<^bold>\\<^bold>\ u) = Dom u" using assms CompDiag_preserves_Diag(2) by simp lemma Cod_CompDiag: assumes "Diag t" and "Diag u" and "Dom t = Cod u" shows "Cod (t \<^bold>\\<^bold>\\<^bold>\ u) = Cod t" using assms CompDiag_preserves_Diag(3) by simp lemma CompDiag_Cod_Diag [simp]: assumes "Diag t" shows "Cod t \<^bold>\\<^bold>\\<^bold>\ t = t" proof - have "Diag t \ Cod t \<^bold>\\<^bold>\\<^bold>\ t = t" using C.comp_cod_arr apply (induct t, auto) by (auto simp add: Diag_TensorE) thus ?thesis using assms by blast qed lemma CompDiag_Diag_Dom [simp]: assumes "Diag t" shows "t \<^bold>\\<^bold>\\<^bold>\ Dom t = t" proof - have "Diag t \ t \<^bold>\\<^bold>\\<^bold>\ Dom t = t" using C.comp_arr_dom apply (induct t, auto) by (auto simp add: Diag_TensorE) thus ?thesis using assms by blast qed lemma CompDiag_Ide_Diag [simp]: assumes "Diag t" and "Ide a" and "Dom a = Cod t" shows "a \<^bold>\\<^bold>\\<^bold>\ t = t" using assms Ide_in_Hom by simp lemma CompDiag_Diag_Ide [simp]: assumes "Diag t" and "Ide a" and "Dom t = Cod a" shows "t \<^bold>\\<^bold>\\<^bold>\ a = t" using assms Ide_in_Hom by auto lemma CompDiag_assoc: assumes "Diag t" and "Diag u" and "Diag v" and "Dom t = Cod u" and "Dom u = Cod v" shows "(t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof - have "\u v. \ Diag t; Diag u; Diag v; Dom t = Cod u; Dom u = Cod v \ \ (t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = t \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" proof (induct t, simp_all) fix f u v assume f: "C.arr f" assume u: "Diag u" assume v: "Diag v" assume 1: "\<^bold>\C.dom f\<^bold>\ = Cod u" assume 2: "Dom u = Cod v" show "(\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = \<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" using C.comp_assoc by (cases u, simp_all; cases v, simp_all) next fix u v w x assume I1: "\u v. \ Diag w; Diag u; Diag v; Dom w = Cod u; Dom u = Cod v \ \ (w \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = w \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" assume I2: "\u v. \ Diag x; Diag u; Diag v; Dom x = Cod u; Dom u = Cod v \ \ (x \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = x \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ v)" assume wx: "Diag (w \<^bold>\ x)" assume u: "Diag u" assume v: "Diag v" assume 1: "(Dom w \<^bold>\ Dom x) = Cod u" assume 2: "Dom u = Cod v" show "((w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v = (w \<^bold>\ x) \<^bold>\\<^bold>\\<^bold>\ u \<^bold>\\<^bold>\\<^bold>\ v" proof - have w: "Diag w" using wx Diag_TensorE by blast have x: "Diag x" using wx Diag_TensorE by blast have "is_Tensor u" using u 1 by (cases u) simp_all thus ?thesis using u v apply (cases u, simp_all, cases v, simp_all) proof - fix u1 u2 v1 v2 assume 3: "u = (u1 \<^bold>\ u2)" assume 4: "v = (v1 \<^bold>\ v2)" show "(w \<^bold>\\<^bold>\\<^bold>\ u1) \<^bold>\\<^bold>\\<^bold>\ v1 = w \<^bold>\\<^bold>\\<^bold>\ u1 \<^bold>\\<^bold>\\<^bold>\ v1 \ (x \<^bold>\\<^bold>\\<^bold>\ u2) \<^bold>\\<^bold>\\<^bold>\ v2 = x \<^bold>\\<^bold>\\<^bold>\ u2 \<^bold>\\<^bold>\\<^bold>\ v2" proof - have "Diag u1 \ Diag u2" using u 3 Diag_TensorE by blast moreover have "Diag v1 \ Diag v2" using v 4 Diag_TensorE by blast ultimately show ?thesis using w x I1 I2 1 2 3 4 by simp qed qed qed qed thus ?thesis using assms by blast qed lemma CompDiag_preserves_Ide: assumes "Ide t" and "Ide u" and "Diag t" and "Diag u" and "Dom t = Cod u" shows "Ide (t \<^bold>\\<^bold>\\<^bold>\ u)" proof - have "\u. \ Ide t; Ide u; Diag t; Diag u; Dom t = Cod u \ \ Ide (CompDiag t u)" by (induct t; simp) thus ?thesis using assms by blast qed lemma CompDiag_preserves_Can: assumes "Can t" and "Can u" and "Diag t" and "Diag u" and "Dom t = Cod u" shows "Can (t \<^bold>\\<^bold>\\<^bold>\ u)" proof - have "\u. \ Can t \ Diag t; Can u \ Diag u; Dom t = Cod u \ \ Can (t \<^bold>\\<^bold>\\<^bold>\ u)" proof (induct t, simp_all) fix t u v assume I1: "\v. \ Diag t; Can v \ Diag v; Dom t = Cod v \ \ Can (t \<^bold>\\<^bold>\\<^bold>\ v)" assume I2: "\v. \ Diag u; Can v \ Diag v; Dom u = Cod v \ \ Can (u \<^bold>\\<^bold>\\<^bold>\ v)" assume tu: "Can t \ Can u \ Diag (t \<^bold>\ u)" have t: "Can t \ Diag t" using tu Diag_TensorE by blast have u: "Can u \ Diag u" using tu Diag_TensorE by blast assume v: "Can v \ Diag v" assume 1: "(Dom t \<^bold>\ Dom u) = Cod v" show "Can ((t \<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v)" proof - have 2: "(Dom t \<^bold>\ Dom u) = Cod v" using 1 by simp show ?thesis using v 2 proof (cases v; simp) fix w x assume wx: "v = (w \<^bold>\ x)" have "Can w \ Diag w" using v wx Diag_TensorE [of w x] by auto moreover have "Can x \ Diag x" using v wx Diag_TensorE [of w x] by auto moreover have "Dom t = Cod w" using 2 wx by simp moreover have ux: "Dom u = Cod x" using 2 wx by simp ultimately show "Can (t \<^bold>\\<^bold>\\<^bold>\ w) \ Can (u \<^bold>\\<^bold>\\<^bold>\ x)" using t u I1 I2 by simp qed qed qed thus ?thesis using assms by blast qed lemma Inv_CompDiag: assumes "Can t" and "Can u" and "Diag t" and "Diag u" and "Dom t = Cod u" shows "Inv (t \<^bold>\\<^bold>\\<^bold>\ u) = Inv u \<^bold>\\<^bold>\\<^bold>\ Inv t" proof - have "\u. \ Can t \ Diag t; Can u \ Diag u; Dom t = Cod u \ \ Inv (t \<^bold>\\<^bold>\\<^bold>\ u) = Inv u \<^bold>\\<^bold>\\<^bold>\ Inv t" proof (induct t, simp_all) show "\x u. \ C.ide x \ C.arr x; Can u \ Diag u; \<^bold>\x\<^bold>\ = Cod u \ \ Inv u = Inv u \<^bold>\\<^bold>\\<^bold>\ Inv (Cod u)" by (metis CompDiag_Diag_Dom Inv_Ide Inv_preserves_Can(2) Inv_preserves_Diag Ide.simps(1)) show "\u. Can u \ Diag u \ \<^bold>\ = Cod u \ Inv u = Inv u \<^bold>\\<^bold>\\<^bold>\ \<^bold>\" by (simp add: Inv_preserves_Can(2) Inv_preserves_Diag) fix t u v assume tu: "Can t \ Can u \ Diag (t \<^bold>\ u)" have t: "Can t \ Diag t" using tu Diag_TensorE by blast have u: "Can u \ Diag u" using tu Diag_TensorE by blast assume I1: "\v. \ Diag t; Can v \ Diag v; Dom t = Cod v \ \ Inv (t \<^bold>\\<^bold>\\<^bold>\ v) = Inv v \<^bold>\\<^bold>\\<^bold>\ Inv t" assume I2: "\v. \ Diag u; Can v \ Diag v; Dom u = Cod v \ \ Inv (u \<^bold>\\<^bold>\\<^bold>\ v) = Inv v \<^bold>\\<^bold>\\<^bold>\ Inv u" assume v: "Can v \ Diag v" assume 1: "(Dom t \<^bold>\ Dom u) = Cod v" show "Inv ((t \<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ v) = Inv v \<^bold>\\<^bold>\\<^bold>\ (Inv t \<^bold>\ Inv u)" using v 1 proof (cases v, simp_all) fix w x assume wx: "v = (w \<^bold>\ x)" have "Can w \ Diag w" using v wx Diag_TensorE [of w x] by auto moreover have "Can x \ Diag x" using v wx Diag_TensorE [of w x] by auto moreover have "Dom t = Cod w" using wx 1 by simp moreover have "Dom u = Cod x" using wx 1 by simp ultimately show "Inv (t \<^bold>\\<^bold>\\<^bold>\ w) = Inv w \<^bold>\\<^bold>\\<^bold>\ Inv t \ Inv (u \<^bold>\\<^bold>\\<^bold>\ x) = Inv x \<^bold>\\<^bold>\\<^bold>\ Inv u" using t u I1 I2 by simp qed qed thus ?thesis using assms by blast qed lemma Can_and_Diag_implies_Ide: assumes "Can t" and "Diag t" shows "Ide t" proof - have "\ Can t; Diag t \ \ Ide t" apply (induct t, simp_all) by (simp add: Diag_TensorE) thus ?thesis using assms by blast qed lemma CompDiag_Can_Inv [simp]: assumes "Can t" and "Diag t" shows "t \<^bold>\\<^bold>\\<^bold>\ Inv t = Cod t" using assms Can_and_Diag_implies_Ide Ide_in_Hom by simp lemma CompDiag_Inv_Can [simp]: assumes "Can t" and "Diag t" shows "Inv t \<^bold>\\<^bold>\\<^bold>\ t = Dom t" using assms Can_and_Diag_implies_Ide Ide_in_Hom by simp text\ The next fact is a syntactic version of the interchange law, for diagonal terms. \ lemma CompDiag_TensorDiag: assumes "Diag t" and "Diag u" and "Diag v" and "Diag w" and "Seq t v" and "Seq u w" shows "(t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = (t \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" proof - have "\u v w. \ Diag t; Diag u; Diag v; Diag w; Seq t v; Seq u w \ \ (t \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = (t \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" proof (induct t, simp_all) fix u v w assume u: "Diag u" assume v: "Diag v" assume w: "Diag w" assume uw: "Seq u w" show "Arr v \ \<^bold>\ = Cod v \ u \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = v \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" using u v w uw by (cases v) simp_all show "\f. \ C.arr f; Arr v \ \<^bold>\C.dom f\<^bold>\ = Cod v \ \ (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" proof - fix f assume f: "C.arr f" assume 1: "Arr v \ \<^bold>\C.dom f\<^bold>\ = Cod v" show "(\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" proof - have 2: "v = \<^bold>\un_Prim v\<^bold>\ \ C.arr (un_Prim v)" using v 1 by (cases v) simp_all have "u = \<^bold>\ \ ?thesis" using v w uw 1 2 Cod.simps(3) CompDiag_Cod_Diag Dom.simps(2) TensorDiag_Prim TensorDiag_term_Unity TensorDiag_preserves_Diag(3) by (cases w) simp_all moreover have "u \ \<^bold>\ \ ?thesis" proof - assume 3: "u \ \<^bold>\" hence 4: "w \ \<^bold>\" using u w uw by (cases u, simp_all; cases w, simp_all) have "(\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = (\<^bold>\f\<^bold>\ \<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\ w)" proof - have "\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u = \<^bold>\f\<^bold>\ \<^bold>\ u" using u f 3 TensorDiag_Diag by (cases u) simp_all moreover have "v \<^bold>\\<^bold>\\<^bold>\ w = v \<^bold>\ w" using w 2 4 TensorDiag_Diag by (cases v, simp_all; cases w, simp_all) ultimately show ?thesis by simp qed also have 5: "... = (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" by simp also have "... = (\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" using f u w uw 1 2 3 4 5 TensorDiag_Diag TensorDiag_Prim TensorDiag_preserves_Diag(1) CompDiag_preserves_Diag(1) by (metis Cod.simps(3) Dom.simps(1) Dom.simps(3) Diag.simps(2)) finally show ?thesis by blast qed ultimately show ?thesis by blast qed qed fix t1 t2 assume I2: "\u v w. \ Diag t2; Diag u; Diag v; Diag w; Arr v \ Dom t2 = Cod v; Seq u w \ \ (t2 \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = (t2 \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" assume t12: "Diag (t1 \<^bold>\ t2)" have t1: "t1 = \<^bold>\un_Prim t1\<^bold>\ \ C.arr (un_Prim t1) \ Diag t1" using t12 by (cases t1) simp_all have t2: "Diag t2 \ t2 \ \<^bold>\" using t12 by (cases t1) simp_all assume 1: "Arr t1 \ Arr t2 \ Arr v \ Dom t1 \<^bold>\ Dom t2 = Cod v" show "((t1 \<^bold>\ t2) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v \<^bold>\\<^bold>\\<^bold>\ w) = ((t1 \<^bold>\ t2) \<^bold>\\<^bold>\\<^bold>\ v) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" proof - have "u = \<^bold>\ \ ?thesis" using w uw TensorDiag_term_Unity CompDiag_Cod_Diag by (cases w) simp_all moreover have "u \ \<^bold>\ \ ?thesis" proof - assume u': "u \ \<^bold>\" hence w': "w \ \<^bold>\" using u w uw by (cases u; simp; cases w; simp) show ?thesis using 1 v proof (cases v, simp_all) fix v1 v2 assume v12: "v = Tensor v1 v2" have v1: "v1 = \<^bold>\un_Prim v1\<^bold>\ \ C.arr (un_Prim v1) \ Diag v1" using v v12 by (cases v1) simp_all have v2: "Diag v2 \ v2 \ \<^bold>\" using v v12 by (cases v1) simp_all have 2: "v = (\<^bold>\un_Prim v1\<^bold>\ \<^bold>\ v2)" using v1 v12 by simp show "((t1 \<^bold>\ t2) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ ((v1 \<^bold>\ v2) \<^bold>\\<^bold>\\<^bold>\ w) = ((t1 \<^bold>\\<^bold>\\<^bold>\ v1) \<^bold>\ (t2 \<^bold>\\<^bold>\\<^bold>\ v2)) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" proof - have 3: "(t1 \<^bold>\ t2) \<^bold>\\<^bold>\\<^bold>\ u = t1 \<^bold>\\<^bold>\\<^bold>\ (t2 \<^bold>\\<^bold>\\<^bold>\ u)" using u u' by (cases u) simp_all have 4: "v \<^bold>\\<^bold>\\<^bold>\ w = v1 \<^bold>\\<^bold>\\<^bold>\ (v2 \<^bold>\\<^bold>\\<^bold>\ w)" using v w v1 v2 2 TensorDiag_assoc TensorDiag_Diag by metis have "((t1 \<^bold>\ t2) \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ ((v1 \<^bold>\ v2) \<^bold>\\<^bold>\\<^bold>\ w) = (t1 \<^bold>\\<^bold>\\<^bold>\ (t2 \<^bold>\\<^bold>\\<^bold>\ u)) \<^bold>\\<^bold>\\<^bold>\ (v1 \<^bold>\\<^bold>\\<^bold>\ (v2 \<^bold>\\<^bold>\\<^bold>\ w))" using 3 4 v12 by simp also have "... = (t1 \<^bold>\\<^bold>\\<^bold>\ v1) \<^bold>\\<^bold>\\<^bold>\ ((t2 \<^bold>\\<^bold>\\<^bold>\ u) \<^bold>\\<^bold>\\<^bold>\ (v2 \<^bold>\\<^bold>\\<^bold>\ w))" proof - have "is_Tensor (t2 \<^bold>\\<^bold>\\<^bold>\ u)" using t2 u u' not_is_Tensor_TensorDiagE by auto moreover have "is_Tensor (v2 \<^bold>\\<^bold>\\<^bold>\ w)" using v2 v12 w w' not_is_Tensor_TensorDiagE by auto ultimately show ?thesis using u u' v w t1 v1 t12 v12 TensorDiag_Prim not_is_Tensor_Unity by (metis (no_types, lifting) CompDiag.simps(2) CompDiag.simps(3) is_Tensor_def) qed also have "... = (t1 \<^bold>\\<^bold>\\<^bold>\ v1) \<^bold>\\<^bold>\\<^bold>\ (t2 \<^bold>\\<^bold>\\<^bold>\ v2) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" using u w uw t2 v2 1 2 Diag_implies_Arr I2 by fastforce also have "... = ((t1 \<^bold>\\<^bold>\\<^bold>\ v1) \<^bold>\ (t2 \<^bold>\\<^bold>\\<^bold>\ v2)) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w)" proof - have "u \<^bold>\\<^bold>\\<^bold>\ w \ Unity" proof - have "Arr v1 \ \<^bold>\C.dom (un_Prim t1)\<^bold>\ = Cod v1" using t1 v1 1 2 by (cases t1, auto) thus ?thesis using t1 t2 v1 v2 u w uw u' CompDiag_preserves_Diag TensorDiag_preserves_Diag TensorDiag_Prim by (metis (mono_tags, lifting) Cod.simps(2) Cod.simps(3) TensorDiag.simps(2) term.distinct(3)) qed hence "((t1 \<^bold>\\<^bold>\\<^bold>\ v1) \<^bold>\ (t2 \<^bold>\\<^bold>\\<^bold>\ v2)) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w) = (t1 \<^bold>\\<^bold>\\<^bold>\ v1) \<^bold>\\<^bold>\\<^bold>\ ((t2 \<^bold>\\<^bold>\\<^bold>\ v2) \<^bold>\\<^bold>\\<^bold>\ (u \<^bold>\\<^bold>\\<^bold>\ w))" by (cases "u \<^bold>\\<^bold>\\<^bold>\ w") simp_all thus ?thesis by argo qed finally show ?thesis by blast qed qed qed ultimately show ?thesis by blast qed qed thus ?thesis using assms by blast qed text\ The following function reduces an arrow to diagonal form. The precise relationship between a term and its diagonalization is developed below. \ fun Diagonalize :: "'a term \ 'a term" ("\<^bold>\_\<^bold>\") where "\<^bold>\\<^bold>\f\<^bold>\\<^bold>\ = \<^bold>\f\<^bold>\" | "\<^bold>\\<^bold>\\<^bold>\ = \<^bold>\" | "\<^bold>\t \<^bold>\ u\<^bold>\ = \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\" | "\<^bold>\t \<^bold>\ u\<^bold>\ = \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\" | "\<^bold>\\<^bold>\\<^bold>[t\<^bold>]\<^bold>\ = \<^bold>\t\<^bold>\" | "\<^bold>\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<^bold>\ = \<^bold>\t\<^bold>\" | "\<^bold>\\<^bold>\\<^bold>[t\<^bold>]\<^bold>\ = \<^bold>\t\<^bold>\" | "\<^bold>\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\<^bold>\ = \<^bold>\t\<^bold>\" | "\<^bold>\\<^bold>\\<^bold>[t, u, v\<^bold>]\<^bold>\ = (\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\" | "\<^bold>\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<^bold>\ = \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\)" lemma Diag_Diagonalize: assumes "Arr t" shows "Diag \<^bold>\t\<^bold>\" and "Dom \<^bold>\t\<^bold>\ = \<^bold>\Dom t\<^bold>\" and "Cod \<^bold>\t\<^bold>\ = \<^bold>\Cod t\<^bold>\" proof - have 0: "Arr t \ Diag \<^bold>\t\<^bold>\ \ Dom \<^bold>\t\<^bold>\ = \<^bold>\Dom t\<^bold>\ \ Cod \<^bold>\t\<^bold>\ = \<^bold>\Cod t\<^bold>\" using TensorDiag_preserves_Diag CompDiag_preserves_Diag TensorDiag_assoc apply (induct t) apply auto apply (metis (full_types)) by (metis (full_types)) show "Diag \<^bold>\t\<^bold>\" using assms 0 by blast show "Dom \<^bold>\t\<^bold>\ = \<^bold>\Dom t\<^bold>\" using assms 0 by blast show "Cod \<^bold>\t\<^bold>\ = \<^bold>\Cod t\<^bold>\" using assms 0 by blast qed lemma Diagonalize_in_Hom: assumes "Arr t" shows "\<^bold>\t\<^bold>\ \ Hom \<^bold>\Dom t\<^bold>\ \<^bold>\Cod t\<^bold>\" using assms Diag_Diagonalize Diag_implies_Arr by blast lemma Diagonalize_Dom: assumes "Arr t" shows "\<^bold>\Dom t\<^bold>\ = Dom \<^bold>\t\<^bold>\" using assms Diagonalize_in_Hom by simp lemma Diagonalize_Cod: assumes "Arr t" shows "\<^bold>\Cod t\<^bold>\ = Cod \<^bold>\t\<^bold>\" using assms Diagonalize_in_Hom by simp lemma Diagonalize_preserves_Ide: assumes "Ide a" shows "Ide \<^bold>\a\<^bold>\" proof - have "Ide a \ Ide \<^bold>\a\<^bold>\" using Ide_implies_Arr TensorDiag_preserves_Ide Diag_Diagonalize by (induct a) simp_all thus ?thesis using assms by blast qed text\ The diagonalizations of canonical arrows are identities. \ lemma Ide_Diagonalize_Can: assumes "Can t" shows "Ide \<^bold>\t\<^bold>\" proof - have "Can t \ Ide \<^bold>\t\<^bold>\" using Can_implies_Arr TensorDiag_preserves_Ide Diag_Diagonalize CompDiag_preserves_Ide TensorDiag_preserves_Diag by (induct t) auto thus ?thesis using assms by blast qed lemma Diagonalize_preserves_Can: assumes "Can t" shows "Can \<^bold>\t\<^bold>\" using assms Ide_Diagonalize_Can Ide_implies_Can by auto lemma Diagonalize_Diag [simp]: assumes "Diag t" shows "\<^bold>\t\<^bold>\ = t" proof - have "Diag t \ \<^bold>\t\<^bold>\ = t" apply (induct t, simp_all) using TensorDiag_Prim Diag_TensorE by metis thus ?thesis using assms by blast qed lemma Diagonalize_Diagonalize [simp]: assumes "Arr t" shows "\<^bold>\\<^bold>\t\<^bold>\\<^bold>\ = \<^bold>\t\<^bold>\" using assms Diag_Diagonalize Diagonalize_Diag by blast lemma Diagonalize_Tensor: assumes "Arr t" and "Arr u" shows "\<^bold>\t \<^bold>\ u\<^bold>\ = \<^bold>\\<^bold>\t\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\\<^bold>\" using assms Diagonalize_Diagonalize by simp lemma Diagonalize_Tensor_Unity_Arr [simp]: assumes "Arr u" shows "\<^bold>\\<^bold>\ \<^bold>\ u\<^bold>\ = \<^bold>\u\<^bold>\" using assms by simp lemma Diagonalize_Tensor_Arr_Unity [simp]: assumes "Arr t" shows "\<^bold>\t \<^bold>\ \<^bold>\\<^bold>\ = \<^bold>\t\<^bold>\" using assms by simp lemma Diagonalize_Tensor_Prim_Arr [simp]: assumes "arr f" and "Arr u" and "\<^bold>\u\<^bold>\ \ Unity" shows "\<^bold>\\<^bold>\f\<^bold>\ \<^bold>\ u\<^bold>\ = \<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\u\<^bold>\" using assms by simp lemma Diagonalize_Tensor_Tensor: assumes "Arr t" and "Arr u" and "Arr v" shows "\<^bold>\(t \<^bold>\ u) \<^bold>\ v\<^bold>\ = \<^bold>\\<^bold>\t\<^bold>\ \<^bold>\ (\<^bold>\u\<^bold>\ \<^bold>\ \<^bold>\v\<^bold>\)\<^bold>\" using assms Diag_Diagonalize Diagonalize_Diagonalize by (simp add: TensorDiag_assoc) lemma Diagonalize_Comp_Cod_Arr: assumes "Arr t" shows "\<^bold>\Cod t \<^bold>\ t\<^bold>\ = \<^bold>\t\<^bold>\" proof - have "Arr t \ \<^bold>\Cod t \<^bold>\ t\<^bold>\ = \<^bold>\t\<^bold>\" using C.comp_cod_arr apply (induct t, simp_all) using CompDiag_TensorDiag Arr_implies_Ide_Cod Ide_in_Hom Diag_Diagonalize Diagonalize_in_Hom apply simp using CompDiag_preserves_Diag CompDiag_Cod_Diag Diag_Diagonalize apply metis using CompDiag_TensorDiag Arr_implies_Ide_Cod Ide_in_Hom TensorDiag_in_Hom TensorDiag_preserves_Diag Diag_Diagonalize Diagonalize_in_Hom TensorDiag_assoc by simp_all thus ?thesis using assms by blast qed lemma Diagonalize_Comp_Arr_Dom: assumes "Arr t" shows "\<^bold>\t \<^bold>\ Dom t\<^bold>\ = \<^bold>\t\<^bold>\" proof - have "Arr t \ \<^bold>\t \<^bold>\ Dom t\<^bold>\ = \<^bold>\t\<^bold>\" by (metis CompDiag_Diag_Dom Diag_Diagonalize(1-2) Diagonalize.simps(4)) thus ?thesis using assms by blast qed lemma Diagonalize_Inv: assumes "Can t" shows "\<^bold>\Inv t\<^bold>\ = Inv \<^bold>\t\<^bold>\" proof - have "Can t \ \<^bold>\Inv t\<^bold>\ = Inv \<^bold>\t\<^bold>\" proof (induct t, simp_all) fix u v assume I1: "\<^bold>\Inv u\<^bold>\ = Inv \<^bold>\u\<^bold>\" assume I2: "\<^bold>\Inv v\<^bold>\ = Inv \<^bold>\v\<^bold>\" show "Can u \ Can v \ Inv \<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Inv \<^bold>\v\<^bold>\ = Inv (\<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\)" using Inv_TensorDiag Diag_Diagonalize Can_implies_Arr Diagonalize_preserves_Can I1 I2 by simp show "Can u \ Can v \ Dom u = Cod v \ Inv \<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Inv \<^bold>\u\<^bold>\ = Inv (\<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\)" using Inv_CompDiag Diag_Diagonalize Can_implies_Arr Diagonalize_in_Hom Diagonalize_preserves_Can I1 I2 by simp fix w assume I3: "\<^bold>\Inv w\<^bold>\ = Inv \<^bold>\w\<^bold>\" assume uvw: "Can u \ Can v \ Can w" show "Inv \<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (Inv \<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Inv \<^bold>\w\<^bold>\) = Inv ((\<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\)" using uvw I1 I2 I3 Inv_TensorDiag Diag_Diagonalize Can_implies_Arr Diagonalize_preserves_Can TensorDiag_preserves_Diag TensorDiag_preserves_Can TensorDiag_assoc by simp show "(Inv \<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ Inv \<^bold>\v\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ Inv \<^bold>\w\<^bold>\ = Inv (\<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ (\<^bold>\v\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\w\<^bold>\))" - using uvw I1 I2 I3 - Inv_TensorDiag Diag_Diagonalize Can_implies_Arr Diagonalize_preserves_Can - TensorDiag_preserves_Diag TensorDiag_preserves_Can - apply simp - using TensorDiag_assoc [of "\<^bold>\u\<^bold>\" "\<^bold>\v\<^bold>\" "\<^bold>\w\<^bold>\"] by metis + by (simp add: Can_implies_Arr Ide_Diagonalize_Can TensorDiag_assoc + TensorDiag_preserves_Diag(1) TensorDiag_preserves_Ide Diag_Diagonalize(1) uvw) qed thus ?thesis using assms by blast qed text\ Our next objective is to begin making the connection, to be completed in a subsequent section, between arrows and their diagonalizations. To summarize, an arrow @{term t} and its diagonalization @{term "\<^bold>\t\<^bold>\"} are opposite sides of a square whose other sides are certain canonical terms \Dom t\<^bold>\ \ Hom (Dom t) \<^bold>\Dom t\<^bold>\\ and \Cod t\<^bold>\ \ Hom (Cod t) \<^bold>\Cod t\<^bold>\\, where \Dom t\<^bold>\\ and \Cod t\<^bold>\\ are defined by the function \red\ below. The coherence theorem amounts to the statement that every such square commutes when the formal terms involved are evaluated in the evident way in any monoidal category. Function \red\ defined below takes an identity term @{term a} to a canonical arrow \a\<^bold>\ \ Hom a \<^bold>\a\<^bold>\\. The auxiliary function \red2\ takes a pair @{term "(a, b)"} of diagonal identity terms and produces a canonical arrow \a \<^bold>\ b \ Hom (a \<^bold>\ b) \<^bold>\a \<^bold>\ b\<^bold>\\. The canonical arrow \a\<^bold>\\ amounts to a ``parallel innermost reduction'' from @{term a} to @{term "\<^bold>\a\<^bold>\"}, where the reduction steps are canonical arrows that involve the unitors and associator only in their uninverted forms. In general, a parallel innermost reduction from @{term a} will not be unique: at some points there is a choice available between left and right unitors and at other points there are choices between unitors and associators. These choices are inessential, and the ordering of the clauses in the function definitions below resolves them in an arbitrary way. What is more important is having chosen an innermost reduction, which is what allows us to write these definitions in structurally recursive form. The essence of coherence is that the axioms for a monoidal category allow us to prove that any reduction from @{term a} to @{term "\<^bold>\a\<^bold>\"} is equivalent (under evaluation of terms) to a parallel innermost reduction. The problematic cases are terms of the form @{term "((a \<^bold>\ b) \<^bold>\ c) \<^bold>\ d"}, which present a choice between an inner and outer reduction that lead to terms with different structures. It is of course the pentagon axiom that ensures the confluence (under evaluation) of the two resulting paths. Although simple in appearance, the structurally recursive definitions below were difficult to get right even after I started to understand what I was doing. I wish I could have just written them down straightaway. If so, then I could have avoided laboriously constructing and then throwing away thousands of lines of proof text that used a non-structural, ``operational'' approach to defining a reduction from @{term a} to @{term "\<^bold>\a\<^bold>\"}. \ fun red2 (infixr "\<^bold>\" 53) where "\<^bold>\ \<^bold>\ a = \<^bold>\\<^bold>[a\<^bold>]" | "\<^bold>\f\<^bold>\ \<^bold>\ \<^bold>\ = \<^bold>\\<^bold>[\<^bold>\f\<^bold>\\<^bold>]" | "\<^bold>\f\<^bold>\ \<^bold>\ a = \<^bold>\f\<^bold>\ \<^bold>\ a" | "(a \<^bold>\ b) \<^bold>\ \<^bold>\ = \<^bold>\\<^bold>[a \<^bold>\ b\<^bold>]" | "(a \<^bold>\ b) \<^bold>\ c = (a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\) \<^bold>\ (a \<^bold>\ (b \<^bold>\ c)) \<^bold>\ \<^bold>\\<^bold>[a, b, c\<^bold>]" | "a \<^bold>\ b = undefined" fun red ("_\<^bold>\" [56] 56) where "\<^bold>\\<^bold>\ = \<^bold>\" | "\<^bold>\f\<^bold>\\<^bold>\ = \<^bold>\f\<^bold>\" | "(a \<^bold>\ b)\<^bold>\ = (if Diag (a \<^bold>\ b) then a \<^bold>\ b else (\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\) \<^bold>\ (a\<^bold>\ \<^bold>\ b\<^bold>\))" | "a\<^bold>\ = undefined" lemma red_Diag [simp]: assumes "Diag a" shows "a\<^bold>\ = a" using assms by (cases a) simp_all lemma red2_Diag: assumes "Diag (a \<^bold>\ b)" shows "a \<^bold>\ b = a \<^bold>\ b" proof - have a: "a = \<^bold>\un_Prim a\<^bold>\" using assms Diag_TensorE by metis have b: "Diag b \ b \ \<^bold>\" using assms Diag_TensorE by metis show ?thesis using a b apply (cases b) apply simp_all apply (metis red2.simps(3)) by (metis red2.simps(4)) qed lemma Can_red2: assumes "Ide a" and "Diag a" and "Ide b" and "Diag b" shows "Can (a \<^bold>\ b)" and "a \<^bold>\ b \ Hom (a \<^bold>\ b) \<^bold>\a \<^bold>\ b\<^bold>\" proof - have 0: "\b. \ Ide a \ Diag a; Ide b \ Diag b \ \ Can (a \<^bold>\ b) \ a \<^bold>\ b \ Hom (a \<^bold>\ b) \<^bold>\a \<^bold>\ b\<^bold>\" proof (induct a, simp_all) fix b show "Ide b \ Diag b \ Can b \ Dom b = b \ Cod b = b" using Ide_implies_Can Ide_in_Hom Diagonalize_Diag by auto fix f show "\ C.ide f \ C.arr f; Ide b \ Diag b \ \ Can (\<^bold>\f\<^bold>\ \<^bold>\ b) \ Arr (\<^bold>\f\<^bold>\ \<^bold>\ b) \ Dom (\<^bold>\f\<^bold>\ \<^bold>\ b) = \<^bold>\f\<^bold>\ \<^bold>\ b \ Cod (\<^bold>\f\<^bold>\ \<^bold>\ b) = \<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ b" using Ide_implies_Can Ide_in_Hom by (cases b; auto) next fix a b c assume ab: "Ide a \ Ide b \ Diag (Tensor a b)" assume c: "Ide c \ Diag c" assume I1: "\c. \ Diag a; Ide c \ Diag c \ \ Can (a \<^bold>\ c) \ Arr (a \<^bold>\ c) \ Dom (a \<^bold>\ c) = a \<^bold>\ c \ Cod (a \<^bold>\ c) = a \<^bold>\\<^bold>\\<^bold>\ c" assume I2: "\c. \ Diag b; Ide c \ Diag c \ \ Can (b \<^bold>\ c) \ Arr (b \<^bold>\ c) \ Dom (b \<^bold>\ c) = b \<^bold>\ c \ Cod (b \<^bold>\ c) = b \<^bold>\\<^bold>\\<^bold>\ c" show "Can ((a \<^bold>\ b) \<^bold>\ c) \ Arr ((a \<^bold>\ b) \<^bold>\ c) \ Dom ((a \<^bold>\ b) \<^bold>\ c) = (a \<^bold>\ b) \<^bold>\ c \ Cod ((a \<^bold>\ b) \<^bold>\ c) = (\<^bold>\a\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ c" proof - have a: "Diag a \ Ide a" using ab Diag_TensorE by blast have b: "Diag b \ Ide b" using ab Diag_TensorE by blast have "c = \<^bold>\ \ ?thesis" proof - assume 1: "c = \<^bold>\" have 2: "(a \<^bold>\ b) \<^bold>\ c = \<^bold>\\<^bold>[a \<^bold>\ b\<^bold>]" using 1 by simp have 3: "Can (a \<^bold>\ b) \ Arr (a \<^bold>\ b) \ Dom (a \<^bold>\ b) = a \<^bold>\ b \ Cod (a \<^bold>\ b) = a \<^bold>\ b" using a b ab 1 2 I1 Diagonalize_Diag Diagonalize.simps(3) by metis hence 4: "Seq (a \<^bold>\ b) \<^bold>\\<^bold>[a \<^bold>\ b\<^bold>]" using ab by (metis (mono_tags, lifting) Arr.simps(7) Cod.simps(3) Cod.simps(7) Diag_implies_Arr Ide_in_Hom mem_Collect_eq) have "Can ((a \<^bold>\ b) \<^bold>\ c)" using 1 2 3 4 ab by (simp add: Ide_implies_Can) moreover have "Dom ((a \<^bold>\ b) \<^bold>\ c) = (a \<^bold>\ b) \<^bold>\ c" using 1 2 3 4 a b ab I1 Ide_in_Hom TensorDiag_preserves_Diag by simp moreover have "Cod ((a \<^bold>\ b) \<^bold>\ c) = \<^bold>\(a \<^bold>\ b) \<^bold>\ c\<^bold>\" using 1 2 3 4 ab using Diagonalize_Diag by fastforce ultimately show ?thesis using Can_implies_Arr by (simp add: 1 ab) qed moreover have "c \ \<^bold>\ \ ?thesis" proof - assume 1: "c \ \<^bold>\" have 2: "(a \<^bold>\ b) \<^bold>\ c = (a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\) \<^bold>\ (a \<^bold>\ b \<^bold>\ c) \<^bold>\ \<^bold>\\<^bold>[a, b, c\<^bold>]" using 1 a b ab c by (cases c; simp) have 3: "Can (a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\) \ Dom (a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\) = a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\ \ Cod (a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\) = \<^bold>\a \<^bold>\ (b \<^bold>\ c)\<^bold>\" proof - have "Can (a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\) \ Dom (a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\) = a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\ \ Cod (a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\) = \<^bold>\a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\\<^bold>\" using a c ab 1 2 I1 Diag_implies_Arr Diag_Diagonalize(1) Diagonalize_preserves_Ide TensorDiag_preserves_Ide TensorDiag_preserves_Diag(1) by auto moreover have "\<^bold>\a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\\<^bold>\ = \<^bold>\a \<^bold>\ (b \<^bold>\ c)\<^bold>\" using ab c Diagonalize_Tensor Diagonalize_Diagonalize Diag_implies_Arr by (metis Arr.simps(3) Diagonalize.simps(3)) ultimately show ?thesis by metis qed have 4: "Can (b \<^bold>\ c) \ Dom (b \<^bold>\ c) = b \<^bold>\ c \ Cod (b \<^bold>\ c) = \<^bold>\b \<^bold>\ c\<^bold>\" using b c ab 1 2 I2 by simp hence "Can (a \<^bold>\ (b \<^bold>\ c)) \ Dom (a \<^bold>\ (b \<^bold>\ c)) = a \<^bold>\ (b \<^bold>\ c) \ Cod (a \<^bold>\ (b \<^bold>\ c)) = a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\" using ab Ide_implies_Can Ide_in_Hom by force moreover have "\<^bold>\a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\\<^bold>\ = \<^bold>\a \<^bold>\ b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\" proof - have "\<^bold>\a \<^bold>\ \<^bold>\b \<^bold>\ c\<^bold>\\<^bold>\ = a \<^bold>\\<^bold>\\<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)" using a b c 4 by (metis Arr_implies_Ide_Dom Can_implies_Arr Ide_implies_Arr Diag_Diagonalize(1) Diagonalize.simps(3) Diagonalize_Diag) also have "... = (a \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c" using a b ab c TensorDiag_assoc by metis also have "... = \<^bold>\a \<^bold>\ b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\" using a b c by (metis Diagonalize.simps(3) Diagonalize_Diag) finally show ?thesis by blast qed moreover have "Can \<^bold>\\<^bold>[a, b, c\<^bold>] \ Dom \<^bold>\\<^bold>[a, b, c\<^bold>] = (a \<^bold>\ b) \<^bold>\ c \ Cod \<^bold>\\<^bold>[a, b, c\<^bold>] = a \<^bold>\ (b \<^bold>\ c)" using ab c Ide_implies_Can Ide_in_Hom by auto ultimately show ?thesis - using ab c 2 3 4 Diag_implies_Arr Diagonalize_Diagonalize Ide_implies_Can + using c 2 3 4 Diagonalize_Diagonalize Ide_implies_Can Diagonalize_Diag Arr_implies_Ide_Dom Can_implies_Arr by (metis Can.simps(4) Cod.simps(4) Dom.simps(4) Diagonalize.simps(3)) qed ultimately show ?thesis by blast qed qed show "Can (a \<^bold>\ b)" using assms 0 by blast show "a \<^bold>\ b \ Hom (a \<^bold>\ b) \<^bold>\a \<^bold>\ b\<^bold>\" using 0 assms by blast qed lemma red2_in_Hom: assumes "Ide a" and "Diag a" and "Ide b" and "Diag b" shows "a \<^bold>\ b \ Hom (a \<^bold>\ b) \<^bold>\a \<^bold>\ b\<^bold>\" using assms Can_red2 Can_implies_Arr by simp lemma Can_red: assumes "Ide a" shows "Can (a\<^bold>\)" and "a\<^bold>\ \ Hom a \<^bold>\a\<^bold>\" proof - have 0: "Ide a \ Can (a\<^bold>\) \ a\<^bold>\ \ Hom a \<^bold>\a\<^bold>\" proof (induct a, simp_all) fix b c assume b: "Can (b\<^bold>\) \ Arr (b\<^bold>\) \ Dom (b\<^bold>\) = b \ Cod (b\<^bold>\) = \<^bold>\b\<^bold>\" assume c: "Can (c\<^bold>\) \ Arr (c\<^bold>\) \ Dom (c\<^bold>\) = c \ Cod (c\<^bold>\) = \<^bold>\c\<^bold>\" assume bc: "Ide b \ Ide c" show "(Diag (b \<^bold>\ c) \ Can b \ Can c \ Dom b = b \ Dom c = c \ Cod b \<^bold>\ Cod c = \<^bold>\b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\) \ (\ Diag (b \<^bold>\ c) \ Can (\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\) \ Dom (\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\) = \<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\ \ Arr (\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\) \ Dom (\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\) = \<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\ \ Cod (\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\) = \<^bold>\b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\)" proof show "Diag (b \<^bold>\ c) \ Can b \ Can c \ Dom b = b \ Dom c = c \ Cod b \<^bold>\ Cod c = \<^bold>\b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\" using bc Diag_TensorE Ide_implies_Can Inv_preserves_Can(2) CompDiag_Ide_Diag Inv_Ide Diagonalize.simps(3) Diagonalize_Diag by (metis CompDiag_Inv_Can) show "\ Diag (b \<^bold>\ c) \ Can (\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\) \ Dom (\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\) = \<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\ \ Arr (\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\) \ Dom (\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\) = \<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\ \ Cod (\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\) = \<^bold>\b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\" - using b c bc Ide_in_Hom Ide_implies_Can Diagonalize_Diag Can_red2 Diag_Diagonalize - Ide_implies_Arr Diagonalize_Tensor Diagonalize_preserves_Ide - TensorDiag_preserves_Diag TensorDiag_preserves_Ide - TensorDiag_preserves_Can - by (cases b) simp_all + using b c bc Ide_in_Hom Ide_implies_Can Can_red2 Diag_Diagonalize + Diagonalize_preserves_Ide TensorDiag_preserves_Diag TensorDiag_preserves_Ide + by force qed qed show "Can (a\<^bold>\)" using assms 0 by blast show "a\<^bold>\ \ Hom a \<^bold>\a\<^bold>\" using assms 0 by blast qed lemma red_in_Hom: assumes "Ide a" shows "a\<^bold>\ \ Hom a \<^bold>\a\<^bold>\" using assms Can_red Can_implies_Arr by simp lemma Diagonalize_red [simp]: assumes "Ide a" shows "\<^bold>\a\<^bold>\\<^bold>\ = \<^bold>\a\<^bold>\" using assms Can_red Ide_Diagonalize_Can Diagonalize_in_Hom Ide_in_Hom by fastforce lemma Diagonalize_red2 [simp]: assumes "Ide a" and "Ide b" and "Diag a" and "Diag b" shows "\<^bold>\a \<^bold>\ b\<^bold>\ = \<^bold>\a \<^bold>\ b\<^bold>\" using assms Can_red2 Ide_Diagonalize_Can Diagonalize_in_Hom [of "a \<^bold>\ b"] red2_in_Hom Ide_in_Hom by simp end section "Coherence" text\ If @{term D} is a monoidal category, then a functor \V: C \ D\ extends in an evident way to an evaluation map that interprets each formal arrow of the monoidal language of @{term C} as an arrow of @{term D}. \ locale evaluation_map = monoidal_language C + monoidal_category D T \ \ + V: "functor" C D V for C :: "'c comp" (infixr "\\<^sub>C" 55) and D :: "'d comp" (infixr "\" 55) and T :: "'d * 'd \ 'd" and \ :: "'d * 'd * 'd \ 'd" and \ :: 'd and V :: "'c \ 'd" begin no_notation C.in_hom ("\_ : _ \ _\") notation unity ("\") notation runit ("\[_]") notation lunit ("\[_]") notation assoc' ("\\<^sup>-\<^sup>1[_, _, _]") notation runit' ("\\<^sup>-\<^sup>1[_]") notation lunit' ("\\<^sup>-\<^sup>1[_]") primrec eval :: "'c term \ 'd" ("\_\") where "\\<^bold>\f\<^bold>\\ = V f" | "\\<^bold>\\ = \" | "\t \<^bold>\ u\ = \t\ \ \u\" | "\t \<^bold>\ u\ = \t\ \ \u\" | "\\<^bold>\\<^bold>[t\<^bold>]\ = \ \t\" | "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \' \t\" | "\\<^bold>\\<^bold>[t\<^bold>]\ = \ \t\" | "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \' \t\" | "\\<^bold>\\<^bold>[t, u, v\<^bold>]\ = \ (\t\, \u\, \v\)" | "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = \' (\t\, \u\, \v\)" text\ Identity terms evaluate to identities of \D\ and evaluation preserves domain and codomain. \ lemma ide_eval_Ide [simp]: shows "Ide t \ ide \t\" by (induct t, auto) lemma eval_in_hom: shows "Arr t \ \\t\ : \Dom t\ \ \Cod t\\" apply (induct t) apply auto[4] apply fastforce proof - fix t u v assume I: "Arr t \ \\t\ : \Dom t\ \ \Cod t\\" show "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ \\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ : \Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ \ \Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\\" using I arr_dom_iff_arr [of "\t\"] by force show "Arr \<^bold>\\<^bold>[t\<^bold>] \ \\\<^bold>\\<^bold>[t\<^bold>]\ : \Dom \<^bold>\\<^bold>[t\<^bold>]\ \ \Cod \<^bold>\\<^bold>[t\<^bold>]\\" using I arr_cod_iff_arr [of "\t\"] by force show "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ \\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ : \Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ \ \Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\\" using I arr_dom_iff_arr [of "\t\"] by force assume I1: "Arr t \ \\t\ : \Dom t\ \ \Cod t\\" assume I2: "Arr u \ \\u\ : \Dom u\ \ \Cod u\\" assume I3: "Arr v \ \\v\ : \Dom v\ \ \Cod v\\" show "Arr \<^bold>\\<^bold>[t, u, v\<^bold>] \ \\\<^bold>\\<^bold>[t, u, v\<^bold>]\ : \Dom \<^bold>\\<^bold>[t, u, v\<^bold>]\ \ \Cod \<^bold>\\<^bold>[t, u, v\<^bold>]\\" proof - assume 1: "Arr \<^bold>\\<^bold>[t, u, v\<^bold>]" have t: "\\t\ : dom \t\ \ cod \t\\" using 1 I1 by auto have u: "\\u\ : dom \u\ \ cod \u\\" using 1 I2 by auto have v: "\\v\ : dom \v\ \ cod \v\\" using 1 I3 by auto have "\\<^bold>\\<^bold>[t, u, v\<^bold>]\ = (\t\ \ \u\ \ \v\) \ \[dom \t\, dom \u\, dom \v\]" using t u v \_simp [of "\t\" "\u\" "\v\"] by auto moreover have "\(\t\ \ \u\ \ \v\) \ \[dom \t\, dom \u\, dom \v\] : (dom \t\ \ dom \u\) \ dom \v\ \ cod \t\ \ cod \u\ \ cod \v\\" using t u v by (elim in_homE, auto) moreover have "\Dom t\ = dom \t\ \ \Dom u\ = dom \u\ \ \Dom v\ = dom \v\ \ \Cod t\ = cod \t\ \ \Cod u\ = cod \u\ \ \Cod v\ = cod \v\" using 1 I1 I2 I3 by auto ultimately show "\\\<^bold>\\<^bold>[t, u, v\<^bold>]\ : \Dom \<^bold>\\<^bold>[t, u, v\<^bold>]\ \ \Cod \<^bold>\\<^bold>[t, u, v\<^bold>]\\" by simp qed show "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] \ \\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ : \Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ \ \Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\\" proof - assume 1: "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]" have t: "\\t\ : dom \t\ \ cod \t\\" using 1 I1 by auto have u: "\\u\ : dom \u\ \ cod \u\\" using 1 I2 by auto have v: "\\v\ : dom \v\ \ cod \v\\" using 1 I3 by auto have "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = ((\t\ \ \u\) \ \v\) \ \\<^sup>-\<^sup>1[dom \t\, dom \u\, dom \v\]" using 1 I1 I2 I3 \'_simp [of "\t\" "\u\" "\v\"] by auto moreover have "\((\t\ \ \u\) \ \v\) \ \\<^sup>-\<^sup>1[dom \t\, dom \u\, dom \v\] : dom \t\ \ dom \u\ \ dom \v\ \ (cod \t\ \ cod \u\) \ cod \v\\" using t u v assoc'_in_hom [of "dom \t\" "dom \u\" "dom \v\"] by (elim in_homE, auto) moreover have "\Dom t\ = dom \t\ \ \Dom u\ = dom \u\ \ \Dom v\ = dom \v\ \ \Cod t\ = cod \t\ \ \Cod u\ = cod \u\ \ \Cod v\ = cod \v\" using 1 I1 I2 I3 by auto ultimately show " \\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ : \Dom \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ \ \Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\\" by simp qed qed lemma arr_eval [simp]: assumes "Arr f" shows "arr \f\" using assms eval_in_hom by auto lemma dom_eval [simp]: assumes "Arr f" shows "dom \f\ = \Dom f\" using assms eval_in_hom by auto lemma cod_eval [simp]: assumes "Arr f" shows "cod \f\ = \Cod f\" using assms eval_in_hom by auto lemma eval_Prim [simp]: assumes "C.arr f" shows "\\<^bold>\f\<^bold>\\ = V f" by simp lemma eval_Tensor [simp]: assumes "Arr t" and "Arr u" shows "\t \<^bold>\ u\ = \t\ \ \u\" using assms eval_in_hom by auto lemma eval_Comp [simp]: assumes "Arr t" and "Arr u" and "Dom t = Cod u" shows " \t \<^bold>\ u\ = \t\ \ \u\" using assms by simp lemma eval_Lunit [simp]: assumes "Arr t" shows "\\<^bold>\\<^bold>[t\<^bold>]\ = \[\Cod t\] \ (\ \ \t\)" using assms lunit_naturality [of "\t\"] by simp lemma eval_Lunit' [simp]: assumes "Arr t" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \\<^sup>-\<^sup>1[\Cod t\] \ \t\" using assms lunit'_naturality [of "\t\"] \'.map_simp [of "\t\"] \_ide_simp Arr_implies_Ide_Cod by simp lemma eval_Runit [simp]: assumes "Arr t" shows "\\<^bold>\\<^bold>[t\<^bold>]\ = \[\Cod t\] \ (\t\ \ \)" using assms runit_naturality [of "\t\"] by simp lemma eval_Runit' [simp]: assumes "Arr t" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = \\<^sup>-\<^sup>1[\Cod t\] \ \t\" using assms runit'_naturality [of "\t\"] \'.map_simp [of "\t\"] \_ide_simp Arr_implies_Ide_Cod by simp lemma eval_Assoc [simp]: assumes "Arr t" and "Arr u" and "Arr v" shows "\\<^bold>\\<^bold>[t, u, v\<^bold>]\ = \[cod \t\, cod \u\, cod \v\] \ ((\t\ \ \u\) \ \v\)" using assms \.is_natural_2 [of "(\t\, \u\, \v\)"] by auto lemma eval_Assoc' [simp]: assumes "Arr t" and "Arr u" and "Arr v" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = \\<^sup>-\<^sup>1[cod \t\, cod \u\, cod \v\] \ (\t\ \ \u\ \ \v\)" using assms \'_simp [of "\t\" "\u\" "\v\"] assoc'_naturality [of "\t\" "\u\" "\v\"] by simp text\ The following are conveniences for the case of identity arguments to avoid having to get rid of the extra identities that are introduced by the general formulas above. \ lemma eval_Lunit_Ide [simp]: assumes "Ide a" shows "\\<^bold>\\<^bold>[a\<^bold>]\ = \[\a\]" using assms comp_cod_arr by simp lemma eval_Lunit'_Ide [simp]: assumes "Ide a" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[a\<^bold>]\ = \\<^sup>-\<^sup>1[\a\]" using assms comp_cod_arr by simp lemma eval_Runit_Ide [simp]: assumes "Ide a" shows "\\<^bold>\\<^bold>[a\<^bold>]\ = \[\a\]" using assms comp_cod_arr by simp lemma eval_Runit'_Ide [simp]: assumes "Ide a" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[a\<^bold>]\ = \\<^sup>-\<^sup>1[\a\]" using assms comp_cod_arr by simp lemma eval_Assoc_Ide [simp]: assumes "Ide a" and "Ide b" and "Ide c" shows "\\<^bold>\\<^bold>[a, b, c\<^bold>]\ = \[\a\, \b\, \c\]" using assms by simp lemma eval_Assoc'_Ide [simp]: assumes "Ide a" and "Ide b" and "Ide c" shows "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[a, b, c\<^bold>]\ = \\<^sup>-\<^sup>1[\a\, \b\, \c\]" using assms \'_ide_simp by simp text\ Canonical arrows evaluate to isomorphisms in \D\, and formal inverses evaluate to inverses in \D\. \ lemma iso_eval_Can: shows "Can t \ iso \t\" - using Can_implies_Arr iso_is_arr \'.preserves_iso \'.preserves_iso - \.preserves_iso \'.preserves_iso Arr_implies_Ide_Dom Arr_implies_Ide_Cod + using Can_implies_Arr \'.preserves_iso \'.preserves_iso \.preserves_iso \'.preserves_iso + Arr_implies_Ide_Dom by (induct t) auto lemma eval_Inv_Can: shows "Can t \ \Inv t\ = inv \t\" apply (induct t) using iso_eval_Can inv_comp Can_implies_Arr apply auto[4] proof - fix t assume I: "Can t \ \Inv t\ = inv \t\" show "Can \<^bold>\\<^bold>[t\<^bold>] \ \Inv \<^bold>\\<^bold>[t\<^bold>]\ = inv \\<^bold>\\<^bold>[t\<^bold>]\" using I \'.is_natural_2 [of "inv \t\"] iso_eval_Can \_ide_simp iso_is_arr comp_cod_arr inv_comp by simp show "Can \<^bold>\\<^bold>[t\<^bold>] \ \Inv \<^bold>\\<^bold>[t\<^bold>]\ = inv \\<^bold>\\<^bold>[t\<^bold>]\" using I \'.is_natural_2 [of "inv \t\"] iso_eval_Can \_ide_simp iso_is_arr comp_cod_arr inv_comp by simp show "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ \Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\" proof - assume t: "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]" hence 1: "iso \t\" using iso_eval_Can by simp have "inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = inv (\' \t\)" using t by simp also have "... = inv (\\<^sup>-\<^sup>1[cod \t\] \ \t\)" using 1 \'.is_natural_2 [of "\t\"] \'_ide_simp iso_is_arr by auto also have "... = \Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\" using t I 1 iso_is_arr inv_comp by auto finally show ?thesis by simp qed show "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ \Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\" proof - assume t: "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]" hence 1: "iso \t\" using iso_eval_Can by simp have "inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\ = inv (\' \t\)" using t by simp also have "... = inv (\\<^sup>-\<^sup>1[cod \t\] \ \t\)" using 1 \'.is_natural_2 [of "\t\"] \'_ide_simp iso_is_arr by auto also have "... = \Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]\" using t I 1 iso_is_arr inv_comp by auto finally show ?thesis by simp qed next fix t u v assume I1: "Can t \ \Inv t\ = inv \t\" assume I2: "Can u \ \Inv u\ = inv \u\" assume I3: "Can v \ \Inv v\ = inv \v\" show "Can \<^bold>\\<^bold>[t, u, v\<^bold>] \ \Inv \<^bold>\\<^bold>[t, u, v\<^bold>]\ = inv \\<^bold>\\<^bold>[t, u, v\<^bold>]\" proof - assume tuv: "Can \<^bold>\\<^bold>[t, u, v\<^bold>]" have t: "iso \t\" using tuv iso_eval_Can by auto have u: "iso \u\" using tuv iso_eval_Can by auto have v: "iso \v\" using tuv iso_eval_Can by auto have "\Inv \<^bold>\\<^bold>[t, u, v\<^bold>]\ = \' (inv \t\, inv \u\, inv \v\)" using tuv I1 I2 I3 by simp also have "... = inv (\[cod \t\, cod \u\, cod \v\] \ ((\t\ \ \u\) \ \v\))" using t u v \'_simp iso_is_arr inv_comp by auto also have "... = inv ((\t\ \ \u\ \ \v\) \ \[dom \t\, dom \u\, dom \v\])" using t u v iso_is_arr assoc_naturality by simp also have "... = inv \\<^bold>\\<^bold>[t, u, v\<^bold>]\" using t u v iso_is_arr \_simp [of "\t\" "\u\" "\v\"] by simp finally show ?thesis by simp qed show "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] \ \Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\" proof - assume tuv: "Can \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]" have t: "iso \t\" using tuv iso_eval_Can by auto have u: "iso \u\" using tuv iso_eval_Can by auto have v: "iso \v\" using tuv iso_eval_Can by auto have "\Inv \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = \ (inv \t\, inv \u\, inv \v\)" using tuv I1 I2 I3 by simp also have "... = (inv \t\ \ inv \u\ \ inv \v\) \ \[cod \t\, cod \u\, cod \v\]" using t u v iso_is_arr \_simp [of "inv \t\" "inv \u\" "inv \v\"] by simp also have "... = inv (\\<^sup>-\<^sup>1[cod \t\, cod \u\, cod \v\] \ (\t\ \ \u\ \ \v\))" using t u v iso_is_arr inv_comp by auto also have "... = inv (((\t\ \ \u\) \ \v\) \ \\<^sup>-\<^sup>1[dom \t\, dom \u\, dom \v\])" using t u v iso_is_arr assoc'_naturality by simp also have "... = inv \\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\" using t u v iso_is_arr \'_simp by auto finally show ?thesis by blast qed qed text\ The operation \\<^bold>\\<^bold>\\<^bold>\\ evaluates to composition in \D\. \ lemma eval_CompDiag: assumes "Diag t" and "Diag u" and "Seq t u" shows "\t \<^bold>\\<^bold>\\<^bold>\ u\ = \t\ \ \u\" proof - have "\u. \ Diag t; Diag u; Seq t u \ \ \t \<^bold>\\<^bold>\\<^bold>\ u\ = \t\ \ \u\" using eval_in_hom comp_cod_arr proof (induct t, simp_all) fix u f assume u: "Diag u" assume f: "C.arr f" assume 1: "Arr u \ \<^bold>\C.dom f\<^bold>\ = Cod u" show "\\<^bold>\f\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ u\ = V f \ \u\" using f u 1 as_nat_trans.preserves_comp_2 by (cases u; simp) next fix u v w assume I1: "\u. \ Diag v; Diag u; Arr u \ Dom v = Cod u \ \ \v \<^bold>\\<^bold>\\<^bold>\ u\ = \v\ \ \u\" assume I2: "\u. \ Diag w; Diag u; Arr u \ Dom w = Cod u \ \ \w \<^bold>\\<^bold>\\<^bold>\ u\ = \w\ \ \u\" assume vw: "Diag (Tensor v w)" have v: "Diag v \ v = Prim (un_Prim v)" using vw by (simp add: Diag_TensorE) have w: "Diag w" using vw by (simp add: Diag_TensorE) assume u: "Diag u" assume 1: "Arr v \ Arr w \ Arr u \ Dom v \<^bold>\ Dom w = Cod u" show "\(v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u\ = (\v\ \ \w\) \ \u\" using u 1 eval_in_hom CompDiag_in_Hom proof (cases u, simp_all) fix x y assume 3: "u = x \<^bold>\ y" assume 4: "Arr v \ Arr w \ Dom v = Cod x \ Dom w = Cod y" have x: "Diag x" using u 1 3 Diag_TensorE [of x y] by simp have y: "Diag y" using u x 1 3 Diag_TensorE [of x y] by simp show "\v \<^bold>\\<^bold>\\<^bold>\ x\ \ \w \<^bold>\\<^bold>\\<^bold>\ y\ = (\v\ \ \w\) \ (\x\ \ \y\)" using v w x y 4 I1 I2 CompDiag_in_Hom eval_in_hom Diag_implies_Arr interchange by auto qed qed thus ?thesis using assms by blast qed text\ For identity terms @{term a} and @{term b}, the reduction @{term "(a \<^bold>\ b)\<^bold>\"} factors (under evaluation in \D\) into the parallel reduction @{term "a\<^bold>\ \<^bold>\ b\<^bold>\"}, followed by a reduction of its codomain @{term "\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\"}. \ lemma eval_red_Tensor: assumes "Ide a" and "Ide b" shows "\(a \<^bold>\ b)\<^bold>\\ = \\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\\ \ (\a\<^bold>\\ \ \b\<^bold>\\)" proof - have "Diag (a \<^bold>\ b) \ ?thesis" using assms Can_red2 Ide_implies_Arr red_Diag Diagonalize_Diag red2_Diag Can_implies_Arr iso_eval_Can iso_is_arr apply simp using Diag_TensorE eval_Tensor Diagonalize_Diag Diag_implies_Arr red_Diag tensor_preserves_ide ide_eval_Ide dom_eval comp_arr_dom by metis moreover have "\ Diag (a \<^bold>\ b) \ ?thesis" using assms Can_red2 by (simp add: Can_red(1) iso_eval_Can) ultimately show ?thesis by blast qed lemma eval_red2_Diag_Unity: assumes "Ide a" and "Diag a" shows "\a \<^bold>\ \<^bold>\\ = \[\a\]" using assms tensor_preserves_ide \_ide_simp unitor_coincidence \_in_hom comp_cod_arr by (cases a, auto) text\ Define a formal arrow t to be ``coherent'' if the square formed by @{term t}, @{term "\<^bold>\t\<^bold>\"} and the reductions @{term "Dom t\<^bold>\"} and @{term "Cod t\<^bold>\"} commutes under evaluation in \D\. We will show that all formal arrows are coherent. Since the diagonalizations of canonical arrows are identities, a corollary is that parallel canonical arrows have equal evaluations. \ abbreviation coherent where "coherent t \ \Cod t\<^bold>\\ \ \t\ = \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" text\ Diagonal arrows are coherent, since for such arrows @{term t} the reductions @{term "Dom t\<^bold>\"} and @{term "Cod t\<^bold>\"} are identities. \ lemma Diag_implies_coherent: assumes "Diag t" shows "coherent t" using assms Diag_implies_Arr Arr_implies_Ide_Dom Arr_implies_Ide_Cod Dom_preserves_Diag Cod_preserves_Diag Diagonalize_Diag red_Diag comp_arr_dom comp_cod_arr by simp text\ The evaluation of a coherent arrow @{term t} has a canonical factorization in \D\ into the evaluations of a reduction @{term "Dom t\<^bold>\"}, diagonalization @{term "\<^bold>\t\<^bold>\"}, and inverse reduction @{term "Inv (Cod t\<^bold>\)"}. This will later allow us to use the term @{term "Inv (Cod t\<^bold>\) \<^bold>\ \<^bold>\t\<^bold>\ \<^bold>\ Dom t\<^bold>\"} as a normal form for @{term t}. \ lemma canonical_factorization: assumes "Arr t" shows "coherent t \ \t\ = inv \Cod t\<^bold>\\ \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" proof assume 1: "coherent t" have "inv \Cod t\<^bold>\\ \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\ = inv \Cod t\<^bold>\\ \ \Cod t\<^bold>\\ \ \t\" using 1 by simp also have "... = (inv \Cod t\<^bold>\\ \ \Cod t\<^bold>\\) \ \t\" using comp_assoc by simp also have "... = \t\" using assms 1 red_in_Hom inv_in_hom Arr_implies_Ide_Cod Can_red iso_eval_Can comp_cod_arr Ide_in_Hom inv_is_inverse by (simp add: comp_inv_arr) finally show "\t\ = inv \Cod t\<^bold>\\ \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" by simp next assume 1: "\t\ = inv \Cod t\<^bold>\\ \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" hence "\Cod t\<^bold>\\ \ \t\ = \Cod t\<^bold>\\ \ inv \Cod t\<^bold>\\ \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" by simp also have "... = (\Cod t\<^bold>\\ \ inv \Cod t\<^bold>\\) \ \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" using comp_assoc by simp also have "... = \\<^bold>\t\<^bold>\\ \ \Dom t\<^bold>\\" using assms 1 red_in_Hom Arr_implies_Ide_Cod Can_red iso_eval_Can inv_is_inverse Diagonalize_in_Hom comp_arr_inv comp_cod_arr Arr_implies_Ide_Dom Diagonalize_in_Hom by auto finally show "coherent t" by blast qed text\ A canonical arrow is coherent if and only if its formal inverse is. \ lemma Can_implies_coherent_iff_coherent_Inv: assumes "Can t" shows "coherent t \ coherent (Inv t)" proof have 1: "\t. Can t \ coherent t \ coherent (Inv t)" proof - fix t assume "Can t" hence t: "Can t \ Arr t \ Ide (Dom t) \ Ide (Cod t) \ arr \t\ \ iso \t\ \ inverse_arrows \t\ (inv \t\) \ Can \<^bold>\t\<^bold>\ \ Arr \<^bold>\t\<^bold>\ \ arr \\<^bold>\t\<^bold>\\ \ iso \\<^bold>\t\<^bold>\\ \ \<^bold>\t\<^bold>\ \ Hom \<^bold>\Dom t\<^bold>\ \<^bold>\Cod t\<^bold>\ \ inverse_arrows \\<^bold>\t\<^bold>\\ (inv \\<^bold>\t\<^bold>\\) \ Inv t \ Hom (Cod t) (Dom t)" using assms Can_implies_Arr Arr_implies_Ide_Dom Arr_implies_Ide_Cod iso_eval_Can inv_is_inverse Diagonalize_in_Hom Diagonalize_preserves_Can Inv_in_Hom by simp assume coh: "coherent t" have "\Cod (Inv t)\<^bold>\\ \ \Inv t\ = (inv \\<^bold>\t\<^bold>\\ \ \\<^bold>\t\<^bold>\\) \ \Cod (Inv t)\<^bold>\\ \ \Inv t\" using t red_in_Hom comp_cod_arr comp_inv_arr by (simp add: canonical_factorization coh Diagonalize_preserves_Can \Can t\ inv_is_inverse) also have "... = inv \\<^bold>\t\<^bold>\\ \ (\Cod t\<^bold>\\ \ \t\) \ inv \t\" using t eval_Inv_Can coh comp_assoc by auto also have "... = \\<^bold>\Inv t\<^bold>\\ \ \Dom (Inv t)\<^bold>\\" using t Diagonalize_Inv eval_Inv_Can comp_arr_inv red_in_Hom comp_arr_dom comp_assoc by auto finally show "coherent (Inv t)" by blast qed show "coherent t \ coherent (Inv t)" using assms 1 by simp show "coherent (Inv t) \ coherent t" proof - assume "coherent (Inv t)" hence "coherent (Inv (Inv t))" using assms 1 Inv_preserves_Can by blast thus ?thesis using assms by simp qed qed text\ Some special cases of coherence are readily dispatched. \ lemma coherent_Unity: shows "coherent \<^bold>\" by simp lemma coherent_Prim: assumes "Arr \<^bold>\f\<^bold>\" shows "coherent \<^bold>\f\<^bold>\" using assms by simp lemma coherent_Lunit_Ide: assumes "Ide a" shows "coherent \<^bold>\\<^bold>[a\<^bold>]" proof - have a: "Ide a \ Arr a \ Dom a = a \ Cod a = a \ ide \a\ \ ide \\<^bold>\a\<^bold>\\ \ \a\<^bold>\\ \ hom \a\ \\<^bold>\a\<^bold>\\" using assms Ide_implies_Arr Ide_in_Hom Diagonalize_preserves_Ide red_in_Hom by auto thus ?thesis using a lunit_naturality [of "\a\<^bold>\\"] comp_cod_arr by auto qed lemma coherent_Runit_Ide: assumes "Ide a" shows "coherent \<^bold>\\<^bold>[a\<^bold>]" proof - have a: "Ide a \ Arr a \ Dom a = a \ Cod a = a \ ide \a\ \ ide \\<^bold>\a\<^bold>\\ \ \a\<^bold>\\ \ hom \a\ \\<^bold>\a\<^bold>\\" using assms Ide_implies_Arr Ide_in_Hom Diagonalize_preserves_Ide red_in_Hom by auto have "\Cod \<^bold>\\<^bold>[a\<^bold>]\<^bold>\\ \ \\<^bold>\\<^bold>[a\<^bold>]\ = \a\<^bold>\\ \ \[\a\]" using a runit_in_hom comp_cod_arr by simp also have "... = \[\\<^bold>\a\<^bold>\\] \ (\a\<^bold>\\ \ \)" using a eval_Runit runit_naturality [of "\red a\"] by auto also have "... = \\<^bold>\\<^bold>\\<^bold>[a\<^bold>]\<^bold>\\ \ \Dom \<^bold>\\<^bold>[a\<^bold>]\<^bold>\\" proof - have "\ Diag (a \<^bold>\ \<^bold>\)" by (cases a; simp) thus ?thesis using a comp_cod_arr red2_in_Hom eval_red2_Diag_Unity Diag_Diagonalize Diagonalize_preserves_Ide by auto qed finally show ?thesis by blast qed lemma coherent_Lunit'_Ide: assumes "Ide a" shows "coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[a\<^bold>]" using assms Ide_implies_Can coherent_Lunit_Ide Can_implies_coherent_iff_coherent_Inv [of "Lunit a"] by simp lemma coherent_Runit'_Ide: assumes "Ide a" shows "coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[a\<^bold>]" using assms Ide_implies_Can coherent_Runit_Ide Can_implies_coherent_iff_coherent_Inv [of "Runit a"] by simp text\ To go further, we need the next result, which is in some sense the crux of coherence: For diagonal identities @{term a}, @{term b}, and @{term c}, the reduction @{term "((a \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c) \<^bold>\ ((a \<^bold>\ b) \<^bold>\ c)"} from @{term "(a \<^bold>\ b) \<^bold>\ c"} that first reduces the subterm @{term "a \<^bold>\ b"} and then reduces the result, is equivalent under evaluation in \D\ to the reduction that first applies the associator @{term "\<^bold>\\<^bold>[a, b, c\<^bold>]"} and then applies the reduction @{term "(a \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)) \<^bold>\ (a \<^bold>\ (b \<^bold>\ c))"} from @{term "a \<^bold>\ (b \<^bold>\ c)"}. The triangle and pentagon axioms are used in the proof. \ lemma coherence_key_fact: assumes "Ide a \ Diag a" and "Ide b \ Diag b" and "Ide c \ Diag c" shows "\(a \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\a \<^bold>\ b\ \ \c\) = (\a \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (\a\ \ \b \<^bold>\ c\)) \ \[\a\, \b\, \c\]" proof - have "b = \<^bold>\ \ ?thesis" using assms not_is_Tensor_TensorDiagE eval_red2_Diag_Unity triangle comp_cod_arr comp_assoc by simp text \The triangle is used!\ moreover have "c = \<^bold>\ \ ?thesis" using assms TensorDiag_preserves_Diag TensorDiag_preserves_Ide not_is_Tensor_TensorDiagE eval_red2_Diag_Unity red2_in_Hom runit_tensor runit_naturality [of "\a \<^bold>\ b\"] comp_assoc by simp moreover have "\ b \ \<^bold>\; c \ \<^bold>\ \ \ ?thesis" proof - assume b': "b \ \<^bold>\" hence b: "Ide b \ Diag b \ Arr b \ b \ \<^bold>\ \ ide \b\ \ arr \b\ \ \<^bold>\b\<^bold>\ = b \ b\<^bold>\ = b \ Dom b = b \ Cod b = b" using assms Diagonalize_preserves_Ide Ide_in_Hom by simp assume c': "c \ \<^bold>\" hence c: "Ide c \ Diag c \ Arr c \ c \ \<^bold>\ \ ide \c\ \ arr \c\ \ \<^bold>\c\<^bold>\ = c \ c\<^bold>\ = c \ Dom c = c \ Cod c = c" using assms Diagonalize_preserves_Ide Ide_in_Hom by simp have "\a. Ide a \ Diag a \ \(a \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\a \<^bold>\ b\ \ \c\) = (\a \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (\a\ \ \b \<^bold>\ c\)) \ \[\a\, \b\, \c\]" proof - fix a :: "'c term" show "Ide a \ Diag a \ \(a \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\a \<^bold>\ b\ \ \c\) = (\a \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (\a\ \ \b \<^bold>\ c\)) \ \[\a\, \b\, \c\]" apply (induct a) using b c TensorDiag_in_Hom apply simp_all proof - show "\b \<^bold>\ c\ \ (\b\ \ \[\b\] \ \c\) = ((\b \<^bold>\\<^bold>\\<^bold>\ c\ \ \[\b \<^bold>\\<^bold>\\<^bold>\ c\]) \ (\ \ \b \<^bold>\ c\)) \ \[\, \b\, \c\]" proof - have "\b \<^bold>\\<^bold>\\<^bold>\ c\ \ (\[\b \<^bold>\\<^bold>\\<^bold>\ c\] \ (\ \ \b \<^bold>\ c\)) \ \[\, \b\, \c\] = \b \<^bold>\\<^bold>\\<^bold>\ c\ \ (\b \<^bold>\ c\ \ \[\b\ \ \c\]) \ \[\, \b\, \c\]" using b c red2_in_Hom lunit_naturality [of "\b \<^bold>\ c\"] by simp thus ?thesis using b c red2_in_Hom lunit_tensor comp_arr_dom comp_cod_arr comp_assoc by simp qed show "\f. C.ide f \ C.arr f \ \(\<^bold>\f\<^bold>\ \<^bold>\ b) \<^bold>\ c\ \ (\\<^bold>\f\<^bold>\ \<^bold>\ b\ \ \c\) = (\\<^bold>\f\<^bold>\ \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (V f \ \b \<^bold>\ c\)) \ \[V f, \b\, \c\]" proof - fix f assume f: "C.ide f \ C.arr f" show "\(\<^bold>\f\<^bold>\ \<^bold>\ b) \<^bold>\ c\ \ (\\<^bold>\f\<^bold>\ \<^bold>\ b\ \ \c\) = (\\<^bold>\f\<^bold>\ \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (V f \ \b \<^bold>\ c\)) \ \[V f, \b\, \c\]" proof - have "\(\<^bold>\f\<^bold>\ \<^bold>\ b) \<^bold>\ c\ \ (\\<^bold>\f\<^bold>\ \<^bold>\ b\ \ \c\) = ((V f \ \b \<^bold>\\<^bold>\\<^bold>\ c\) \ (V f \ \b \<^bold>\ c\) \ \[V f, \b\, \c\]) \ ((V f \ \b\) \ \c\)" proof - have "\\<^bold>\f\<^bold>\ \<^bold>\ b\ = V f \ \b\" using assms f b c red2_Diag by simp moreover have "\\<^bold>\f\<^bold>\ \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\ = V f \ \b \<^bold>\\<^bold>\\<^bold>\ c\" proof - have "is_Tensor (b \<^bold>\\<^bold>\\<^bold>\ c)" using assms b c not_is_Tensor_TensorDiagE by blast thus ?thesis using assms f b c red2_Diag TensorDiag_preserves_Diag(1) by (cases "b \<^bold>\\<^bold>\\<^bold>\ c"; simp) qed ultimately show ?thesis using assms b c by (cases c, simp_all) qed also have "... = ((V f \ \b \<^bold>\\<^bold>\\<^bold>\ c\) \ (V f \ \b \<^bold>\ c\)) \ \[V f, \b\, \c\]" using b c f TensorDiag_in_Hom red2_in_Hom comp_arr_dom comp_cod_arr by simp also have "... = (\\<^bold>\f\<^bold>\ \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ (V f \ \b \<^bold>\ c\)) \ \[V f, \b\, \c\]" using b c f Ide_implies_Arr TensorDiag_preserves_Ide not_is_Tensor_TensorDiagE by (cases "b \<^bold>\\<^bold>\\<^bold>\ c", simp_all; blast) finally show ?thesis by blast qed qed fix d e assume I: "Diag e \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\e \<^bold>\ b\ \ \c\) = (\e \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\ \ (\e\ \ \b \<^bold>\ c\)) \ \[\e\, \b\, \c\]" assume de: "Ide d \ Ide e \ Diag (d \<^bold>\ e)" show "\((d \<^bold>\ e) \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\(d \<^bold>\ e) \<^bold>\ b\ \ \c\) = (\(d \<^bold>\ e) \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ ((\d\ \ \e\) \ \b \<^bold>\ c\)) \ \[\d\ \ \e\, \b\, \c\]" proof - let ?f = "un_Prim d" have "is_Prim d" using de by (cases d, simp_all) hence "d = \<^bold>\?f\<^bold>\ \ C.ide ?f" using de by (cases d, simp_all) hence d: "Ide d \ Arr d \ Dom d = d \ Cod d = d \ Diag d \ d = \<^bold>\?f\<^bold>\ \ C.ide ?f \ ide \d\ \ arr \d\" using de ide_eval_Ide Ide_implies_Arr Diag_Diagonalize(1) Ide_in_Hom Diag_TensorE [of d e] by simp have "Diag e \ e \ \<^bold>\" using de Diag_TensorE by metis hence e: "Ide e \ Arr e \ Dom e = e \ Cod e = e \ Diag e \ e \ \<^bold>\ \ ide \e\ \ arr \e\" using de Ide_in_Hom by simp have 1: "is_Tensor (e \<^bold>\\<^bold>\\<^bold>\ b) \ is_Tensor (b \<^bold>\\<^bold>\\<^bold>\ c) \ is_Tensor (e \<^bold>\\<^bold>\\<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))" using b c e de not_is_Tensor_TensorDiagE TensorDiag_preserves_Diag not_is_Tensor_TensorDiagE [of e "b \<^bold>\\<^bold>\\<^bold>\ c"] by auto have "\((d \<^bold>\ e) \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\(d \<^bold>\ e) \<^bold>\ b\ \ \c\) = ((\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c\) \ (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ \[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\]) \ ((\d\ \ \e \<^bold>\\<^bold>\\<^bold>\ b\) \ (\d\ \ \e \<^bold>\ b\) \ \[\d\, \e\, \b\] \ \c\)" proof - have "\((d \<^bold>\ e) \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ = (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c\) \ (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ \[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\]" proof - have "((d \<^bold>\ e) \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ b)) \<^bold>\ c" using b c d e de 1 TensorDiag_Diag TensorDiag_preserves_Diag TensorDiag_assoc TensorDiag_Prim not_is_Tensor_Unity by metis also have "... = (d \<^bold>\ (\<^bold>\e \<^bold>\\<^bold>\\<^bold>\ b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ c)) \<^bold>\ (d \<^bold>\ ((e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c)) \<^bold>\ \<^bold>\\<^bold>[d, e \<^bold>\\<^bold>\\<^bold>\ b, c\<^bold>]" using c d 1 by (cases c) simp_all also have "... = (d \<^bold>\ ((e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\\<^bold>\\<^bold>\ c)) \<^bold>\ (d \<^bold>\ ((e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c)) \<^bold>\ \<^bold>\\<^bold>[d, e \<^bold>\\<^bold>\\<^bold>\ b, c\<^bold>]" - using b c d e 1 TensorDiag_preserves_Diag red2_Diag TensorDiag_assoc - by (cases d) auto + by (metis 1 Diagonalize_Diag TensorDiag_assoc TensorDiag_preserves_Diag(1) + b c d e is_Tensor_def red2.simps(4)) finally show ?thesis using b c d e TensorDiag_in_Hom red2_in_Hom TensorDiag_preserves_Diag TensorDiag_preserves_Ide by simp qed moreover have "\(d \<^bold>\ e) \<^bold>\ b\ = (\d\ \ \e \<^bold>\\<^bold>\\<^bold>\ b\) \ (\d\ \ \e \<^bold>\ b\) \ \[\d\, \e\, \b\]" proof - have "(d \<^bold>\ e) \<^bold>\ b = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ b)) \<^bold>\ (d \<^bold>\ (e \<^bold>\ b)) \<^bold>\ \<^bold>\\<^bold>[d, e, b\<^bold>]" using b c d e de 1 TensorDiag_Prim Diagonalize_Diag by (cases b) simp_all also have "... = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ b)) \<^bold>\ (d \<^bold>\ (e \<^bold>\ b)) \<^bold>\ \<^bold>\\<^bold>[d, e, b\<^bold>]" using b d e 1 TensorDiag_preserves_Diag red2_Diag - by (cases d) auto + by (metis Diag.simps(3) de term.disc(12)) finally have "(d \<^bold>\ e) \<^bold>\ b = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ b)) \<^bold>\ (d \<^bold>\ (e \<^bold>\ b)) \<^bold>\ \<^bold>\\<^bold>[d, e, b\<^bold>]" by simp thus ?thesis using b d e eval_in_hom TensorDiag_in_Hom red2_in_Hom by simp qed ultimately show ?thesis by argo qed also have "... = (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ \[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\] \ ((\d\ \ \e \<^bold>\ b\) \ \c\) \ (\[\d\, \e\, \b\] \ \c\)" using b c d e red2_in_Hom TensorDiag_preserves_Ide TensorDiag_preserves_Diag interchange comp_cod_arr comp_assoc by simp also have "... = (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\) \ (\d\ \ (\e \<^bold>\ b\ \ \c\)) \ \[\d\, \e\ \ \b\, \c\] \ (\[\d\, \e\, \b\] \ \c\)" using b c d e TensorDiag_in_Hom red2_in_Hom TensorDiag_preserves_Ide TensorDiag_preserves_Diag assoc_naturality [of "\d\" "\e \<^bold>\ b\" "\c\"] comp_permute [of "\[\d\, \e \<^bold>\\<^bold>\\<^bold>\ b\, \c\]" "(\d\ \ \e \<^bold>\ b\) \ \c\" "\d\ \ (\e \<^bold>\ b\ \ \c\)" "\[\d\, \e\ \ \b\, \c\]"] by simp also have "... = (\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\e \<^bold>\ b\ \ \c\)) \ \[\d\, \e\ \ \b\, \c\] \ (\[\d\, \e\, \b\] \ \c\)" using b c d e TensorDiag_in_Hom red2_in_Hom TensorDiag_preserves_Ide TensorDiag_preserves_Diag interchange comp_reduce [of "\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\" "\d\ \ (\e \<^bold>\ b\ \ \c\)" "\d\ \ \(e \<^bold>\\<^bold>\\<^bold>\ b) \<^bold>\ c\ \ (\e \<^bold>\ b\ \ \c\)" "\[\d\, \e\ \ \b\, \c\] \ (\[\d\, \e\, \b\] \ \c\)"] by simp also have "... = (((\d\ \ \e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\) \ (\d\ \ \e\ \ \b \<^bold>\ c\)) \ (\d\ \ \[\e\, \b\, \c\])) \ \[\d\, \e\ \ \b\, \c\] \ (\[\d\, \e\, \b\] \ \c\)" using b c d e I TensorDiag_in_Hom red2_in_Hom TensorDiag_preserves_Ide TensorDiag_preserves_Diag interchange by simp also have "... = ((\d\ \ \e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\) \ (\d\ \ (\e\ \ \b \<^bold>\ c\))) \ \[\d\, \e\, \b\ \ \c\] \ \[\d\ \ \e\, \b\, \c\]" using b c d e comp_assoc red2_in_Hom TensorDiag_in_Hom ide_eval_Ide TensorDiag_preserves_Diag tensor_preserves_ide TensorDiag_preserves_Ide pentagon by simp text \The pentagon is used!\ also have "... = (((\d\ \ \e \<^bold>\\<^bold>\\<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\) \ (\d\ \ \e \<^bold>\ b \<^bold>\\<^bold>\\<^bold>\ c\) \ \[\d\, \e\, \b \<^bold>\\<^bold>\\<^bold>\ c\]) \ ((\d\ \ \e\) \ \b \<^bold>\ c\)) \ \[\d\ \ \e\, \b\, \c\]" - using b c d e 1 red2_in_Hom TensorDiag_preserves_Ide - TensorDiag_preserves_Diag assoc_naturality [of "\d\" "\e\" "\b \<^bold>\ c\"] - comp_cod_arr comp_assoc + using b c d e red2_in_Hom TensorDiag_preserves_Ide TensorDiag_preserves_Diag + assoc_naturality [of "\d\" "\e\" "\b \<^bold>\ c\"] comp_cod_arr comp_assoc by simp also have "... = (\(d \<^bold>\ e) \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ \ ((\d\ \ \e\) \ \b \<^bold>\ c\)) \ \[\d\ \ \e\, \b\, \c\]" proof - have "\(d \<^bold>\ e) \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\ = (\d\ \ \e \<^bold>\\<^bold>\\<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\) \ (\d\ \ \e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c)\) \ \[\d\, \e\, \b \<^bold>\\<^bold>\\<^bold>\ c\]" proof - have "(d \<^bold>\ e) \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c) = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b \<^bold>\\<^bold>\\<^bold>\ c\<^bold>\)) \<^bold>\ (d \<^bold>\ (e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ \<^bold>\\<^bold>[d, e, b \<^bold>\\<^bold>\\<^bold>\ c\<^bold>]" using b c e not_is_Tensor_TensorDiagE by (cases "b \<^bold>\\<^bold>\\<^bold>\ c") auto also have "... = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ (d \<^bold>\ (e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ \<^bold>\\<^bold>[d, e, b \<^bold>\\<^bold>\\<^bold>\ c\<^bold>]" using b c d e 1 TensorDiag_preserves_Diag Diagonalize_Diag by simp also have "... = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ (d \<^bold>\ (e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ \<^bold>\\<^bold>[d, e, b \<^bold>\\<^bold>\\<^bold>\ c\<^bold>]" using b c d e 1 TensorDiag_preserves_Diag(1) red2_Diag - by (cases d) auto + by (metis Diag.simps(3) de not_is_Tensor_Unity) finally have "(d \<^bold>\ e) \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c) = (d \<^bold>\ (e \<^bold>\\<^bold>\\<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ (d \<^bold>\ (e \<^bold>\ (b \<^bold>\\<^bold>\\<^bold>\ c))) \<^bold>\ \<^bold>\\<^bold>[d, e, b \<^bold>\\<^bold>\\<^bold>\ c\<^bold>]" by blast thus ?thesis using b c d e red2_in_Hom TensorDiag_in_Hom TensorDiag_preserves_Diag TensorDiag_preserves_Ide by simp qed thus ?thesis using d e b c by simp qed finally show ?thesis by simp qed qed qed thus ?thesis using assms(1) by blast qed ultimately show ?thesis by blast qed lemma coherent_Assoc_Ide: assumes "Ide a" and "Ide b" and "Ide c" shows "coherent \<^bold>\\<^bold>[a, b, c\<^bold>]" proof - have a: "Ide a \ Arr a \ Dom a = a \ Cod a = a \ ide \a\ \ ide \\<^bold>\a\<^bold>\\ \ \\a\<^bold>\\ : \a\ \ \\<^bold>\a\<^bold>\\\" using assms Ide_implies_Arr Ide_in_Hom Diagonalize_preserves_Ide red_in_Hom by auto have b: "Ide b \ Arr b \ Dom b = b \ Cod b = b \ ide \b\ \ ide \\<^bold>\b\<^bold>\\ \ \\b\<^bold>\\ : \b\ \ \\<^bold>\b\<^bold>\\\" using assms Ide_implies_Arr Ide_in_Hom Diagonalize_preserves_Ide red_in_Hom by auto have c: "Ide c \ Arr c \ Dom c = c \ Cod c = c \ ide \c\ \ ide \\<^bold>\c\<^bold>\\ \ \\c\<^bold>\\ : \c\ \ \\<^bold>\c\<^bold>\\\" using assms Ide_implies_Arr Ide_in_Hom Diagonalize_preserves_Ide red_in_Hom by auto have "\Cod \<^bold>\\<^bold>[a, b, c\<^bold>]\<^bold>\\ \ \\<^bold>\\<^bold>[a, b, c\<^bold>]\ = (\\<^bold>\a\<^bold>\ \<^bold>\ (\<^bold>\b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\)\ \ (\\<^bold>\a\<^bold>\\ \ (\\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\\)) \ (\a\<^bold>\\ \ \b\<^bold>\\ \ \c\<^bold>\\)) \ \[\a\, \b\, \c\]" using a b c red_in_Hom red2_in_Hom Diagonalize_in_Hom Diag_Diagonalize Diagonalize_preserves_Ide interchange Ide_in_Hom eval_red_Tensor comp_cod_arr [of "\a\<^bold>\\"] by simp also have "... = ((\\<^bold>\a\<^bold>\ \<^bold>\ (\<^bold>\b\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\c\<^bold>\)\ \ (\\<^bold>\a\<^bold>\\ \ \\<^bold>\b\<^bold>\ \<^bold>\ \<^bold>\c\<^bold>\\)) \ \[\\<^bold>\a\<^bold>\\, \\<^bold>\b\<^bold>\\, \\<^bold>\c\<^bold>\\]) \ ((\a\<^bold>\\ \ \b\<^bold>\\) \ \c\<^bold>\\)" using a b c red_in_Hom Diag_Diagonalize TensorDiag_preserves_Diag assoc_naturality [of "\a\<^bold>\\" "\b\<^bold>\\" "\c\<^bold>\\"] comp_assoc by simp also have "... = (\(\<^bold>\a\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b\<^bold>\) \<^bold>\ \<^bold>\c\<^bold>\\ \ (\\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\\ \ \\<^bold>\c\<^bold>\\)) \ ((\a\<^bold>\\ \ \b\<^bold>\\) \ \c\<^bold>\\)" using a b c Diag_Diagonalize Diagonalize_preserves_Ide coherence_key_fact by simp also have "... = \\<^bold>\\<^bold>\\<^bold>[a, b, c\<^bold>]\<^bold>\\ \ \Dom \<^bold>\\<^bold>[a, b, c\<^bold>]\<^bold>\\" using a b c red_in_Hom red2_in_Hom TensorDiag_preserves_Diag Diagonalize_preserves_Ide TensorDiag_preserves_Ide Diag_Diagonalize interchange eval_red_Tensor TensorDiag_assoc comp_cod_arr [of "\c\<^bold>\\"] comp_cod_arr [of "\(\<^bold>\a\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\b\<^bold>\) \<^bold>\ \<^bold>\c\<^bold>\\ \ (\\<^bold>\a\<^bold>\ \<^bold>\ \<^bold>\b\<^bold>\\ \ (\a\<^bold>\\ \ \b\<^bold>\\) \ \c\<^bold>\\)"] comp_assoc by simp finally show ?thesis by blast qed lemma coherent_Assoc'_Ide: assumes "Ide a" and "Ide b" and "Ide c" shows "coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[a, b, c\<^bold>]" proof - have "Can \<^bold>\\<^bold>[a, b, c\<^bold>]" using assms Ide_implies_Can by simp moreover have "\<^bold>\\<^sup>-\<^sup>1\<^bold>[a, b, c\<^bold>] = Inv \<^bold>\\<^bold>[a, b, c\<^bold>]" using assms Inv_Ide by simp ultimately show ?thesis using assms Ide_implies_Can coherent_Assoc_Ide Inv_Ide Can_implies_coherent_iff_coherent_Inv by metis qed text\ The next lemma implies coherence for the special case of a term that is the tensor of two diagonal arrows. \ lemma eval_red2_naturality: assumes "Diag t" and "Diag u" shows "\Cod t \<^bold>\ Cod u\ \ (\t\ \ \u\) = \t \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom t \<^bold>\ Dom u\" proof - have *: "\t u. Diag (t \<^bold>\ u) \ arr \t\ \ arr \u\" using Diag_implies_Arr by force have "t = \<^bold>\ \ ?thesis" using assms Diag_implies_Arr lunit_naturality [of "\u\"] Arr_implies_Ide_Dom Arr_implies_Ide_Cod comp_cod_arr by simp moreover have "t \ \<^bold>\ \ u = \<^bold>\ \ ?thesis" using assms Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diag_implies_Arr Dom_preserves_Diag Cod_preserves_Diag eval_red2_Diag_Unity runit_naturality [of "\t\"] by simp moreover have "t \ \<^bold>\ \ u \ \<^bold>\ \ ?thesis" using assms * Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diag_implies_Arr Dom_preserves_Diag Cod_preserves_Diag apply (induct t, simp_all) proof - fix f assume f: "C.arr f" assume "u \ \<^bold>\" hence u: "u \ \<^bold>\ \ Diag u \ Diag (Dom u) \ Diag (Cod u) \ Ide (Dom u) \ Ide (Cod u) \ arr \u\ \ arr \Dom u\ \ arr \Cod u\ \ ide \Dom u\ \ ide \Cod u\" using assms(2) Diag_implies_Arr Dom_preserves_Diag Cod_preserves_Diag Arr_implies_Ide_Dom Arr_implies_Ide_Cod by simp hence 1: "Dom u \ \<^bold>\ \ Cod u \ \<^bold>\" using u by (cases u, simp_all) show "\\<^bold>\C.cod f\<^bold>\ \<^bold>\ Cod u\ \ (V f \ \u\) = (V f \ \u\) \ \\<^bold>\C.dom f\<^bold>\ \<^bold>\ Dom u\" using f u 1 Diag_implies_Arr red2_Diag comp_arr_dom comp_cod_arr by simp next fix v w assume I2: "\ w \ Unity; Diag w \ \ \Cod w \<^bold>\ Cod u\ \ (\w\ \ \u\) = \w \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom w \<^bold>\ Dom u\" assume "u \ \<^bold>\" hence u: "u \ \<^bold>\ \ Arr u \ Arr (Dom u) \ Arr (Cod u) \ Diag u \ Diag (Dom u) \ Diag (Cod u) \ Ide (Dom u) \ Ide (Cod u) \ arr \u\ \ arr \Dom u\ \ arr \Cod u\ \ ide \Dom u\ \ ide \Cod u\" using assms(2) Diag_implies_Arr Dom_preserves_Diag Cod_preserves_Diag Arr_implies_Ide_Dom Arr_implies_Ide_Cod by simp assume vw: "Diag (v \<^bold>\ w)" let ?f = "un_Prim v" have "v = \<^bold>\?f\<^bold>\ \ C.arr ?f" using vw by (metis Diag_TensorE(1) Diag_TensorE(2)) hence "Arr v \ v = \<^bold>\un_Prim v\<^bold>\ \ C.arr ?f \ Diag v" by (cases v; simp) hence v: "v = \<^bold>\?f\<^bold>\ \ C.arr ?f \ Arr v \ Ide (Dom v) \ Ide (Cod v) \ Diag v \ Diag (Dom v) \ arr \v\ \ arr \Dom v\ \ arr \Cod v\ \ ide \Dom v\ \ ide \Cod v\" by (cases v, simp_all) have "Diag w \ w \ \<^bold>\" using vw v by (metis Diag.simps(3)) hence w: "w \ \<^bold>\ \ Arr w \ Arr (Dom w) \ Arr (Cod w) \ Diag w \ Diag (Dom w) \ Diag (Cod w) \ Ide (Dom w) \ Ide (Cod w) \ arr \w\ \ arr \Dom w\ \ arr \Cod w\ \ ide \Dom w\ \ ide \Cod w\" using vw * Diag_implies_Arr Dom_preserves_Diag Cod_preserves_Diag Arr_implies_Ide_Dom Arr_implies_Ide_Cod ide_eval_Ide Ide_implies_Arr Ide_in_Hom by simp show "\(Cod v \<^bold>\ Cod w) \<^bold>\ Cod u\ \ ((\v\ \ \w\) \ \u\) = \(v \<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u\ \ \(Dom v \<^bold>\ Dom w) \<^bold>\ Dom u\" proof - have u': "Dom u \ \<^bold>\ \ Cod u \ \<^bold>\" using u by (cases u) simp_all have w': "Dom w \ \<^bold>\ \ Cod w \ \<^bold>\" using w by (cases w) simp_all have D: "Diag (Dom v \<^bold>\ (Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u))" proof - have "Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u \ \<^bold>\" using u u' w w' not_is_Tensor_TensorDiagE by blast moreover have "Diag (Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u)" using u w TensorDiag_preserves_Diag by simp moreover have "Dom v = \<^bold>\C.dom ?f\<^bold>\" using v by (cases v, simp_all) ultimately show ?thesis using u v w TensorDiag_preserves_Diag by auto qed have C: "Diag (Cod v \<^bold>\ (Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u))" proof - have "Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u \ \<^bold>\" using u u' w w' not_is_Tensor_TensorDiagE by blast moreover have "Diag (Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u)" using u w TensorDiag_preserves_Diag by simp moreover have "Cod v = \<^bold>\C.cod ?f\<^bold>\" using v by (cases v, simp_all) ultimately show ?thesis using u v w by (cases "Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u") simp_all qed have "\(Cod v \<^bold>\ Cod w) \<^bold>\ Cod u\ \ ((\v\ \ \w\) \ \u\) = (\Cod v \<^bold>\ (Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u)\ \ (\Cod v\ \ \Cod w \<^bold>\ Cod u\) \ \[\Cod v\, \Cod w\, \Cod u\]) \ ((\v\ \ \w\) \ \u\)" proof - have "(Cod v \<^bold>\ Cod w) \<^bold>\ Cod u = (Cod v \<^bold>\ (Cod w \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Cod u\<^bold>\)) \<^bold>\ (Cod v \<^bold>\ Cod w \<^bold>\ Cod u) \<^bold>\ \<^bold>\\<^bold>[Cod v, Cod w, Cod u\<^bold>]" using u v w by (cases u, simp_all) hence "\(Cod v \<^bold>\ Cod w) \<^bold>\ Cod u\ = \Cod v \<^bold>\ (Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u)\ \ (\Cod v\ \ \Cod w \<^bold>\ Cod u\) \ \[\Cod v\, \Cod w\, \Cod u\]" using u v w by simp thus ?thesis by argo qed also have "... = ((\Cod v\ \ \Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u\) \ (\Cod v\ \ \Cod w \<^bold>\ Cod u\) \ \[\Cod v\, \Cod w\, \Cod u\]) \ ((\v\ \ \w\) \ \u\)" using u v w C red2_Diag by simp also have "... = ((\Cod v\ \ \Cod w \<^bold>\ Cod u\) \ \[\Cod v\, \Cod w\, \Cod u\]) \ ((\v\ \ \w\) \ \u\)" proof - have "(\Cod v\ \ \Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u\) \ (\Cod v\ \ \Cod w \<^bold>\ Cod u\) = \Cod v\ \ \Cod w \<^bold>\ Cod u\" using u v w comp_cod_arr red2_in_Hom by simp moreover have "seq (\Cod v\ \ \Cod w \<^bold>\\<^bold>\\<^bold>\ Cod u\) (\Cod v\ \ \Cod w \<^bold>\ Cod u\)" using u v w red2_in_Hom TensorDiag_in_Hom Ide_in_Hom by simp moreover have "seq (\Cod v\ \ \Cod w \<^bold>\ Cod u\) \[\Cod v\, \Cod w\, \Cod u\]" using u v w red2_in_Hom by simp ultimately show ?thesis using u v w comp_reduce by presburger qed also have "... = (\v\ \ \w \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using u v w I2 red2_in_Hom TensorDiag_in_Hom interchange comp_reduce assoc_naturality [of "\v\" "\w\" "\u\"] comp_cod_arr comp_assoc by simp also have "... = (\v\ \ \w \<^bold>\\<^bold>\\<^bold>\ u\) \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using u v w red2_in_Hom TensorDiag_in_Hom interchange comp_reduce comp_arr_dom by simp also have "... = \v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using u u' v w not_is_Tensor_TensorDiagE TensorDiag_Prim [of "w \<^bold>\\<^bold>\\<^bold>\ u" ?f] by force also have "... = \v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom v \<^bold>\\<^bold>\\<^bold>\ Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" proof - have "\v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\] = (\v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom v \<^bold>\\<^bold>\\<^bold>\ Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u\) \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using u v w comp_arr_dom TensorDiag_in_Hom TensorDiag_preserves_Diag by simp also have "... = \v \<^bold>\\<^bold>\\<^bold>\ w \<^bold>\\<^bold>\\<^bold>\ u\ \ \Dom v \<^bold>\\<^bold>\\<^bold>\ Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using comp_assoc by simp finally show ?thesis by blast qed also have "... = \(v \<^bold>\\<^bold>\\<^bold>\ w) \<^bold>\\<^bold>\\<^bold>\ u\ \ \(Dom v \<^bold>\ Dom w) \<^bold>\ Dom u\" proof - have "\(Dom v \<^bold>\ Dom w) \<^bold>\ Dom u\ = \Dom v \<^bold>\ (Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u)\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" proof - have "(Dom v \<^bold>\ Dom w) \<^bold>\ Dom u = (Dom v \<^bold>\ (Dom w \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom u\<^bold>\)) \<^bold>\ (Dom v \<^bold>\ (Dom w \<^bold>\ Dom u)) \<^bold>\ \<^bold>\\<^bold>[Dom v, Dom w, Dom u\<^bold>]" using u u' v w red2_in_Hom TensorDiag_in_Hom tensor_in_hom Ide_in_Hom by (cases u) auto thus ?thesis using u v w red2_in_Hom by simp qed also have "... = \Dom v \<^bold>\\<^bold>\\<^bold>\ Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" using D TensorDiag_Diag red2_Diag by simp finally have "\(Dom v \<^bold>\ Dom w) \<^bold>\ Dom u\ = \Dom v \<^bold>\\<^bold>\\<^bold>\ Dom w \<^bold>\\<^bold>\\<^bold>\ Dom u\ \ (\Dom v\ \ \Dom w \<^bold>\ Dom u\) \ \[\Dom v\, \Dom w\, \Dom u\]" by blast thus ?thesis using assms v w TensorDiag_assoc by auto qed finally show ?thesis using vw TensorDiag_Diag by simp qed qed ultimately show ?thesis by blast qed lemma Tensor_preserves_coherent: assumes "Arr t" and "Arr u" and "coherent t" and "coherent u" shows "coherent (t \<^bold>\ u)" proof - have t: "Arr t \ Ide (Dom t) \ Ide (Cod t) \ Ide \<^bold>\Dom t\<^bold>\ \ Ide \<^bold>\Cod t\<^bold>\ \ arr \t\ \ arr \Dom t\ \ ide \Dom t\ \ arr \Cod t\ \ ide \Cod t\" using assms Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diagonalize_preserves_Ide by auto have u: "Arr u \ Ide (Dom u) \ Ide (Cod u) \ Ide \<^bold>\Dom u\<^bold>\ \ Ide \<^bold>\Cod u\<^bold>\ \ arr \u\ \ arr \Dom u\ \ ide \Dom u\ \ arr \Cod u\ \ ide \Cod u\" using assms Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diagonalize_preserves_Ide by auto have "\Cod (t \<^bold>\ u)\<^bold>\\ \ (\t\ \ \u\) = (\\<^bold>\Cod t\<^bold>\ \<^bold>\ \<^bold>\Cod u\<^bold>\\ \ (\Cod t\<^bold>\\ \ \Cod u\<^bold>\\)) \ (\t\ \ \u\)" using t u eval_red_Tensor by simp also have "... = \\<^bold>\Cod t\<^bold>\ \<^bold>\ \<^bold>\Cod u\<^bold>\\ \ (\Cod t\<^bold>\\ \ \Cod u\<^bold>\\) \ (\t\ \ \u\)" using comp_assoc by simp also have "... = \\<^bold>\Cod t\<^bold>\ \<^bold>\ \<^bold>\Cod u\<^bold>\\ \ (\\<^bold>\t\<^bold>\\ \ \\<^bold>\u\<^bold>\\) \ (\Dom t\<^bold>\\ \ \Dom u\<^bold>\\)" using assms t u Diagonalize_in_Hom red_in_Hom interchange by simp also have "... = (\\<^bold>\Cod t\<^bold>\ \<^bold>\ \<^bold>\Cod u\<^bold>\\ \ (\\<^bold>\t\<^bold>\\ \ \\<^bold>\u\<^bold>\\)) \ (\Dom t\<^bold>\\ \ \Dom u\<^bold>\\)" using comp_assoc by simp also have "... = (\\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\\ \ \\<^bold>\Dom t\<^bold>\ \<^bold>\ \<^bold>\Dom u\<^bold>\\) \ (\Dom t\<^bold>\\ \ \Dom u\<^bold>\\)" using assms t u Diag_Diagonalize Diagonalize_in_Hom eval_red2_naturality [of "Diagonalize t" "Diagonalize u"] by simp also have "... = \\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\\ \ \\<^bold>\Dom t\<^bold>\ \<^bold>\ \<^bold>\Dom u\<^bold>\\ \ (\Dom t\<^bold>\\ \ \Dom u\<^bold>\\)" using comp_assoc by simp also have "... = \\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\\ \ \(Dom t \<^bold>\ Dom u)\<^bold>\\" using t u eval_red_Tensor by simp finally have "\Cod (t \<^bold>\ u)\<^bold>\\ \ (\t\ \ \u\) = \\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\\ \ \(Dom t \<^bold>\ Dom u)\<^bold>\\" by blast thus ?thesis using t u by simp qed lemma Comp_preserves_coherent: assumes "Arr t" and "Arr u" and "Dom t = Cod u" and "coherent t" and "coherent u" shows "coherent (t \<^bold>\ u)" proof - have t: "Arr t \ Ide (Dom t) \ Ide (Cod t) \ Ide \<^bold>\Dom t\<^bold>\ \ Ide \<^bold>\Cod t\<^bold>\ \ arr \t\ \ arr \Dom t\ \ ide \Dom t\ \ arr \Cod t\ \ ide \Cod t\" using assms Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diagonalize_preserves_Ide by auto have u: "Arr u \ Ide (Dom u) \ Ide (Cod u) \ Ide \<^bold>\Dom u\<^bold>\ \ Ide \<^bold>\Cod u\<^bold>\ \ arr \u\ \ arr \Dom u\ \ ide \Dom u\ \ arr \Cod u\ \ ide \Cod u\" using assms Arr_implies_Ide_Dom Arr_implies_Ide_Cod Diagonalize_preserves_Ide by auto have "\Cod (t \<^bold>\ u)\<^bold>\\ \ \t \<^bold>\ u\ = \Cod t\<^bold>\\ \ \t\ \ \u\" using t u by simp also have "... = (\Cod t\<^bold>\\ \ \t\) \ \u\" proof - (* TODO: I still haven't figured out how to do this without spoon-feeding it. *) have "seq \Cod t\<^bold>\\ \t\" using assms t red_in_Hom by (intro seqI, auto) moreover have "seq \t\ \u\" using assms t u by auto ultimately show ?thesis using comp_assoc by auto qed also have "... = \\<^bold>\t \<^bold>\ u\<^bold>\\ \ \Dom (t \<^bold>\ u)\<^bold>\\" using t u assms red_in_Hom Diag_Diagonalize comp_assoc by (simp add: Diag_implies_Arr eval_CompDiag) finally show "coherent (t \<^bold>\ u)" by blast qed text\ The main result: ``Every formal arrow is coherent.'' \ theorem coherence: assumes "Arr t" shows "coherent t" proof - have "Arr t \ coherent t" proof (induct t) fix u v show "\ Arr u \ coherent u; Arr v \ coherent v \ \ Arr (u \<^bold>\ v) \ coherent (u \<^bold>\ v)" using Tensor_preserves_coherent by simp show "\ Arr u \ coherent u; Arr v \ coherent v \ \ Arr (u \<^bold>\ v) \ coherent (u \<^bold>\ v)" using Comp_preserves_coherent by simp next show "coherent \<^bold>\" by simp fix f show "Arr \<^bold>\f\<^bold>\ \ coherent \<^bold>\f\<^bold>\" by simp next fix t assume I: "Arr t \ coherent t" show Lunit: "Arr \<^bold>\\<^bold>[t\<^bold>] \ coherent \<^bold>\\<^bold>[t\<^bold>]" using I Arr_implies_Ide_Dom coherent_Lunit_Ide Ide_in_Hom Ide_implies_Arr Comp_preserves_coherent [of t "\<^bold>\\<^bold>[Dom t\<^bold>]"] Diagonalize_Comp_Arr_Dom \_ide_simp by auto show Runit: "Arr \<^bold>\\<^bold>[t\<^bold>] \ coherent \<^bold>\\<^bold>[t\<^bold>]" using I Arr_implies_Ide_Dom coherent_Runit_Ide Ide_in_Hom Ide_implies_Arr Comp_preserves_coherent [of t "\<^bold>\\<^bold>[Dom t\<^bold>]"] Diagonalize_Comp_Arr_Dom \_ide_simp by auto show "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]" proof - assume "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]" hence t: "Arr t" by simp have "coherent (\<^bold>\\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>] \<^bold>\ t)" using t I Arr_implies_Ide_Cod coherent_Lunit'_Ide Ide_in_Hom Comp_preserves_coherent [of "\<^bold>\\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>]" t] by fastforce thus ?thesis using t Arr_implies_Ide_Cod Ide_implies_Arr Ide_in_Hom Diagonalize_Comp_Cod_Arr eval_in_hom \'.is_natural_2 [of "\t\"] by force qed show "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>] \ coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]" proof - assume "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t\<^bold>]" hence t: "Arr t" by simp have "coherent (\<^bold>\\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>] \<^bold>\ t)" using t I Arr_implies_Ide_Cod coherent_Runit'_Ide Ide_in_Hom Comp_preserves_coherent [of "\<^bold>\\<^sup>-\<^sup>1\<^bold>[Cod t\<^bold>]" t] by fastforce thus ?thesis using t Arr_implies_Ide_Cod Ide_implies_Arr Ide_in_Hom Diagonalize_Comp_Cod_Arr eval_in_hom \'.is_natural_2 [of "\t\"] by force qed next fix t u v assume I1: "Arr t \ coherent t" assume I2: "Arr u \ coherent u" assume I3: "Arr v \ coherent v" show "Arr \<^bold>\\<^bold>[t, u, v\<^bold>] \ coherent \<^bold>\\<^bold>[t, u, v\<^bold>]" proof - assume tuv: "Arr \<^bold>\\<^bold>[t, u, v\<^bold>]" have t: "Arr t" using tuv by simp have u: "Arr u" using tuv by simp have v: "Arr v" using tuv by simp have "coherent ((t \<^bold>\ u \<^bold>\ v) \<^bold>\ \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>])" proof - have "Arr (t \<^bold>\ u \<^bold>\ v) \ coherent (t \<^bold>\ u \<^bold>\ v)" proof have 1: "Arr t \ coherent t" using t I1 by simp have 2: "Arr (u \<^bold>\ v) \ coherent (u \<^bold>\ v)" using u v I2 I3 Tensor_preserves_coherent by force show "Arr (t \<^bold>\ u \<^bold>\ v) " using 1 2 by simp show "coherent (t \<^bold>\ u \<^bold>\ v)" using 1 2 Tensor_preserves_coherent by blast qed moreover have "Arr \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v Arr_implies_Ide_Dom by simp moreover have "coherent \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v Arr_implies_Ide_Dom coherent_Assoc_Ide by blast moreover have "Dom (t \<^bold>\ u \<^bold>\ v) = Cod \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v Arr_implies_Ide_Dom Ide_in_Hom by simp ultimately show ?thesis using t u v Arr_implies_Ide_Dom Ide_implies_Arr Comp_preserves_coherent [of "t \<^bold>\ u \<^bold>\ v" "\<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>]"] by blast qed moreover have "Par \<^bold>\\<^bold>[t, u, v\<^bold>] ((t \<^bold>\ u \<^bold>\ v) \<^bold>\ \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>])" using t u v Arr_implies_Ide_Dom Ide_implies_Arr Ide_in_Hom by simp moreover have "\<^bold>\\<^bold>\\<^bold>[t, u, v\<^bold>]\<^bold>\ = \<^bold>\(t \<^bold>\ u \<^bold>\ v) \<^bold>\ \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>]\<^bold>\" proof - have "(\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\ = (\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ ((\<^bold>\Dom t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom u\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom v\<^bold>\)" proof - have 1: "Diag \<^bold>\t\<^bold>\ \ Diag \<^bold>\u\<^bold>\ \ Diag \<^bold>\v\<^bold>\ \ Dom \<^bold>\t\<^bold>\ = \<^bold>\Dom t\<^bold>\ \ Dom \<^bold>\u\<^bold>\ = \<^bold>\Dom u\<^bold>\ \ Dom \<^bold>\v\<^bold>\ = \<^bold>\Dom v\<^bold>\" using t u v Diag_Diagonalize by blast moreover have "Diag (\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\)" using 1 TensorDiag_preserves_Diag(1) by blast moreover have "\t. Arr t \ \<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\Dom t\<^bold>\ = \<^bold>\t\<^bold>\" using t Diagonalize_Comp_Arr_Dom by simp moreover have "Dom \<^bold>\\<^bold>\\<^bold>[t, u, v\<^bold>]\<^bold>\ = \<^bold>\Dom \<^bold>\\<^bold>[t, u, v\<^bold>]\<^bold>\" using Diag_Diagonalize tuv by blast ultimately show ?thesis using t u v tuv 1 TensorDiag_assoc TensorDiag_preserves_Diag(2) by (metis (no_types) Diagonalize.simps(9)) qed thus ?thesis using t u v Diagonalize_Comp_Arr_Dom CompDiag_TensorDiag Diag_Diagonalize by simp qed moreover have "\\<^bold>\\<^bold>[t, u, v\<^bold>]\ = \(t \<^bold>\ u \<^bold>\ v) \<^bold>\ \<^bold>\\<^bold>[Dom t, Dom u, Dom v\<^bold>]\" using t u v Arr_implies_Ide_Dom Ide_implies_Arr \_simp [of "\t\" "\u\" "\v\"] by simp ultimately show "coherent \<^bold>\\<^bold>[t, u, v\<^bold>]" by argo qed show "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] \ coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]" proof - assume tuv: "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]" have t: "Arr t" using tuv by simp have u: "Arr u" using tuv by simp have v: "Arr v" using tuv by simp have "coherent (((t \<^bold>\ u) \<^bold>\ v) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>])" proof - have "Arr ((t \<^bold>\ u) \<^bold>\ v) \ coherent ((t \<^bold>\ u) \<^bold>\ v)" proof have 1: "Arr v \ coherent v" using v I3 by simp have 2: "Arr (t \<^bold>\ u) \ coherent (t \<^bold>\ u)" using t u I1 I2 Tensor_preserves_coherent by force show "Arr ((t \<^bold>\ u) \<^bold>\ v)" using 1 2 by simp show "coherent ((t \<^bold>\ u) \<^bold>\ v)" using 1 2 Tensor_preserves_coherent by blast qed moreover have "Arr \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v Arr_implies_Ide_Dom by simp moreover have "coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v Arr_implies_Ide_Dom coherent_Assoc'_Ide by blast moreover have "Dom ((t \<^bold>\ u) \<^bold>\ v) = Cod \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]" using t u v Arr_implies_Ide_Dom Ide_in_Hom by simp ultimately show ?thesis using t u v Arr_implies_Ide_Dom Ide_implies_Arr Comp_preserves_coherent [of "((t \<^bold>\ u) \<^bold>\ v)" "\<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]"] by metis qed moreover have "Par \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>] (((t \<^bold>\ u) \<^bold>\ v) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>])" using t u v Arr_implies_Ide_Dom Ide_implies_Arr Ide_in_Hom by simp moreover have "\<^bold>\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\<^bold>\ = \<^bold>\((t \<^bold>\ u) \<^bold>\ v) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]\<^bold>\" using t u v Diagonalize_Comp_Arr_Dom CompDiag_TensorDiag Diag_Diagonalize TensorDiag_assoc TensorDiag_preserves_Diag TensorDiag_in_Hom CompDiag_Diag_Dom [of "(\<^bold>\t\<^bold>\ \<^bold>\\<^bold>\\<^bold>\ \<^bold>\u\<^bold>\) \<^bold>\\<^bold>\\<^bold>\ \<^bold>\v\<^bold>\"] by simp moreover have "\\<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]\ = \((t \<^bold>\ u) \<^bold>\ v) \<^bold>\ \<^bold>\\<^sup>-\<^sup>1\<^bold>[Dom t, Dom u, Dom v\<^bold>]\" using t u v Arr_implies_Ide_Dom Ide_implies_Arr eval_in_hom comp_cod_arr \'.is_natural_1 \'_simp by simp ultimately show "coherent \<^bold>\\<^sup>-\<^sup>1\<^bold>[t, u, v\<^bold>]" by argo qed qed thus ?thesis using assms by blast qed text\ MacLane \cite{MacLane71} says: ``A coherence theorem asserts `Every diagram commutes','' but that is somewhat misleading. A coherence theorem provides some kind of hopefully useful way of distinguishing diagrams that definitely commute from diagrams that might not. The next result expresses coherence for monoidal categories in this way. As the hypotheses can be verified algorithmically (using the functions @{term Dom}, @{term Cod}, @{term Arr}, and @{term Diagonalize}) if we are given an oracle for equality of arrows in \C\, the result provides a decision procedure, relative to \C\, for the word problem for the free monoidal category generated by \C\. \ corollary eval_eqI: assumes "Par t u" and "\<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" shows "\t\ = \u\" using assms coherence canonical_factorization by simp text\ Our final corollary expresses coherence in a more ``MacLane-like'' fashion: parallel canonical arrows are equivalent under evaluation. \ corollary maclane_coherence: assumes "Par t u" and "Can t" and "Can u" shows "\t\ = \u\" proof (intro eval_eqI) show "Par t u" by fact show "\<^bold>\t\<^bold>\ = \<^bold>\u\<^bold>\" proof - have "Ide \<^bold>\t\<^bold>\ \ Ide \<^bold>\u\<^bold>\ \ Par \<^bold>\t\<^bold>\ \<^bold>\u\<^bold>\" using assms eval_eqI Ide_Diagonalize_Can Diagonalize_in_Hom by simp thus ?thesis using Ide_in_Hom by auto qed qed end end

\<^sub>1[?R, ?w1]" let ?\ = "\?R \r0, r0\ ?R\" have P: "?P = ru.apex" using ru.apex_composite by auto have Chn_\: "\?\ : ?V \\<^sub>C ?P\" using P Chn_in_hom \ by force have p0\: "\?p0 \ ?\ : ?V \\<^sub>C ?U\" using Chn_\ ru.legs_form_cospan by auto have w1: "\?w1 : ?V \\<^sub>C ?R\" using Chn_\ ru.legs_form_cospan r.dom.apex_def by blast have r1w1: "r1 \ ?w1 = ?v1" by (metis C.comp_assoc T.base_simps(3) \ \.leg1_commutes arrow_of_spans_data.select_convs(3) cod_char dom_char ideD(1) ideD(2) in_homE ru.composite_in_hom ru.leg1_composite u v) have w: "ide w" unfolding w_def using P \ w1 by (intro ide_mkIde, auto) interpret w: identity_arrow_of_spans C w using w_def w ide_char' by auto have hseq_fw: "hseq f w" using w_def ide_f w src_def trg_def w1 by auto interpret fw: two_composable_arrows_of_spans C prj0 prj1 f w using w_def hseq_fw hseq_char by unfold_locales auto interpret fw: two_composable_identity_arrows_of_spans C prj0 prj1 f w .. have hseq_gw: "hseq g w" using w_def ide_g w src_def trg_def w1 by auto interpret gw: two_composable_arrows_of_spans C prj0 prj1 g w using hseq_gw hseq_char by unfold_locales auto interpret gw: two_composable_identity_arrows_of_spans C prj0 prj1 g w .. interpret rfw: three_composable_arrows_of_spans C prj0 prj1 r f w .. interpret rfw: three_composable_identity_arrows_of_spans C prj0 prj1 r f w .. have arfw: "\\[r, f, w] : (r \ f) \ w \ r \ f \ w\" using fw.composable ide_f ide_r w rf.composable by auto have fw_apex_eq: "fw.apex = ?R \\ ?w1" using w_def fw.dom.apex_def hcomp_def hseq_fw hseq_char \ C.arr_dom_iff_arr C.pbdom_def fw.chine_eq_apex fw.chine_simps(1) by auto have gw_apex_eq: "gw.apex = ?R \\ ?w1" using w_def \ w1 gw.dom.apex_def hcomp_def hseq_gw hseq_char by auto text \ Well, CKS say take \\ = p\<^sub>0\\, but taking this literally and trying to define \\\ so that \Chn \ = ?p\<^sub>0 \ ?\\, does not yield the required \dom \ = ?R \\ ?w1\. We need \\Chn \ : ?R \\ ?w1 \\<^sub>C ?U\\, so we have to compose with a projection, which leads to: \Chn \ = ?p0 \ ?\ \ \